diff --git a/metadata/metadata b/metadata/metadata --- a/metadata/metadata +++ b/metadata/metadata @@ -1,10689 +1,10777 @@ [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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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)
[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 operations 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)
[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 = eberlm@in.tum.de [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 = eberlm@in.tum.de [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 = eberlm@in.tum.de [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 = eberlm@in.tum.de [Euler_MacLaurin] title = The Euler–MacLaurin Formula author = Manuel Eberl topic = Mathematics/Analysis date = 2017-03-10 notify = eberlm@in.tum.de 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 [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 = eberlm@in.tum.de [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 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 input-output automata. 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 [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 = eberlm@in.tum.de [Error_Function] title = The Error Function author = Manuel Eberl topic = Mathematics/Analysis date = 2018-02-06 notify = eberlm@in.tum.de 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 = eberlm@in.tum.de [Dirichlet_Series] title = Dirichlet Series author = Manuel Eberl topic = Mathematics/Number theory date = 2017-10-12 notify = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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, eberlm@in.tum.de +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 = eberlm@in.tum.de 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 [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 [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 = eberlm@in.tum.de [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 = eberlm@in.tum.de [Median_Of_Medians_Selection] title = The Median-of-Medians Selection Algorithm author = Manuel Eberl topic = Computer science/Algorithms date = 2017-12-21 notify = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 a rectangle box or a 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). [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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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.

[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. [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 = eberlm@in.tum.de 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)
[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. [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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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 = eberlm@in.tum.de 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. + diff --git a/thys/IMP_Compiler/Compiler.thy b/thys/IMP_Compiler/Compiler.thy new file mode 100755 --- /dev/null +++ b/thys/IMP_Compiler/Compiler.thy @@ -0,0 +1,227 @@ +(* Title: A Shorter Compiler Correctness Proof for Language IMP + Author: Pasquale Noce + Software Engineer at HID Global, Italy + pasquale dot noce dot lavoro at gmail dot com + pasquale dot noce at hidglobal dot com +*) + +section "Compiler formalization" + +theory Compiler + imports + "HOL-IMP.BExp" + "HOL-IMP.Star" +begin + + +subsection "Introduction" + +text \ +This paper presents a compiler correctness proof for the didactic imperative programming language +IMP, introduced in @{cite "Nipkow14"}, shorter than the proof described in @{cite "Nipkow14"} and +included in the Isabelle2021 distribution @{cite "Klein21"}. Actually, the size of the presented +proof is just two thirds of the book's proof in the number of formal text lines, and as such it +promises to constitute a further enhanced reference for the formal verification of compilers meant +for larger, real-world programming languages. + +Given compiler \emph{completeness}, viz. the simulation of source code execution by compiled code, +"in a deterministic language like IMP", compiler correctness "reduces to preserving termination: if +the machine program terminates, so must the source program", even though proving this "is not much +easier" (@{cite "Nipkow14"}, section 8.4). However, the presented proof does not depend on language +determinism, so that the proposed approach is applicable to non-deterministic languages as well. + +As a confirmation, this paper extends IMP with an additional command @{text "c\<^sub>1 OR c\<^sub>2"}, standing +for the non-deterministic choice between commands @{text c\<^sub>1} and @{text c\<^sub>2}, and proves compiler +\emph{correctness}, viz. the simulation of compiled code execution by source code, for such extended +language. Of course, the aforesaid comparison between proof sizes does not consider the lines in the +proof of lemma @{text ccomp_correct} (which proves compiler correctness for commands) pertaining to +non-deterministic choice, since this command is not included in the original language IMP. Anyway, +non-deterministic choice turns out to extend that proof just by a modest number of lines. + +For further information about the formal definitions and proofs contained in this paper, see +Isabelle documentation, particularly @{cite "Paulson21"}, @{cite "Nipkow21"}, @{cite "Krauss21"}, +and @{cite "Nipkow11"}. +\ + + +subsection "Definitions" + +text \ +Here below are the definitions of IMP commands, extended with non-deterministic choice, as well as +of their big-step semantics. + +As in the original theory file @{cite "Klein21"}, program counter's values are modeled using type +@{typ int} rather than @{typ nat}. As a result, the same declarations and definitions used in +@{cite "Klein21"} to deal with this modeling choice are adopted here as well. + +\null +\ + +declare [[coercion_enabled]] +declare [[coercion "int :: nat \ int"]] +declare [[syntax_ambiguity_warning = false]] + +datatype com = + SKIP | + Assign vname aexp ("_ ::= _" [1000, 61] 61) | + Seq com com ("_;;/ _" [60, 61] 60) | + If bexp com com ("(IF _/ THEN _/ ELSE _)" [0, 0, 61] 61) | + Or com com ("(_ OR _)" [60, 61] 61) | + While bexp com ("(WHILE _/ DO _)" [0, 61] 61) + +inductive big_step :: "com \ state \ state \ bool" (infix "\" 55) where +Skip: "(SKIP, s) \ s" | +Assign: "(x ::= a, s) \ s(x := aval a s)" | +Seq: "\(c\<^sub>1, s\<^sub>1) \ s\<^sub>2; (c\<^sub>2, s\<^sub>2) \ s\<^sub>3\ \ (c\<^sub>1;; c\<^sub>2, s\<^sub>1) \ s\<^sub>3" | +IfTrue: "\bval b s; (c\<^sub>1, s) \ t\ \ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \ t" | +IfFalse: "\\ bval b s; (c\<^sub>2, s) \ t\ \ (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s) \ t" | +Or1: "(c\<^sub>1, s) \ t \ (c\<^sub>1 OR c\<^sub>2, s) \ t" | +Or2: "(c\<^sub>2, s) \ t \ (c\<^sub>1 OR c\<^sub>2, s) \ t" | +WhileFalse: "\ bval b s \ (WHILE b DO c, s) \ s" | +WhileTrue: "\bval b s\<^sub>1; (c, s\<^sub>1) \ s\<^sub>2; (WHILE b DO c, s\<^sub>2) \ s\<^sub>3\ \ + (WHILE b DO c, s\<^sub>1) \ s\<^sub>3" + +declare big_step.intros [intro] + +abbreviation (output) +"isize xs \ int (length xs)" + +notation isize ("size") + +primrec (nonexhaustive) inth :: "'a list \ int \ 'a" (infixl "!!" 100) where +"(x # xs) !! i = (if i = 0 then x else xs !! (i - 1))" + +lemma inth_append [simp]: + "0 \ i \ + (xs @ ys) !! i = (if i < size xs then xs !! i else ys !! (i - size xs))" +by (induction xs arbitrary: i, auto simp: algebra_simps) + +text \ +\null + +Next, the instruction set and its semantics are defined. Particularly, to allow for the compilation +of non-deterministic choice commands, the instruction set is extended with an additional instruction +@{text JMPND} performing a non-deterministic jump -- viz. as a result of its execution, the program +counter unconditionally either jumps by the specified offset, or just moves to the next instruction. + +As instruction execution can be non-deterministic, an inductively defined predicate @{text iexec}, +rather than a simple non-recursive function as the one used in @{cite "Klein21"}, must be introduced +to define instruction semantics. + +\null +\ + +datatype instr = + LOADI int | LOAD vname | ADD | STORE vname | + JMP int | JMPLESS int | JMPGE int | JMPND int + +type_synonym stack = "val list" +type_synonym config = "int \ state \ stack" + +abbreviation "hd2 xs \ hd (tl xs)" +abbreviation "tl2 xs \ tl (tl xs)" + +inductive iexec :: "instr \ config \ config \ bool" (infix "\" 55) where +LoadI: "(LOADI i, pc, s, stk) \ (pc + 1, s, i # stk)" | +Load: "(LOAD x, pc, s, stk) \ (pc + 1, s, s x # stk)" | +Add: "(ADD, pc, s, stk) \ (pc + 1, s, (hd2 stk + hd stk) # tl2 stk)" | +Store: "(STORE x, pc, s, stk) \ (pc + 1, s(x := hd stk), tl stk)" | +Jmp: "(JMP i, pc, s, stk) \ (pc + i + 1, s, stk)" | +JmpLessY: "hd2 stk < hd stk \ + (JMPLESS i, pc, s, stk) \ (pc + i + 1, s, tl2 stk)" | +JmpLessN: "hd stk \ hd2 stk \ + (JMPLESS i, pc, s, stk) \ (pc + 1, s, tl2 stk)" | +JmpGeY: "hd stk \ hd2 stk \ + (JMPGE i, pc, s, stk) \ (pc + i + 1, s, tl2 stk)" | +JmpGeN: "hd2 stk < hd stk \ + (JMPGE i, pc, s, stk) \ (pc + 1, s, tl2 stk)" | +JmpNdY: "(JMPND i, pc, s, stk) \ (pc + i + 1, s, stk)" | +JmpNdN: "(JMPND i, pc, s, stk) \ (pc + 1, s, stk)" + +declare iexec.intros [intro] + +inductive_cases LoadIE [elim!]: "(LOADI i, pc, s, stk) \ cf" +inductive_cases LoadE [elim!]: "(LOAD x, pc, s, stk) \ cf" +inductive_cases AddE [elim!]: "(ADD, pc, s, stk) \ cf" +inductive_cases StoreE [elim!]: "(STORE x, pc, s, stk) \ cf" +inductive_cases JmpE [elim!]: "(JMP i, pc, s, stk) \ cf" +inductive_cases JmpLessE [elim!]: "(JMPLESS i, pc, s, stk) \ cf" +inductive_cases JmpGeE [elim!]: "(JMPGE i, pc, s, stk) \ cf" +inductive_cases JmpNdE [elim!]: "(JMPND i, pc, s, stk) \ cf" + +definition exec1 :: "instr list \ config \ config \ bool" + ("(_/ \/ _/ \/ _)" 55) where +"P \ cf \ cf' \ (P !! fst cf, cf) \ cf' \ 0 \ fst cf \ fst cf < size P" + +abbreviation exec :: "instr list \ config \ config \ bool" + ("(_/ \/ _/ \*/ _)" 55) where +"exec P \ star (exec1 P)" + +text \ +\null + +Next, compilation is formalized for arithmetic and boolean expressions (functions @{text acomp} and +@{text bcomp}), as well as for commands (function @{text ccomp}). Particularly, as opposed to what +happens in @{cite "Klein21"}, here @{text bcomp} takes a single input, viz. a 3-tuple comprised of +a boolean expression, a flag, and a jump offset. In this way, all three functions accept a single +input, which enables to streamline the compiler correctness proof developed in what follows. + +\null +\ + +primrec acomp :: "aexp \ instr list" where +"acomp (N i) = [LOADI i]" | +"acomp (V x) = [LOAD x]" | +"acomp (Plus a\<^sub>1 a\<^sub>2) = acomp a\<^sub>1 @ acomp a\<^sub>2 @ [ADD]" + +fun bcomp :: "bexp \ bool \ int \ instr list" where +"bcomp (Bc v, f, i) = (if v = f then [JMP i] else [])" | +"bcomp (Not b, f, i) = bcomp (b, \ f, i)" | +"bcomp (And b\<^sub>1 b\<^sub>2, f, i) = + (let cb\<^sub>2 = bcomp (b\<^sub>2, f, i); + cb\<^sub>1 = bcomp (b\<^sub>1, False, size cb\<^sub>2 + (if f then 0 else i)) + in cb\<^sub>1 @ cb\<^sub>2)" | +"bcomp (Less a\<^sub>1 a\<^sub>2, f, i) = + acomp a\<^sub>1 @ acomp a\<^sub>2 @ (if f then [JMPLESS i] else [JMPGE i])" + +primrec ccomp :: "com \ instr list" where +"ccomp SKIP = []" | +"ccomp (x ::= a) = acomp a @ [STORE x]" | +"ccomp (c\<^sub>1;; c\<^sub>2) = ccomp c\<^sub>1 @ ccomp c\<^sub>2" | +"ccomp (IF b THEN c\<^sub>1 ELSE c\<^sub>2) = + (let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2; cb = bcomp (b, False, size cc\<^sub>1 + 1) + in cb @ cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" | +"ccomp (c\<^sub>1 OR c\<^sub>2) = + (let cc\<^sub>1 = ccomp c\<^sub>1; cc\<^sub>2 = ccomp c\<^sub>2 + in JMPND (size cc\<^sub>1 + 1) # cc\<^sub>1 @ JMP (size cc\<^sub>2) # cc\<^sub>2)" | +"ccomp (WHILE b DO c) = + (let cc = ccomp c; cb = bcomp (b, False, size cc + 1) + in cb @ cc @ [JMP (- (size cb + size cc + 1))])" + +text \ +\null + +Finally, two lemmas are proven automatically (both seem not to be included in the standard library, +though being quite basic) and registered for use by automatic proof tactics. In more detail: + + \<^item> The former lemma is an elimination rule similar to @{thm [source] impCE}, with the difference +that it retains the antecedent of the implication in the premise where the consequent is assumed to +hold. This rule permits to have both assumptions @{prop "\ bval b s"} and @{prop "bval b s"} in the +respective cases resulting from the execution of boolean expression @{text b} in state @{text s}. + + \<^item> The latter one is an introduction rule similar to @{thm [source] Suc_lessI}, with the difference +that its second assumption is more convenient for proving statements of the form @{prop "Suc m < n"} +arising from the compiler correctness proof developed in what follows. + +\null +\ + +lemma impCE2 [elim!]: + "\P \ Q; \ P \ R; P \ Q \ R\ \ R" +by blast + +lemma Suc_lessI2 [intro!]: + "\m < n; m \ n - 1\ \ Suc m < n" +by simp + +end diff --git a/thys/IMP_Compiler/Compiler2.thy b/thys/IMP_Compiler/Compiler2.thy new file mode 100755 --- /dev/null +++ b/thys/IMP_Compiler/Compiler2.thy @@ -0,0 +1,549 @@ +(* Title: A Shorter Compiler Correctness Proof for Language IMP + Author: Pasquale Noce + Software Engineer at HID Global, Italy + pasquale dot noce dot lavoro at gmail dot com + pasquale dot noce at hidglobal dot com +*) + +section "Compiler correctness" + +theory Compiler2 + imports Compiler +begin + + +subsection "Preliminary definitions and lemmas" + +text \ +Now everything is ready for the compiler correctness proof. First, two predicates are introduced, +@{text execl} and @{text execl_all}, both taking as inputs a program, i.e. a list of instructions, +@{text P} and a list of program configurations @{text cfs}, and respectively denoted using notations +@{text "P \ cfs"} and @{text "P \ cfs\"}. In more detail: + + \<^item> @{text "P \ cfs"} means that program @{text P} \emph{may} transform each configuration within +@{text cfs} into the subsequent one, if any (word \emph{may} reflects the fact that programs can be +non-deterministic in this case study). +\\Thus, @{text execl} formalizes the notion of a \emph{small-step} program execution. + + \<^item> @{text "P \ cfs\"} reinforces @{text "P \ cfs"} by additionally requiring that @{text cfs} be +nonempty, the initial program counter be zero (viz. execution starts from the first instruction in +@{text P}), and the final program counter falls outside @{text P} (viz. execution terminates). +\\Thus, @{text execl_all} formalizes the notion of a \emph{complete small-step} program execution, +so that assumptions @{text "acomp a \ cfs\"}, @{text "bcomp x \ cfs\"}, @{text "ccomp c \ cfs\"} +will be used in the compiler correctness proofs for arithmetic/boolean expressions and commands. + +Moreover, predicates @{text apred}, @{text bpred}, and @{text cpred} are defined to capture the link +between the initial and the final configuration upon the execution of an arithmetic expression, a +boolean expression, and a whole program, respectively, and abbreviation @{text off} is introduced as +a commodity to shorten the subsequent formal text. + +\null +\ + +fun execl :: "instr list \ config list \ bool" (infix "\" 55) where +"P \ cf # cf' # cfs = (P \ cf \ cf' \ P \ cf' # cfs)" | +"P \ _ = True" + +definition execl_all :: "instr list \ config list \ bool" ("(_/ \/ _\)" 55) where +"P \ cfs\ \ P \ cfs \ cfs \ [] \ + fst (cfs ! 0) = 0 \ fst (cfs ! (length cfs - 1)) \ {0.. config \ config \ bool" where +"apred \ \a (pc, s, stk) (pc', s', stk'). + pc' = pc + size (acomp a) \ s' = s \ stk' = aval a s # stk" + +definition bpred :: "bexp \ bool \ int \ config \ config \ bool" where +"bpred \ \(b, f, i) (pc, s, stk) (pc', s', stk'). + pc' = pc + size (bcomp (b, f, i)) + (if bval b s = f then i else 0) \ + s' = s \ stk' = stk" + +definition cpred :: "com \ config \ config \ bool" where +"cpred \ \c (pc, s, stk) (pc', s', stk'). + pc' = pc + size (ccomp c) \ (c, s) \ s' \ stk' = stk" + +abbreviation off :: "instr list \ config \ config" where +"off P cf \ (fst cf - size P, snd cf)" + +text \ +\null + +Next, some lemmas about @{const [source] execl} and @{const execl_all} are proven. In more detail, +given a program @{text P} and a list of configurations @{text cfs} such that @{prop "P \ cfs"}: + + \<^item> Lemma @{text execl_next} states that for any configuration in @{text cfs} but the last one, the +subsequent configuration must result from the execution of the referenced instruction of @{text P} +in that configuration. +\\Thus, @{text execl_next} permits to reproduce the execution of a single instruction. + + \<^item> Lemma @{text execl_last} states that a configuration in @{text cfs} whose program counter falls +outside @{text P} must be the last one in @{text cfs}. +\\Thus, @{text execl_last} permits to infer the completion of program execution. + + \<^item> Lemma @{text execl_drop} states that @{prop "P \ drop n cfs"} for any natural number @{text n}, +and will be used to prove compiler correctness for loops by induction over the length of the list of +configurations @{text cfs}. + +Furthermore, some other lemmas enabling to prove compiler correctness automatically for constructors +@{const N}, @{const V} (arithmetic expressions), @{const Bc} (boolean expressions) and @{const SKIP} +(commands) are also proven. + +\null +\ + +lemma iexec_offset_aux: + "(i :: int) + 1 - j = i - j + 1 \ i + j - k + 1 = i - k + j + 1" +by simp + +lemma iexec_offset [intro]: + "(ins, pc, s, stk) \ (pc', s', stk') \ + (ins, pc - i, s, stk) \ (pc' - i, s', stk')" +by (erule iexec.cases, auto simp: iexec_offset_aux) + +lemma execl_next: + "\P \ cfs; k < length cfs; k \ length cfs - 1\ \ + (P !! fst (cfs ! k), cfs ! k) \ cfs ! Suc k" +by (induction cfs arbitrary: k rule: execl.induct, auto + simp: nth_Cons exec1_def split: nat.split) + +lemma execl_last: + "\P \ cfs; k < length cfs; fst (cfs ! k) \ {0.. \ + length cfs - 1 = k" +by (induction cfs arbitrary: k rule: execl.induct, auto + simp: nth_Cons exec1_def split: nat.split_asm) + +lemma execl_drop: + "P \ cfs \ P \ drop n cfs" +by (induction cfs arbitrary: n rule: execl.induct, + simp_all add: drop_Cons split: nat.split) + +lemma execl_all_N [simplified, intro]: + "[LOADI i] \ cfs\ \ apred (N i) (cfs ! 0) (cfs ! (length cfs - 1))" +by (clarsimp simp: execl_all_def apred_def, cases "cfs ! 0", + subgoal_tac "length cfs - 1 = 1", frule_tac [!] execl_next, + ((rule ccontr)?, fastforce, (rule execl_last)?)+) + +lemma execl_all_V [simplified, intro]: + "[LOAD x] \ cfs\ \ apred (V x) (cfs ! 0) (cfs ! (length cfs - 1))" +by (clarsimp simp: execl_all_def apred_def, cases "cfs ! 0", + subgoal_tac "length cfs - 1 = 1", frule_tac [!] execl_next, + ((rule ccontr)?, fastforce, (rule execl_last)?)+) + +lemma execl_all_Bc [simplified, intro]: + "\if v = f then [JMP i] else [] \ cfs\; 0 \ i\ \ + bpred (Bc v, f, i) (cfs ! 0) (cfs ! (length cfs - 1))" +by (clarsimp simp: execl_all_def bpred_def split: if_split_asm, cases "cfs ! 0", + subgoal_tac "length cfs - 1 = 1", frule_tac [1-2] execl_next, + ((rule ccontr)?, force, (rule execl_last)?)+, rule execl.cases [of "([], cfs)"], + auto simp: exec1_def) + +lemma execl_all_SKIP [simplified, intro]: + "[] \ cfs\ \ cpred SKIP (cfs ! 0) (cfs ! (length cfs - 1))" +by (rule execl.cases [of "([], cfs)"], auto simp: execl_all_def exec1_def cpred_def) + +text \ +\null + +Next, lemma @{text execl_all_sub} is proven. It states that, if @{prop "P @ P' x @ P'' \ cfs\"}, +configuration @{text cf} within @{text cfs} refers to the start of program @{text "P' x"}, and the +initial and the final configuration in every complete execution of @{text "P' x"} satisfy predicate +@{text "Q x"}, then there exists a configuration @{text cf'} in @{text cfs} such that @{text cf} and +@{text cf'} satisfy @{text "Q x"}. +\\Thus, this lemma permits to reproduce the execution of a subprogram, particularly: + + \<^item> a compiled arithmetic expression @{text a}, where @{prop "Q = apred"} and @{prop "x = a"}, + + \<^item> a compiled boolean expression @{text b}, where @{prop "Q = bpred"} and @{prop "x = (b, f, i)"} +(given a flag @{text f} and a jump offset @{text i}), and + + \<^item> a compiled command @{text c}, where @{prop "Q = cpred"} and @{prop "x = c"}. + +Furthermore, lemma @{text execl_all_sub2} is derived from @{text execl_all_sub} to enable a shorter +symbolical execution of two consecutive subprograms. + +\null +\ + +lemma execl_sub_aux: + "\\m n. \k \ {m.. P' \ + map (off P) (case m of 0 \ (pc, s, stk) # take n cfs | Suc m \ F cfs m n); + \k \ {m.. \ + P' \ (pc - size P, s, stk) # map (off P) (take n cfs)" + (is "\\_ _. \k \ _. Q P' (?F k) \ _; \k \ ?A. Q P' (?G k)\ \ _") +by (subgoal_tac "\k \ {0..k \ {0.. ?A \ ?F k = ?G (k + m + length cfs')", + fastforce, simp add: nth_append) + +lemma execl_sub: + "\P @ P' @ P'' \ cfs; \k \ {m.. fst (cfs ! k) \ fst (cfs ! k) - size P < size P'\ \ + P' \ map (off P) (drop m (take (Suc n) cfs))" + (is "\_; \k \ _. ?P P' cfs k\ \ P' \ map _ (?F cfs m (Suc n))") +proof (induction cfs arbitrary: m n rule: execl.induct [of _ P'], auto + simp: take_Cons drop_Cons exec1_def split: nat.split, force, force, force, + erule execl_sub_aux [where m = 0], subst append_Cons [of _ "[]"], simp, + erule execl_sub_aux [where m = "Suc 0" and cfs' = "[]"], simp) + fix P' pc s stk cfs m n + let ?cf = "(pc, s, stk) :: config" + assume "\m n. \k \ {m.. P' \ + map (off P) (case m of 0 \ ?cf # take n cfs | Suc m \ ?F cfs m n)" + moreover assume "\k \ {Suc (Suc m)..k \ {Suc m.. map (off P) (?F cfs m n)" + by fastforce +qed + +lemma execl_all_sub [rule_format]: + assumes + A: "P @ P' x @ P'' \ cfs\" and + B: "k < length cfs" and + C: "fst (cfs ! k) = size P" and + D: "\cfs. P' x \ cfs\ \ Q x (cfs ! 0) (cfs ! (length cfs - 1))" + shows "\k' < length cfs. Q x (off P (cfs ! k)) (off P (cfs ! k'))" +proof - + let ?P = "\k. size P \ fst (cfs ! k) \ fst (cfs ! k) - size P < size (P' x)" + let ?A = "{k'. k' \ {k.. \ ?P k'}" + have E: "Min ?A \ ?A" + using A and B by (rule_tac Min_in, simp_all add: execl_all_def, + rule_tac exI [of _ "length cfs - 1"], auto) + hence "map (off P) (drop k (take (Suc (Min ?A)) cfs)) ! 0 = off P (cfs ! k)" + (is "?cfs ! _ = _") + by auto + moreover have "length cfs \ Suc (Min ?A) \ Min ?A = length cfs - 1" + using E by auto + with A and B and E have F: "?cfs ! (length ?cfs - 1) = off P (cfs ! Min ?A)" + by (subst nth_map, auto simp: min_def execl_all_def, arith) + hence "?cfs \ [] \ fst (?cfs ! (length ?cfs - 1)) \ {0.. (\k'. k' \ {k'. k' \ {k.. \ ?P k'})" + by (rule notI, erule exE, frule rev_subsetD [of _ _ ?A], rule subsetI, + insert E, simp, subgoal_tac "finite ?A", drule Min_le, force+) + hence "P' x \ ?cfs" + using A by (subst (asm) execl_all_def, rule_tac execl_sub, blast+) + ultimately have "Q x (?cfs ! 0) (?cfs ! (length ?cfs - 1))" + using C and D by (auto simp: execl_all_def) + thus ?thesis + using E and F by (rule_tac exI [of _ "Min ?A"], auto) +qed + +lemma execl_all_sub2: + assumes + A: "P x @ P' x' @ P'' \ cfs\" + (is "?P \ _\") and + B: "\cfs. P x \ cfs\ \ (\(pc, s, stk) (pc', s', stk'). + pc' = pc + size (P x) + I s \ Q s s' \ stk' = F s stk) + (cfs ! 0) (cfs ! (length cfs - 1))" + (is "\cfs. _ \ ?Q x (cfs ! 0) (cfs ! (length cfs - 1))") and + C: "\cfs. P' x' \ cfs\ \ (\(pc, s, stk) (pc', s', stk'). + pc' = pc + size (P' x') + I' s \ Q' s s' \ stk' = F' s stk) + (cfs ! 0) (cfs ! (length cfs - 1))" + (is "\cfs. _ \ ?Q' x' (cfs ! 0) (cfs ! (length cfs - 1))") and + D: "I (fst (snd (cfs ! 0))) = 0" + shows "\k < length cfs. \t. (\(pc, s, stk) (pc', s', stk'). + pc = 0 \ pc' = size (P x) + size (P' x') + I' t \ Q s t \ Q' t s' \ + stk' = F' t (F s stk)) (cfs ! 0) (cfs ! k)" +by (subgoal_tac "[] @ ?P \ cfs\", drule execl_all_sub [where k = 0 and Q = ?Q], + insert A B, (clarsimp simp: execl_all_def)+, insert A C D, drule execl_all_sub + [where Q = ?Q'], simp+, clarify, rule exI, rule conjI, simp, rule exI, auto) + + +subsection "Main theorem" + +text \ +It is time to prove compiler correctness. First, lemmas @{text acomp_acomp}, @{text bcomp_bcomp} are +derived from @{thm [source] execl_all_sub2} to reproduce the execution of two consecutive compiled +arithmetic expressions (possibly generated by both @{const acomp} and @{const bcomp}) and boolean +expressions (possibly generated by @{const bcomp}), respectively. Subsequently, the correctness of +@{const acomp} and @{const bcomp} is proven in lemmas @{text acomp_correct}, @{text bcomp_correct}. + +\null +\ + +lemma acomp_acomp: + "\acomp a\<^sub>1 @ acomp a\<^sub>2 @ P \ cfs\; + \cfs. acomp a\<^sub>1 \ cfs\ \ apred a\<^sub>1 (cfs ! 0) (cfs ! (length cfs - 1)); + \cfs. acomp a\<^sub>2 \ cfs\ \ apred a\<^sub>2 (cfs ! 0) (cfs ! (length cfs - 1))\ \ + case cfs ! 0 of (pc, s, stk) \ pc = 0 \ (\k < length cfs. cfs ! k = + (size (acomp a\<^sub>1 @ acomp a\<^sub>2), s, aval a\<^sub>2 s # aval a\<^sub>1 s # stk))" +by (drule execl_all_sub2 [where I = "\s. 0" and I' = "\s. 0" and Q = "\s s'. s' = s" + and Q' = "\s s'. s' = s" and F = "\s stk. aval a\<^sub>1 s # stk" + and F' = "\s stk. aval a\<^sub>2 s # stk"], auto simp: apred_def) + +lemma bcomp_bcomp: + "\bcomp (b\<^sub>1, f\<^sub>1, i\<^sub>1) @ bcomp (b\<^sub>2, f\<^sub>2, i\<^sub>2) \ cfs\; + \cfs. bcomp (b\<^sub>1, f\<^sub>1, i\<^sub>1) \ cfs\ \ + bpred (b\<^sub>1, f\<^sub>1, i\<^sub>1) (cfs ! 0) (cfs ! (length cfs - 1)); + \cfs. bcomp (b\<^sub>2, f\<^sub>2, i\<^sub>2) \ cfs\ \ + bpred (b\<^sub>2, f\<^sub>2, i\<^sub>2) (cfs ! 0) (cfs ! (length cfs - 1))\ \ + case cfs ! 0 of (pc, s, stk) \ pc = 0 \ (bval b\<^sub>1 s \ f\<^sub>1 \ + (\k < length cfs. cfs ! k = (size (bcomp (b\<^sub>1, f\<^sub>1, i\<^sub>1) @ bcomp (b\<^sub>2, f\<^sub>2, i\<^sub>2)) + + (if bval b\<^sub>2 s = f\<^sub>2 then i\<^sub>2 else 0), s, stk)))" +by (clarify, rule conjI, simp add: execl_all_def, rule impI, subst (asm) append_Nil2 + [symmetric], drule execl_all_sub2 [where I = "\s. if bval b\<^sub>1 s = f\<^sub>1 then i\<^sub>1 else 0" + and I' = "\s. if bval b\<^sub>2 s = f\<^sub>2 then i\<^sub>2 else 0" and Q = "\s s'. s' = s" + and Q' = "\s s'. s' = s" and F = "\s stk. stk" and F' = "\s stk. stk"], + auto simp: bpred_def) + +lemma acomp_correct [simplified, intro]: + "acomp a \ cfs\ \ apred a (cfs ! 0) (cfs ! (length cfs - 1))" +proof (induction a arbitrary: cfs, simp_all, frule_tac [3] acomp_acomp, auto) + fix a\<^sub>1 a\<^sub>2 cfs s stk k + assume A: "acomp a\<^sub>1 @ acomp a\<^sub>2 @ [ADD] \ cfs\" + (is "?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i \ _\") + assume B: "k < length cfs" and + C: "cfs ! k = (size ?ca\<^sub>1 + size ?ca\<^sub>2, s, aval a\<^sub>2 s # aval a\<^sub>1 s # stk)" + hence "cfs ! Suc k = (size (?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i), s, aval (Plus a\<^sub>1 a\<^sub>2) s # stk)" + using A by (insert execl_next [of "?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i" cfs k], + simp add: execl_all_def, drule_tac impI, auto) + thus "apred (Plus a\<^sub>1 a\<^sub>2) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A and B and C by (insert execl_last [of "?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i" cfs "Suc k"], + simp add: execl_all_def apred_def, drule_tac impI, auto) +qed + +lemma bcomp_correct [simplified, intro]: + "\bcomp x \ cfs\; 0 \ snd (snd x)\ \ bpred x (cfs ! 0) (cfs ! (length cfs - 1))" +proof (induction x arbitrary: cfs rule: bcomp.induct, simp_all add: Let_def, + frule_tac [4] acomp_acomp, frule_tac [3] bcomp_bcomp, auto, force simp: bpred_def) + fix b\<^sub>1 b\<^sub>2 f i cfs s stk + assume A: "bcomp (b\<^sub>1, False, size (bcomp (b\<^sub>2, f, i)) + (if f then 0 else i)) @ + bcomp (b\<^sub>2, f, i) \ cfs\" + (is "bcomp (_, _, ?n + ?i) @ ?cb \ _\") + moreover assume B: "cfs ! 0 = (0, s, stk)" and + "\cb cfs. \cb = ?cb; bcomp (b\<^sub>1, False, ?n + ?i) \ cfs\\ \ + bpred (b\<^sub>1, False, ?n + ?i) (cfs ! 0) (cfs ! (length cfs - Suc 0))" + ultimately have "\k < length cfs. bpred (b\<^sub>1, False, ?n + ?i) + (off [] (cfs ! 0)) (off [] (cfs ! k))" + by (rule_tac execl_all_sub, auto simp: execl_all_def) + moreover assume C: "\ bval b\<^sub>1 s" + ultimately obtain k where D: "k < length cfs" and + E: "cfs ! k = (size (bcomp (b\<^sub>1, False, ?n + ?i)) + ?n + ?i, s, stk)" + using B by (auto simp: bpred_def) + assume "0 \ i" + thus "bpred (And b\<^sub>1 b\<^sub>2, f, i) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A and C and D and E by (insert execl_last, auto simp: + execl_all_def bpred_def Let_def) +next + fix b\<^sub>1 b\<^sub>2 f i cfs s stk k + assume A: "bcomp (b\<^sub>1, False, size (bcomp (b\<^sub>2, f, i)) + (if f then 0 else i)) @ + bcomp (b\<^sub>2, f, i) \ cfs\" + (is "?cb\<^sub>1 @ ?cb\<^sub>2 \ _\") + assume "k < length cfs" and "0 \ i" and "bval b\<^sub>1 s" and + "cfs ! k = (size ?cb\<^sub>1 + size ?cb\<^sub>2 + (if bval b\<^sub>2 s = f then i else 0), s, stk)" + thus "bpred (And b\<^sub>1 b\<^sub>2, f, i) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A by (insert execl_last, auto simp: execl_all_def bpred_def Let_def) +next + fix a\<^sub>1 a\<^sub>2 f i cfs s stk k + assume A: "acomp a\<^sub>1 @ acomp a\<^sub>2 @ + (if f then [JMPLESS i] else [JMPGE i]) \ cfs\" + (is "?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i \ _\") + assume B: "k < length cfs" and + C: "cfs ! k = (size ?ca\<^sub>1 + size ?ca\<^sub>2, s, aval a\<^sub>2 s # aval a\<^sub>1 s # stk)" + hence D: "cfs ! Suc k = + (size (?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i) + (if bval (Less a\<^sub>1 a\<^sub>2) s = f then i else 0), s, stk)" + using A by (insert execl_next [of "?ca\<^sub>1 @ ?ca\<^sub>2 @ ?i" cfs k], + simp add: execl_all_def, drule_tac impI, auto split: if_split_asm) + assume "0 \ i" + with A and B and C and D have "length cfs - 1 = Suc k" + by (rule_tac execl_last, auto simp: execl_all_def, simp split: if_split_asm) + thus "bpred (Less a\<^sub>1 a\<^sub>2, f, i) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using D by (simp add: bpred_def) +qed + +text \ +\null + +Next, lemmas @{text bcomp_ccomp}, @{text ccomp_ccomp} are derived to reproduce the execution of a +compiled boolean expression followed by a compiled command and of two consecutive compiled commands, +respectively (possibly generated by @{const ccomp}). Then, compiler correctness for loops and for +all commands is proven in lemmas @{text while_correct} and @{text ccomp_correct}, respectively by +induction over the length of the list of configurations and by structural induction over commands. + +\null +\ + +lemma bcomp_ccomp: + "\bcomp (b, f, i) @ ccomp c @ P \ cfs\; 0 \ i; + \cfs. ccomp c \ cfs\ \ cpred c (cfs ! 0) (cfs ! (length cfs - 1))\ \ + case cfs ! 0 of (pc, s, stk) \ pc = 0 \ (bval b s \ f \ + (\k < length cfs. case cfs ! k of (pc', s', stk') \ + pc' = size (bcomp (b, f, i) @ ccomp c) \ (c, s) \ s' \ stk' = stk))" +by (clarify, rule conjI, simp add: execl_all_def, rule impI, drule execl_all_sub2 + [where I = "\s. if bval b s = f then i else 0" and I' = "\s. 0" + and Q = "\s s'. s' = s" and Q' = "\s s'. (c, s) \ s'" and F = "\s stk. stk" + and F' = "\s stk. stk"], insert bcomp_correct [of "(b, f, i)"], + auto simp: bpred_def cpred_def) + +lemma ccomp_ccomp: + "\ccomp c\<^sub>1 @ ccomp c\<^sub>2 \ cfs\; + \cfs. ccomp c\<^sub>1 \ cfs\ \ cpred c\<^sub>1 (cfs ! 0) (cfs ! (length cfs - 1)); + \cfs. ccomp c\<^sub>2 \ cfs\ \ cpred c\<^sub>2 (cfs ! 0) (cfs ! (length cfs - 1))\ \ + case cfs ! 0 of (pc, s, stk) \ pc = 0 \ (\k < length cfs. \t. + case cfs ! k of (pc', s', stk') \ pc' = size (ccomp c\<^sub>1 @ ccomp c\<^sub>2) \ + (c\<^sub>1, s) \ t \ (c\<^sub>2, t) \ s' \ stk' = stk)" +by (subst (asm) append_Nil2 [symmetric], drule execl_all_sub2 [where I = "\s. 0" + and I' = "\s. 0" and Q = "\s s'. (c\<^sub>1, s) \ s'" and Q' = "\s s'. (c\<^sub>2, s) \ s'" + and F = "\s stk. stk" and F' = "\s stk. stk"], auto simp: cpred_def, force) + +lemma while_correct [simplified, intro]: + "\bcomp (b, False, size (ccomp c) + 1) @ ccomp c @ + [JMP (- (size (bcomp (b, False, size (ccomp c) + 1) @ ccomp c) + 1))] \ cfs\; + \cfs. ccomp c \ cfs\ \ cpred c (cfs ! 0) (cfs ! (length cfs - 1))\ \ + cpred (WHILE b DO c) (cfs ! 0) (cfs ! (length cfs - Suc 0))" + (is "\?cb @ ?cc @ [JMP (- ?n)] \ _\; \_. _ \ _\ \ ?Q cfs") +proof (induction cfs rule: length_induct, frule bcomp_ccomp, auto) + fix cfs s stk + assume A: "?cb @ ?cc @ [JMP (- size ?cb - size ?cc - 1)] \ cfs\" + hence "\k < length cfs. bpred (b, False, size (ccomp c) + 1) + (off [] (cfs ! 0)) (off [] (cfs ! k))" + by (rule_tac execl_all_sub, auto simp: execl_all_def) + moreover assume B: "\ bval b s" and "cfs ! 0 = (0, s, stk)" + ultimately obtain k where "k < length cfs" and "cfs ! k = (?n, s, stk)" + by (auto simp: bpred_def) + thus "cpred (WHILE b DO c) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A and B by (insert execl_last, auto simp: execl_all_def cpred_def Let_def) +next + fix cfs s s' stk k + assume A: "?cb @ ?cc @ [JMP (- size ?cb - size ?cc - 1)] \ cfs\" + (is "?P \ _\") + assume B: "k < length cfs" and "cfs ! k = (size ?cb + size ?cc, s', stk)" + moreover from this have C: "k \ length cfs - 1" + using A by (rule_tac notI, simp add: execl_all_def) + ultimately have D: "cfs ! Suc k = (0, s', stk)" + using A by (insert execl_next [of ?P cfs k], auto simp: execl_all_def) + moreover have E: "Suc k + (length (drop (Suc k) cfs) - 1) = length cfs - 1" + (is "Suc k + (length ?cfs - 1) = _") + using B and C by simp + ultimately have "?P \ ?cfs\" + using A and B and C by (auto simp: execl_all_def intro: execl_drop) + moreover assume "\cfs'. length cfs' < length cfs \ ?P \ cfs'\ \ ?Q cfs'" + hence "length ?cfs < length cfs \ ?P \ ?cfs\ \ ?Q ?cfs" .. + ultimately have "cpred (WHILE b DO c) (cfs ! Suc k) (cfs ! (length cfs - 1))" + using B and C and E by simp + moreover assume "bval b s" and "(c, s) \ s'" + ultimately show "cpred (WHILE b DO c) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using D by (auto simp: cpred_def) +qed + +lemma ccomp_correct: + "ccomp c \ cfs\ \ cpred c (cfs ! 0) (cfs ! (length cfs - 1))" +proof (induction c arbitrary: cfs, simp_all add: Let_def, frule_tac [4] bcomp_ccomp, + frule_tac [3] ccomp_ccomp, auto) + fix a x cfs + assume A: "acomp a @ [STORE x] \ cfs\" + hence "\k < length cfs. apred a (off [] (cfs ! 0)) (off [] (cfs ! k))" + by (rule_tac execl_all_sub, auto simp: execl_all_def) + moreover obtain s stk where B: "cfs ! 0 = (0, s, stk)" + using A by (cases "cfs ! 0", simp add: execl_all_def) + ultimately obtain k where C: "k < length cfs" and + D: "cfs ! k = (size (acomp a), s, aval a s # stk)" + by (auto simp: apred_def) + hence "cfs ! Suc k = (size (acomp a) + 1, s(x := aval a s), stk)" + using A by (insert execl_next [of "acomp a @ [STORE x]" cfs k], + simp add: execl_all_def, drule_tac impI, auto) + thus "cpred (x ::= a) (cfs ! 0) (cfs ! (length cfs - Suc 0))" + using A and B and C and D by (insert execl_last [of "acomp a @ [STORE x]" + cfs "Suc k"], simp add: execl_all_def cpred_def, drule_tac impI, auto) +next + fix c\<^sub>1 c\<^sub>2 cfs s s' t stk k + assume "ccomp c\<^sub>1 @ ccomp c\<^sub>2 \ cfs\" and "k < length cfs" and + "cfs ! k = (size (ccomp c\<^sub>1) + size (ccomp c\<^sub>2), s', stk)" + moreover assume "(c\<^sub>1, s) \ t" and "(c\<^sub>2, t) \ s'" + ultimately show "cpred (c\<^sub>1;; c\<^sub>2) (0, s, stk) (cfs ! (length cfs - Suc 0))" + by (insert execl_last, auto simp: execl_all_def cpred_def) +next + fix b c\<^sub>1 c\<^sub>2 cfs s stk + assume A: "bcomp (b, False, size (ccomp c\<^sub>1) + 1) @ ccomp c\<^sub>1 @ + JMP (size (ccomp c\<^sub>2)) # ccomp c\<^sub>2 \ cfs\" + (is "bcomp ?x @ ?cc\<^sub>1 @ _ # ?cc\<^sub>2 \ _\") + let ?P = "bcomp ?x @ ?cc\<^sub>1 @ [JMP (size ?cc\<^sub>2)]" + have "\k < length cfs. bpred ?x (off [] (cfs ! 0)) (off [] (cfs ! k))" + using A by (rule_tac execl_all_sub, auto simp: execl_all_def) + moreover assume B: "\ bval b s" and "cfs ! 0 = (0, s, stk)" + ultimately obtain k where C: "k < length cfs" and D: "cfs ! k = (size ?P, s, stk)" + by (force simp: bpred_def) + assume "\cfs. ?cc\<^sub>2 \ cfs\ \ cpred c\<^sub>2 (cfs ! 0) (cfs ! (length cfs - Suc 0))" + hence "\k' < length cfs. cpred c\<^sub>2 (off ?P (cfs ! k)) (off ?P (cfs ! k'))" + using A and C and D by (rule_tac execl_all_sub [where P'' = "[]"], auto) + then obtain k' where "k' < length cfs" and "case cfs ! k' of (pc', s', stk') \ + pc' = size (?P @ ?cc\<^sub>2) \ (c\<^sub>2, s) \ s' \ stk' = stk" + using D by (fastforce simp: cpred_def) + thus "cpred (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A and B by (insert execl_last, auto simp: execl_all_def cpred_def Let_def) +next + fix b c\<^sub>1 c\<^sub>2 cfs s s' stk k + assume A: "bcomp (b, False, size (ccomp c\<^sub>1) + 1) @ ccomp c\<^sub>1 @ + JMP (size (ccomp c\<^sub>2)) # ccomp c\<^sub>2 \ cfs\" + (is "?cb @ ?cc\<^sub>1 @ ?i # ?cc\<^sub>2 \ _\") + assume B: "k < length cfs" and C: "cfs ! k = (size ?cb + size ?cc\<^sub>1, s', stk)" + hence D: "cfs ! Suc k = (size (?cb @ ?cc\<^sub>1 @ ?i # ?cc\<^sub>2), s', stk)" + (is "_ = (size ?P, _, _)") + using A by (insert execl_next [of ?P cfs k], simp add: execl_all_def, + drule_tac impI, auto) + assume "bval b s" and "(c\<^sub>1, s) \ s'" + thus "cpred (IF b THEN c\<^sub>1 ELSE c\<^sub>2) (0, s, stk) (cfs ! (length cfs - Suc 0))" + using A and B and C and D by (insert execl_last [of ?P cfs "Suc k"], + simp add: execl_all_def cpred_def Let_def, drule_tac impI, auto) +next + fix c\<^sub>1 c\<^sub>2 cfs + assume A: "JMPND (size (ccomp c\<^sub>1) + 1) # ccomp c\<^sub>1 @ + JMP (size (ccomp c\<^sub>2)) # ccomp c\<^sub>2 \ cfs\" + (is "JMPND (?n\<^sub>1 + 1) # ?cc\<^sub>1 @ JMP ?n\<^sub>2 # ?cc\<^sub>2 \ _\") + let ?P = "JMPND (?n\<^sub>1 + 1) # ?cc\<^sub>1 @ [JMP ?n\<^sub>2]" + assume + B: "\cfs. ?cc\<^sub>1 \ cfs\ \ cpred c\<^sub>1 (cfs ! 0) (cfs ! (length cfs - Suc 0))" and + C: "\cfs. ?cc\<^sub>2 \ cfs\ \ cpred c\<^sub>2 (cfs ! 0) (cfs ! (length cfs - Suc 0))" + from A obtain s stk where D: "cfs ! 0 = (0, s, stk)" + by (cases "cfs ! 0", simp add: execl_all_def) + with A have "cfs ! 1 = (1, s, stk) \ cfs ! 1 = (?n\<^sub>1 + 2, s, stk)" + by (insert execl_next [of "?P @ ?cc\<^sub>2" cfs 0], simp add: execl_all_def, + drule_tac impI, auto) + moreover { + assume E: "cfs ! 1 = (1, s, stk)" + hence "\k < length cfs. cpred c\<^sub>1 (off [hd ?P] (cfs ! 1)) (off [hd ?P] (cfs ! k))" + using A and B by (rule_tac execl_all_sub, auto simp: execl_all_def) + then obtain k where "k < length cfs" and "case cfs ! k of (pc', s', stk') \ + pc' = ?n\<^sub>1 + 1 \ (c\<^sub>1, s) \ s' \ stk' = stk" + using E by (fastforce simp: cpred_def) + moreover from this have "case cfs ! Suc k of (pc', s', stk') \ + pc' = ?n\<^sub>1 + ?n\<^sub>2 + 2 \ (c\<^sub>1, s) \ s' \ stk' = stk" + using A by (insert execl_next [of "?P @ ?cc\<^sub>2" cfs k], simp add: execl_all_def, + drule_tac impI, auto) + ultimately have "cpred (c\<^sub>1 OR c\<^sub>2) (cfs ! 0) (cfs ! (length cfs - Suc 0))" + using A and D by (insert execl_last [of "?P @ ?cc\<^sub>2" cfs "Suc k"], + simp add: execl_all_def cpred_def split_def Let_def, drule_tac impI, auto) + } + moreover { + assume E: "cfs ! 1 = (?n\<^sub>1 + 2, s, stk)" + with A and C have "\k2 (off ?P (cfs!1)) (off ?P (cfs!k))" + by (rule_tac execl_all_sub [where P'' = "[]"], auto simp: execl_all_def) + then obtain k where "k < length cfs" and "case cfs ! k of (pc', s', stk') \ + pc' = ?n\<^sub>1 + ?n\<^sub>2 + 2 \ (c\<^sub>2, s) \ s' \ stk' = stk" + using E by (fastforce simp: cpred_def) + with A and D have "cpred (c\<^sub>1 OR c\<^sub>2) (cfs ! 0) (cfs ! (length cfs - Suc 0))" + by (insert execl_last, auto simp: execl_all_def cpred_def Let_def) + } + ultimately show "cpred (c\<^sub>1 OR c\<^sub>2) (cfs ! 0) (cfs ! (length cfs - Suc 0))" .. +qed + +text \ +\null + +Finally, the main compiler correctness theorem, expressed using predicate @{const exec}, is proven. +First, @{prop "P \ cf \* cf'"} is shown to imply the existence of a nonempty list of configurations +@{text cfs} such that @{prop "P \ cfs"}, whose initial and final configurations match @{text cf} +and @{text cf'}, respectively. Then, the main theorem is derived as a straightforward consequence of +this lemma and of the previous lemma @{thm [source] ccomp_correct}. + +\null +\ + +lemma exec_execl [dest!]: + "P \ cf \* cf' \ \cfs. P \ cfs \ cfs \ [] \ hd cfs = cf \ last cfs = cf'" +by (erule star.induct, force, erule exE, rule list.exhaust, blast, + simp del: last.simps, rule exI, subst execl.simps(1), simp) + +theorem ccomp_exec: + "ccomp c \ (0, s, stk) \* (size (ccomp c), s', stk') \ (c, s) \ s' \ stk' = stk" +by (insert ccomp_correct, force simp: hd_conv_nth last_conv_nth execl_all_def cpred_def) + +end diff --git a/thys/IMP_Compiler/ROOT b/thys/IMP_Compiler/ROOT new file mode 100755 --- /dev/null +++ b/thys/IMP_Compiler/ROOT @@ -0,0 +1,12 @@ +chapter AFP + +session IMP_Compiler (AFP) = HOL + + options [timeout=300] + sessions + "HOL-IMP" + theories + Compiler + Compiler2 + document_files + "root.bib" + "root.tex" diff --git a/thys/IMP_Compiler/document/root.bib b/thys/IMP_Compiler/document/root.bib new file mode 100755 --- /dev/null +++ b/thys/IMP_Compiler/document/root.bib @@ -0,0 +1,48 @@ +@book{ + Nipkow14, + author = {Tobias Nipkow and Gerwin Klein}, + title = {Concrete Semantics with Isabelle/HOL}, + month = mar, + year = 2021, + publisher = {Springer-Verlag}, + note = {(Current version: \url{http://www.concrete-semantics.org/concrete-semantics.pdf})} +} + +@misc{ + Klein21, + author = {Gerwin Klein}, + title = {{Theory HOL-IMP.Compiler2 (included in the Isabelle2021 distribution)}}, + note = {\url{https://isabelle.in.tum.de/website-Isabelle2021/dist/library/HOL/HOL-IMP/Compiler2.html}} +} + +@manual{ + Paulson21, + author = {Tobias Nipkow and Lawrence C. Paulson and Markus Wenzel}, + title = {Isabelle/HOL -- A Proof Assistant for Higher-Order Logic}, + month = feb, + year = 2021, + note = {\url{https://isabelle.in.tum.de/website-Isabelle2021/dist/Isabelle2021/doc/tutorial.pdf}} +} + +@manual{ + Nipkow21, + author = {Tobias Nipkow}, + title = {Programming and Proving in Isabelle/HOL}, + month = feb, + year = 2021, + note = {\url{https://isabelle.in.tum.de/website-Isabelle2021/dist/Isabelle2021/doc/prog-prove.pdf}} +} + +@manual{ + Krauss21, + author = {Alexander Krauss}, + title = {Defining Recursive Functions in Isabelle/HOL}, + note = {\url{https://isabelle.in.tum.de/website-Isabelle2021/dist/Isabelle2021/doc/functions.pdf}} +} + +@manual{ + Nipkow11, + author = {Tobias Nipkow}, + title = {A Tutorial Introduction to Structured Isar Proofs}, + note = {\url{https://isabelle.in.tum.de/website-Isabelle2011/dist/Isabelle2011/doc/isar-overview.pdf}} +} diff --git a/thys/IMP_Compiler/document/root.tex b/thys/IMP_Compiler/document/root.tex new file mode 100755 --- /dev/null +++ b/thys/IMP_Compiler/document/root.tex @@ -0,0 +1,68 @@ +\documentclass[11pt,a4paper,fleqn]{article} +\usepackage{isabelle,isabellesym} +\renewcommand{\isastyletxt}{\isastyletext} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + + +\begin{document} + +\title{A Shorter Compiler Correctness Proof\\for Language IMP} +\author{Pasquale Noce\\Software Engineer at HID Global, Italy\\pasquale dot noce dot lavoro at gmail dot com\\pasquale dot noce at hidglobal dot com} +\maketitle + +\begin{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. +\end{abstract} + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% bibliography +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/MiniSail/BTVSubst.thy b/thys/MiniSail/BTVSubst.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/BTVSubst.thy @@ -0,0 +1,1550 @@ +(*<*) +theory BTVSubst + imports Syntax +begin + (*>*) +chapter \Basic Type Variable Substitution\ + +section \Class\ + +class has_subst_b = fs + + fixes subst_b :: "'a::fs \ bv \ b \ 'a::fs" ("_[_::=_]\<^sub>b" [1000,50,50] 1000) + +assumes fresh_subst_if: "j \ (t[i::=x]\<^sub>b ) \ (atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i))" + and forget_subst[simp]: "atom a \ tm \ tm[a::=x]\<^sub>b = tm" + and subst_id[simp]: "tm[a::=(B_var a)]\<^sub>b = tm" + and eqvt[simp,eqvt]: "(p::perm) \ (subst_b t1 x1 v ) = (subst_b (p \t1) (p \x1) (p \v) )" + and flip_subst[simp]: "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + and flip_subst_subst[simp]: "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" +begin + +lemmas flip_subst_b = flip_subst_subst + +lemma subst_b_simple_commute: + fixes x::bv + assumes "atom x \ c" + shows "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = c[z::=b]\<^sub>b" +proof - + have "(c[z::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (( x \ z) \ c)[x::=b]\<^sub>b" using flip_subst assms by simp + thus ?thesis using flip_subst_subst assms by simp +qed + +lemma subst_b_flip_eq_one: + fixes z1::bv and z2::bv and x1::bv and x2::bv + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + and "atom x1 \ (z1,z2,c1,c2)" + shows "(c1[z1::=B_var x1]\<^sub>b) = (c2[z2::=B_var x1]\<^sub>b)" +proof - + have "(c1[z1::=B_var x1]\<^sub>b) = (x1 \ z1) \ c1" using assms flip_subst by auto + moreover have "(c2[z2::=B_var x1]\<^sub>b) = (x1 \ z2) \ c2" using assms flip_subst by auto + ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms + by (metis Abs1_eq_iff_fresh(3) flip_commute) +qed + +lemma subst_b_flip_eq_two: + fixes z1::bv and z2::bv and x1::bv and x2::bv + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + shows "(c1[z1::=b]\<^sub>b) = (c2[z2::=b]\<^sub>b)" +proof - + obtain x::bv where *:"atom x \ (z1,z2,c1,c2)" using obtain_fresh by metis + hence "(c1[z1::=B_var x]\<^sub>b) = (c2[z2::=B_var x]\<^sub>b)" using subst_b_flip_eq_one[OF assms, of x] by metis + hence "(c1[z1::=B_var x]\<^sub>b)[x::=b]\<^sub>b = (c2[z2::=B_var x]\<^sub>b)[x::=b]\<^sub>b" by auto + thus ?thesis using subst_b_simple_commute * fresh_prod4 by metis +qed + +lemma subst_b_fresh_x: + fixes tm::"'a::fs" and x::x + shows "atom x \ tm = atom x \ tm[bv::=b']\<^sub>b" + using fresh_subst_if[of "atom x" tm bv b'] using x_fresh_b by auto + +lemma subst_b_x_flip[simp]: + fixes x'::x and x::x and bv::bv + shows "((x' \ x) \ tm)[bv::=b']\<^sub>b = (x' \ x) \ (tm[bv::=b']\<^sub>b)" +proof - + have "(x' \ x) \ bv = bv" using pure_supp flip_fresh_fresh by force + moreover have "(x' \ x) \ b' = b'" using x_fresh_b flip_fresh_fresh by auto + ultimately show ?thesis using eqvt by simp +qed + +end + +section \Base Type\ + +nominal_function subst_bb :: "b \ bv \ b \ b" where + "subst_bb (B_var bv2) bv1 b = (if bv1 = bv2 then b else (B_var bv2))" +| "subst_bb B_int bv1 b = B_int" +| "subst_bb B_bool bv1 b = B_bool" +| "subst_bb (B_id s) bv1 b = B_id s " +| "subst_bb (B_pair b1 b2) bv1 b = B_pair (subst_bb b1 bv1 b) (subst_bb b2 bv1 b)" +| "subst_bb B_unit bv1 b = B_unit " +| "subst_bb B_bitvec bv1 b = B_bitvec " +| "subst_bb (B_app s b2) bv1 b = B_app s (subst_bb b2 bv1 b)" + apply (simp add: eqvt_def subst_bb_graph_aux_def ) + apply (simp add: eqvt_def subst_bb_graph_aux_def ) + by (auto,meson b.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_bb_abbrev :: "b \ bv \ b \ b" ("_[_::=_]\<^sub>b\<^sub>b" [1000,50,50] 1000) + where + "b[bv::=b']\<^sub>b\<^sub>b \ subst_bb b bv b' " + +instantiation b :: has_subst_b +begin +definition "subst_b = subst_bb" + +instance proof + fix j::atom and i::bv and x::b and t::b + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof (induct t rule: b.induct) + case (B_id x) + then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto + next + case (B_var x) + then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto + next + case (B_app x1 x2) + then show ?case using subst_bb.simps fresh_def pure_fresh subst_b_b_def by auto + qed(auto simp add: subst_bb.simps fresh_def pure_fresh subst_b_b_def)+ + + fix a::bv and tm::b and x::b + show "atom a \ tm \ tm[a::=x]\<^sub>b = tm" + by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) + + fix a::bv and tm::b + show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def + by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) + + fix p::perm and x1::bv and v::b and t1::b + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + by (induct tm rule: b.induct, auto simp add: fresh_at_base subst_bb.simps subst_b_b_def) + + fix bv::bv and c::b and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+) + + fix bv::bv and c::b and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + by (induct c rule: b.induct, (auto simp add: fresh_at_base subst_bb.simps subst_b_b_def permute_pure pure_supp )+) +qed +end + +lemma subst_bb_inject: + assumes "b1 = b2[bv::=b]\<^sub>b\<^sub>b" and "b2 \ B_var bv" + shows + "b1 = B_int \ b2 = B_int" and + "b1 = B_bool \ b2 = B_bool" and + "b1 = B_unit \ b2 = B_unit" and + "b1 = B_bitvec \ b2 = B_bitvec" and + "b1 = B_pair b11 b12 \ (\b11' b12' . b11 = b11'[bv::=b]\<^sub>b\<^sub>b \ b12 = b12'[bv::=b]\<^sub>b\<^sub>b \ b2 = B_pair b11' b12')" and + "b1 = B_var bv' \ b2 = B_var bv'" and + "b1 = B_id tyid \ b2 = B_id tyid" and + "b1 = B_app tyid b11 \ (\b11'. b11= b11'[bv::=b]\<^sub>b\<^sub>b \ b2 = B_app tyid b11')" + using assms by(nominal_induct b2 rule:b.strong_induct,auto+) + +lemma flip_b_subst4: + fixes b1::b and bv1::bv and c::bv and b::b + assumes "atom c \ (b1,bv1)" + shows "b1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ b1)[c ::= b]\<^sub>b\<^sub>b" + using assms proof(nominal_induct b1 rule: b.strong_induct) + case B_int + then show ?case using subst_bb.simps b.perm_simps by auto +next + case B_bool + then show ?case using subst_bb.simps b.perm_simps by auto +next + case (B_id x) + hence "atom bv1 \ x \ atom c \ x" using fresh_def pure_supp by auto + hence "((bv1 \ c) \ B_id x) = B_id x" using fresh_Pair b.fresh(3) flip_fresh_fresh b.perm_simps fresh_def pure_supp by metis + then show ?case using subst_bb.simps by simp +next + case (B_pair x1 x2) + hence "x1[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x1)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair by metis + moreover have "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps(4) b.fresh(4) fresh_Pair B_pair by metis + ultimately show ?case using subst_bb.simps(5) b.perm_simps(4) b.fresh(4) fresh_Pair by auto +next + case B_unit + then show ?case using subst_bb.simps b.perm_simps by auto +next + case B_bitvec + then show ?case using subst_bb.simps b.perm_simps by auto +next + case (B_var x) + then show ?case proof(cases "x=bv1") + case True + then show ?thesis using B_var subst_bb.simps b.perm_simps by simp + next + case False + moreover have "x\c" using B_var b.fresh fresh_def supp_at_base fresh_Pair by fastforce + ultimately show ?thesis using B_var subst_bb.simps(1) b.perm_simps(7) by simp + qed +next + case (B_app x1 x2) + hence "x2[bv1::=b]\<^sub>b\<^sub>b = ((bv1 \ c) \ x2)[c::=b]\<^sub>b\<^sub>b" using b.perm_simps b.fresh fresh_Pair by metis + thus ?case using subst_bb.simps b.perm_simps b.fresh fresh_Pair B_app + by (simp add: permute_pure) +qed + +lemma subst_bb_flip_sym: + fixes b1::b and b2::b + assumes "atom c \ b" and "atom c \ (bv1,bv2, b1, b2)" and "(bv1 \ c) \ b1 = (bv2 \ c) \ b2" + shows "b1[bv1::=b]\<^sub>b\<^sub>b = b2[bv2::=b]\<^sub>b\<^sub>b" + using assms flip_b_subst4[of c b1 bv1 b] flip_b_subst4[of c b2 bv2 b] fresh_prod4 fresh_Pair by simp + +section \Value\ + +nominal_function subst_vb :: "v \ bv \ b \ v" where + "subst_vb (V_lit l) x v = V_lit l" +| "subst_vb (V_var y) x v = V_var y" +| "subst_vb (V_cons tyid c v') x v = V_cons tyid c (subst_vb v' x v)" +| "subst_vb (V_consp tyid c b v') x v = V_consp tyid c (subst_bb b x v) (subst_vb v' x v)" +| "subst_vb (V_pair v1 v2) x v = V_pair (subst_vb v1 x v ) (subst_vb v2 x v )" + apply (simp add: eqvt_def subst_vb_graph_aux_def) + apply auto + using v.strong_exhaust by meson +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_vb_abbrev :: "v \ bv \ b \ v" ("_[_::=_]\<^sub>v\<^sub>b" [1000,50,50] 500) + where + "e[bv::=b]\<^sub>v\<^sub>b \ subst_vb e bv b" + +instantiation v :: has_subst_b +begin +definition "subst_b = subst_vb" + +instance proof + fix j::atom and i::bv and x::b and t::v + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof (induct t rule: v.induct) + case (V_lit l) + have "j \ subst_b (V_lit l) i x = j \ (V_lit l)" using subst_vb.simps fresh_def pure_fresh + subst_b_v_def v.supp v.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by auto + also have "... = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis + moreover have "(atom i \ (V_lit l) \ j \ (V_lit l) \ j \ x \ (j \ (V_lit l) \ j = atom i)) = True" using fresh_at_base v.fresh l.fresh supp_l_empty fresh_def by metis + ultimately show ?case by simp + next + case (V_var y) + then show ?case using subst_b_v_def subst_vb.simps pure_fresh by force + next + case (V_pair x1a x2a) + then show ?case using subst_b_v_def subst_vb.simps by auto + next + case (V_cons x1a x2a x3) + then show ?case using V_cons subst_b_v_def subst_vb.simps pure_fresh by force + next + case (V_consp x1a x2a x3 x4) + then show ?case using subst_b_v_def subst_vb.simps pure_fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by fastforce + qed + + fix a::bv and tm::v and x::b + show "atom a \ tm \ subst_b tm a x = tm" + apply(induct tm rule: v.induct) + apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.forget_subst by fastforce + + fix a::bv and tm::v + show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def + apply (induct tm rule: v.induct) + apply(auto simp add: fresh_at_base subst_vb.simps subst_b_v_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.subst_id by metis + + fix p::perm and x1::bv and v::b and t1::v + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + apply(induct tm rule: v.induct) + apply(auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) + using has_subst_b_class.eqvt subst_b_b_def e.fresh + using has_subst_b_class.eqvt + by (simp add: subst_b_v_def)+ + + fix bv::bv and c::v and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+) + apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2) + using fresh_at_base flip_fresh_fresh[of bv x z] + apply (simp add: flip_fresh_fresh) + using subst_b_b_def by argo + + fix bv::bv and c::v and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + apply (induct c rule: v.induct, (auto simp add: fresh_at_base subst_vb.simps subst_b_v_def permute_pure pure_supp )+) + apply (metis flip_fresh_fresh flip_l_eq permute_flip_cancel2) + using fresh_at_base flip_fresh_fresh[of bv x z] + apply (simp add: flip_fresh_fresh) + using subst_b_b_def flip_subst_subst by fastforce + +qed +end + +section \Constraints Expressions\ + +nominal_function subst_ceb :: "ce \ bv \ b \ ce" where + "subst_ceb ( (CE_val v') ) bv b = ( CE_val (subst_vb v' bv b) )" +| "subst_ceb ( (CE_op opp v1 v2) ) bv b = ( (CE_op opp (subst_ceb v1 bv b)(subst_ceb v2 bv b)) )" +| "subst_ceb ( (CE_fst v')) bv b = CE_fst (subst_ceb v' bv b)" +| "subst_ceb ( (CE_snd v')) bv b = CE_snd (subst_ceb v' bv b)" +| "subst_ceb ( (CE_len v')) bv b = CE_len (subst_ceb v' bv b)" +| "subst_ceb ( CE_concat v1 v2) bv b = CE_concat (subst_ceb v1 bv b) (subst_ceb v2 bv b)" + apply (simp add: eqvt_def subst_ceb_graph_aux_def) + apply auto + by (meson ce.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_ceb_abbrev :: "ce \ bv \ b \ ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>b" [1000,50,50] 500) + where + "ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b \ subst_ceb ce bv b" + +instantiation ce :: has_subst_b +begin +definition "subst_b = subst_ceb" + +instance proof + fix j::atom and i::bv and x::b and t::ce + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof (induct t rule: ce.induct) + case (CE_val v) + then show ?case using subst_ceb.simps fresh_def pure_fresh subst_b_ce_def ce.supp v.supp ce.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def + by metis + next + case (CE_op opp v1 v2) + + have "(j \ v1[i::=x]\<^sub>c\<^sub>e\<^sub>b \ j \ v2[i::=x]\<^sub>c\<^sub>e\<^sub>b) = ((atom i \ v1 \ atom i \ v2) \ j \ v1 \ j \ v2 \ j \ x \ (j \ v1 \ j \ v2 \ j = atom i))" + using has_subst_b_class.fresh_subst_if subst_b_v_def + using CE_op.hyps(1) CE_op.hyps(2) subst_b_ce_def by auto + + thus ?case unfolding subst_ceb.simps subst_b_ce_def ce.fresh + using fresh_def pure_fresh opp.fresh subst_b_v_def opp.exhaust fresh_e_opp_all + by (metis (full_types)) + next + case (CE_concat x1a x2) + then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by force + next + case (CE_fst x) + then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis + + next + case (CE_snd x) + then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis + next + case (CE_len x) + then show ?case using subst_ceb.simps subst_b_ce_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def ce.fresh by metis + qed + + fix a::bv and tm::ce and x::b + show "atom a \ tm \ subst_b tm a x = tm" + apply(induct tm rule: ce.induct) + apply( auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.forget_subst subst_b_v_def apply metis+ + done + + fix a::bv and tm::ce + show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def + apply (induct tm rule: ce.induct) + apply(auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.subst_id subst_b_v_def apply metis+ + done + + fix p::perm and x1::bv and v::b and t1::ce + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + apply(induct tm rule: ce.induct) + apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) + using has_subst_b_class.eqvt subst_b_b_def ce.fresh + using has_subst_b_class.eqvt + by (simp add: subst_b_ce_def)+ + + fix bv::bv and c::ce and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (induct c rule: ce.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+) + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def apply metis + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def + by (simp add: flip_fresh_fresh fresh_opp_all) + + fix bv::bv and c::ce and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + proof (induct c rule: ce.induct) + case (CE_val x) + then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce + next + case (CE_op x1a x2 x3) + then show ?case unfolding subst_ceb.simps subst_b_ce_def ce.perm_simps using flip_subst_subst subst_b_v_def opp.perm_simps opp.strong_exhaust + by (metis (full_types) ce.fresh(2)) + next + case (CE_concat x1a x2) + then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce + next + case (CE_fst x) + then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce + next + case (CE_snd x) + then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce + next + case (CE_len x) + then show ?case using flip_subst_subst subst_b_v_def subst_ceb.simps using subst_b_ce_def by fastforce + qed +qed +end + +section \Constraints\ + +nominal_function subst_cb :: "c \ bv \ b \ c" where + "subst_cb (C_true) x v = C_true" +| "subst_cb (C_false) x v = C_false" +| "subst_cb (C_conj c1 c2) x v = C_conj (subst_cb c1 x v ) (subst_cb c2 x v )" +| "subst_cb (C_disj c1 c2) x v = C_disj (subst_cb c1 x v ) (subst_cb c2 x v )" +| "subst_cb (C_imp c1 c2) x v = C_imp (subst_cb c1 x v ) (subst_cb c2 x v )" +| "subst_cb (C_eq e1 e2) x v = C_eq (subst_ceb e1 x v ) (subst_ceb e2 x v )" +| "subst_cb (C_not c) x v = C_not (subst_cb c x v )" + apply (simp add: eqvt_def subst_cb_graph_aux_def) + apply auto + using c.strong_exhaust apply metis + done +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_cb_abbrev :: "c \ bv \ b \ c" ("_[_::=_]\<^sub>c\<^sub>b" [1000,50,50] 500) + where + "c[bv::=b]\<^sub>c\<^sub>b \ subst_cb c bv b" + +instantiation c :: has_subst_b +begin +definition "subst_b = subst_cb" + +instance proof + fix j::atom and i::bv and x::b and t::c + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + by (induct t rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, + (metis has_subst_b_class.fresh_subst_if subst_b_ce_def c.fresh)+ + ) + + fix a::bv and tm::c and x::b + show "atom a \ tm \ subst_b tm a x = tm" + by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, + (metis has_subst_b_class.forget_subst subst_b_ce_def)+) + + fix a::bv and tm::c + show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_c_def + by(induct tm rule: c.induct, unfold subst_cb.simps subst_b_c_def c.fresh, + (metis has_subst_b_class.subst_id subst_b_ce_def)+) + + fix p::perm and x1::bv and v::b and t1::c + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + apply(induct tm rule: c.induct,unfold subst_cb.simps subst_b_c_def c.fresh) + by( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) + + fix bv::bv and c::c and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+) + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def apply metis + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def + apply (metis opp.perm_simps(2) opp.strong_exhaust)+ + done + + fix bv::bv and c::c and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + apply (induct c rule: c.induct, (auto simp add: fresh_at_base subst_cb.simps subst_b_c_def permute_pure pure_supp )+) + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def + using flip_subst_subst apply fastforce + using flip_fresh_fresh flip_l_eq permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_ce_def + opp.perm_simps(2) opp.strong_exhaust + proof - + fix x1a :: ce and x2 :: ce + assume a1: "atom bv \ x2" + then have "((bv \ z) \ x2)[bv::=v]\<^sub>b = x2[z::=v]\<^sub>b" + by (metis flip_subst_subst) (* 0.0 ms *) + then show "x2[z::=B_var bv]\<^sub>b[bv::=v]\<^sub>c\<^sub>e\<^sub>b = x2[z::=v]\<^sub>c\<^sub>e\<^sub>b" + using a1 by (simp add: subst_b_ce_def) (* 0.0 ms *) + qed + +qed +end + +section \Types\ + +nominal_function subst_tb :: "\ \ bv \ b \ \" where + "subst_tb (\ z : b2 | c \) bv1 b1 = \ z : b2[bv1::=b1]\<^sub>b\<^sub>b | c[bv1::=b1]\<^sub>c\<^sub>b \" +proof(goal_cases) + case 1 + then show ?case using eqvt_def subst_tb_graph_aux_def by force +next + case (2 x y) + then show ?case by auto +next + case (3 P x) + then show ?case using eqvt_def subst_tb_graph_aux_def \.strong_exhaust + by (metis b_of.cases prod_cases3) +next + case (4 z' b2' c' bv1' b1' z b2 c bv1 b1) + show ?case unfolding \.eq_iff proof + have *:"[[atom z']]lst. c' = [[atom z]]lst. c" using \.eq_iff 4 by auto + show "[[atom z']]lst. c'[bv1'::=b1']\<^sub>c\<^sub>b = [[atom z]]lst. c[bv1::=b1]\<^sub>c\<^sub>b" + proof(subst Abs1_eq_iff_all(3),rule,rule,rule) + fix ca::x + assume "atom ca \ z" and 1:"atom ca \ (z', z, c'[bv1'::=b1']\<^sub>c\<^sub>b, c[bv1::=b1]\<^sub>c\<^sub>b)" + hence 2:"atom ca \ (c',c)" using fresh_subst_if subst_b_c_def fresh_Pair fresh_prod4 fresh_at_base subst_b_fresh_x by metis + hence "(z' \ ca) \ c' = (z \ ca) \ c" using 1 2 * Abs1_eq_iff_all(3) by auto + hence "((z' \ ca) \ c')[bv1'::=b1']\<^sub>c\<^sub>b = ((z \ ca) \ c)[bv1'::=b1']\<^sub>c\<^sub>b" by auto + hence "(z' \ ca) \ c'[(z' \ ca) \ bv1'::=(z' \ ca) \ b1']\<^sub>c\<^sub>b = (z \ ca) \ c[(z \ ca) \ bv1'::=(z \ ca) \ b1']\<^sub>c\<^sub>b" by auto + thus "(z' \ ca) \ c'[bv1'::=b1']\<^sub>c\<^sub>b = (z \ ca) \ c[bv1::=b1]\<^sub>c\<^sub>b" using 4 flip_x_b_cancel by simp + qed + show "b2'[bv1'::=b1']\<^sub>b\<^sub>b = b2[bv1::=b1]\<^sub>b\<^sub>b" using 4 by simp + qed +qed + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_tb_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) + where + "t[bv::=b']\<^sub>\\<^sub>b \ subst_tb t bv b' " + +instantiation \ :: has_subst_b +begin +definition "subst_b = subst_tb" + +instance proof + fix j::atom and i::bv and x::b and t::\ + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof (nominal_induct t avoiding: i x j rule: \.strong_induct) + case (T_refined_type z b c) + then show ?case + unfolding subst_b_\_def subst_tb.simps \.fresh + using fresh_subst_if[of j b i x ] subst_b_b_def subst_b_c_def + by (metis has_subst_b_class.fresh_subst_if list.distinct(1) list.set_cases not_self_fresh set_ConsD) + qed + + fix a::bv and tm::\ and x::b + show "atom a \ tm \ subst_b tm a x = tm" + proof (nominal_induct tm avoiding: a x rule: \.strong_induct) + case (T_refined_type xx bb cc ) + moreover hence "atom a \ bb \ atom a \ cc" using \.fresh by auto + ultimately show ?case + unfolding subst_b_\_def subst_tb.simps + using forget_subst subst_b_b_def subst_b_c_def forget_subst \.fresh by metis + qed + + fix a::bv and tm::\ + show "subst_b tm a (B_var a) = tm" + proof (nominal_induct tm rule: \.strong_induct) + case (T_refined_type xx bb cc) + thus ?case + unfolding subst_b_\_def subst_tb.simps + using subst_id subst_b_b_def subst_b_c_def by metis + qed + + fix p::perm and x1::bv and v::b and t1::\ + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v) " + by (induct tm rule: \.induct, auto simp add: fresh_at_base subst_tb.simps subst_b_\_def subst_bb.simps subst_b_b_def) + + fix bv::bv and c::\ and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (induct c rule: \.induct, (auto simp add: fresh_at_base subst_ceb.simps subst_b_ce_def permute_pure pure_supp )+) + using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_c_def subst_b_b_def + by (simp add: flip_fresh_fresh subst_b_\_def) + + fix bv::bv and c::\ and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + proof (induct c rule: \.induct) + case (T_refined_type x1a x2a x3a) + hence "atom bv \ x2a \ atom bv \ x3a \ atom bv \ x1a" using fresh_at_base \.fresh by simp + then show ?case + unfolding subst_tb.simps subst_b_\_def \.perm_simps + using fresh_at_base flip_fresh_fresh[of bv x1a z] flip_subst_subst subst_b_b_def subst_b_c_def T_refined_type + proof - + have "atom z \ x1a" + by (metis b.fresh(7) fresh_at_base(2) x_fresh_b) (* 0.0 ms *) + then show "\ (bv \ z) \ x1a : ((bv \ z) \ x2a)[bv::=v]\<^sub>b\<^sub>b | ((bv \ z) \ x3a)[bv::=v]\<^sub>c\<^sub>b \ = \ x1a : x2a[z::=v]\<^sub>b\<^sub>b | x3a[z::=v]\<^sub>c\<^sub>b \" + by (metis \\atom bv \ x1a; atom z \ x1a\ \ (bv \ z) \ x1a = x1a\ \atom bv \ x2a \ atom bv \ x3a \ atom bv \ x1a\ flip_subst_subst subst_b_b_def subst_b_c_def) (* 78 ms *) + qed + qed + +qed +end + +lemma subst_bb_commute [simp]: + "atom j \ A \ (subst_bb (subst_bb A i t ) j u ) = subst_bb A i (subst_bb t j u) " + by (nominal_induct A avoiding: i j t u rule: b.strong_induct) (auto simp: fresh_at_base) + +lemma subst_vb_commute [simp]: + "atom j \ A \ (subst_vb (subst_vb A i t )) j u = subst_vb A i (subst_bb t j u ) " + by (nominal_induct A avoiding: i j t u rule: v.strong_induct) (auto simp: fresh_at_base) + +lemma subst_ceb_commute [simp]: + "atom j \ A \ (subst_ceb (subst_ceb A i t )) j u = subst_ceb A i (subst_bb t j u ) " + by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base) + +lemma subst_cb_commute [simp]: + "atom j \ A \ (subst_cb (subst_cb A i t )) j u = subst_cb A i (subst_bb t j u ) " + by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base) + +lemma subst_tb_commute [simp]: + "atom j \ A \ (subst_tb (subst_tb A i t )) j u = subst_tb A i (subst_bb t j u ) " +proof (nominal_induct A avoiding: i j t u rule: \.strong_induct) + case (T_refined_type z b c) + then show ?case using subst_tb.simps subst_bb_commute subst_cb_commute by simp +qed + +section \Expressions\ + +nominal_function subst_eb :: "e \ bv \ b \ e" where + "subst_eb ( (AE_val v')) bv b = ( AE_val (subst_vb v' bv b))" +| "subst_eb ( (AE_app f v') ) bv b = ( (AE_app f (subst_vb v' bv b)) )" +| "subst_eb ( (AE_appP f b' v') ) bv b = ( (AE_appP f (b'[bv::=b]\<^sub>b\<^sub>b) (subst_vb v' bv b)))" +| "subst_eb ( (AE_op opp v1 v2) ) bv b = ( (AE_op opp (subst_vb v1 bv b) (subst_vb v2 bv b)) )" +| "subst_eb ( (AE_fst v')) bv b = AE_fst (subst_vb v' bv b)" +| "subst_eb ( (AE_snd v')) bv b = AE_snd (subst_vb v' bv b)" +| "subst_eb ( (AE_mvar u)) bv b = AE_mvar u" +| "subst_eb ( (AE_len v')) bv b = AE_len (subst_vb v' bv b)" +| "subst_eb ( AE_concat v1 v2) bv b = AE_concat (subst_vb v1 bv b) (subst_vb v2 bv b)" +| "subst_eb ( AE_split v1 v2) bv b = AE_split (subst_vb v1 bv b) (subst_vb v2 bv b)" + apply (simp add: eqvt_def subst_eb_graph_aux_def) + apply auto + by (meson e.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_eb_abbrev :: "e \ bv \ b \ e" ("_[_::=_]\<^sub>e\<^sub>b" [1000,50,50] 500) + where + "e[bv::=b]\<^sub>e\<^sub>b \ subst_eb e bv b" + +instantiation e :: has_subst_b +begin +definition "subst_b = subst_eb" + +instance proof + fix j::atom and i::bv and x::b and t::e + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof (induct t rule: e.induct) + case (AE_val v) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp + e.fresh has_subst_b_class.fresh_subst_if subst_b_e_def subst_b_v_def + by metis + next + case (AE_app f v) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def + e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def + by (metis (mono_tags, hide_lams) e.fresh(2)) + next + case (AE_appP f b' v) + then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using + fresh_def pure_fresh subst_b_e_def e.supp v.supp + e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def) + next + case (AE_op opp v1 v2) + then show ?case unfolding subst_eb.simps subst_b_e_def e.fresh using + fresh_def pure_fresh subst_b_e_def e.supp v.supp fresh_e_opp_all + e.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_vb_def by (metis subst_b_v_def) + next + case (AE_concat x1a x2) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp + has_subst_b_class.fresh_subst_if subst_b_v_def + by (metis subst_vb.simps(5)) + next + case (AE_split x1a x2) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp + has_subst_b_class.fresh_subst_if subst_b_v_def + by (metis subst_vb.simps(5)) + next + case (AE_fst x) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp has_subst_b_class.fresh_subst_if subst_b_v_def by metis + next + case (AE_snd x) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis + next + case (AE_mvar x) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp by auto + next + case (AE_len x) + then show ?case using subst_eb.simps fresh_def pure_fresh subst_b_e_def e.supp v.supp using has_subst_b_class.fresh_subst_if subst_b_v_def by metis + qed + + fix a::bv and tm::e and x::b + show "atom a \ tm \ subst_b tm a x = tm" + apply(induct tm rule: e.induct) + apply( auto simp add: fresh_at_base subst_eb.simps subst_b_e_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.forget_subst subst_b_v_def apply metis+ + done + + fix a::bv and tm::e + show "subst_b tm a (B_var a) = tm" using subst_bb.simps subst_b_b_def + apply (induct tm rule: e.induct) + apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def) + using has_subst_b_class.fresh_subst_if subst_b_b_def e.fresh + using has_subst_b_class.subst_id subst_b_v_def apply metis+ + done + + fix p::perm and x1::bv and v::b and t1::e + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + apply(induct tm rule: e.induct) + apply( auto simp add: fresh_at_base subst_bb.simps subst_b_b_def ) + using has_subst_b_class.eqvt subst_b_b_def e.fresh + using has_subst_b_class.eqvt + by (simp add: subst_b_e_def)+ + + fix bv::bv and c::e and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (induct c rule: e.induct) + apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp ) + using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def + flip_fresh_fresh subst_b_\_def apply metis + apply (metis (full_types) opp.perm_simps opp.strong_exhaust) + done + + fix bv::bv and c::e and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + apply (induct c rule: e.induct) + apply(auto simp add: fresh_at_base subst_eb.simps subst_b_e_def subst_b_v_def permute_pure pure_supp ) + using flip_fresh_fresh permute_flip_cancel2 has_subst_b_class.flip_subst subst_b_v_def subst_b_b_def + flip_fresh_fresh subst_b_\_def apply simp + + apply (metis (full_types) opp.perm_simps opp.strong_exhaust) + done +qed +end + +section \Statements\ + +nominal_function (default "case_sum (\x. Inl undefined) (case_sum (\x. Inl undefined) (\x. Inr undefined))") + subst_sb :: "s \ bv \ b \ s" + and subst_branchb :: "branch_s \ bv \ b \ branch_s" + and subst_branchlb :: "branch_list \ bv \ b \ branch_list" + where + "subst_sb (AS_val v') bv b = (AS_val (subst_vb v' bv b))" + | "subst_sb (AS_let y e s) bv b = (AS_let y (e[bv::=b]\<^sub>e\<^sub>b) (subst_sb s bv b ))" + | "subst_sb (AS_let2 y t s1 s2) bv b = (AS_let2 y (subst_tb t bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b))" + | "subst_sb (AS_match v' cs) bv b = AS_match (subst_vb v' bv b) (subst_branchlb cs bv b)" + | "subst_sb (AS_assign y v') bv b = AS_assign y (subst_vb v' bv b)" + | "subst_sb (AS_if v' s1 s2) bv b = (AS_if (subst_vb v' bv b) (subst_sb s1 bv b ) (subst_sb s2 bv b ) )" + | "subst_sb (AS_var u \ v' s) bv b = AS_var u (subst_tb \ bv b) (subst_vb v' bv b) (subst_sb s bv b )" + | "subst_sb (AS_while s1 s2) bv b = AS_while (subst_sb s1 bv b ) (subst_sb s2 bv b )" + | "subst_sb (AS_seq s1 s2) bv b = AS_seq (subst_sb s1 bv b ) (subst_sb s2 bv b )" + | "subst_sb (AS_assert c s) bv b = AS_assert (subst_cb c bv b ) (subst_sb s bv b )" + +| "subst_branchb (AS_branch dc x1 s') bv b = AS_branch dc x1 (subst_sb s' bv b)" + +| "subst_branchlb (AS_final sb) bv b = AS_final (subst_branchb sb bv b )" +| "subst_branchlb (AS_cons sb ssb) bv b = AS_cons (subst_branchb sb bv b) (subst_branchlb ssb bv b)" + + apply (simp add: eqvt_def subst_sb_subst_branchb_subst_branchlb_graph_aux_def ) + apply (auto,metis s_branch_s_branch_list.exhaust s_branch_s_branch_list.exhaust(2) old.sum.exhaust surj_pair) + +proof(goal_cases) + + have eqvt_at_proj: "\ s xa va . eqvt_at subst_sb_subst_branchb_subst_branchlb_sumC (Inl (s, xa, va)) \ + eqvt_at (\a. projl (subst_sb_subst_branchb_subst_branchlb_sumC (Inl a))) (s, xa, va)" + apply(simp only: eqvt_at_def) + apply(rule) + apply(subst Projl_permute) + apply(thin_tac _)+ + apply(simp add: subst_sb_subst_branchb_subst_branchlb_sumC_def) + apply(simp add: THE_default_def) + apply(case_tac "Ex1 (subst_sb_subst_branchb_subst_branchlb_graph (Inl (s,xa,va)))") + apply simp + apply(auto)[1] + apply(erule_tac x="x" in allE) + apply simp + apply(cases rule: subst_sb_subst_branchb_subst_branchlb_graph.cases) + apply(assumption) + apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+ + apply(blast)+ + apply(simp)+ + done + { + case (1 y s bv b ya sa c) + moreover have "atom y \ (bv, b) \ atom ya \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp + ultimately show ?case + using eqvt_triple eqvt_at_proj by metis + next + case (2 y s1 s2 bv b ya s2a c) + moreover have "atom y \ (bv, b) \ atom ya \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp + ultimately show ?case + using eqvt_triple eqvt_at_proj by metis + next + case (3 u s bv b ua sa c) + moreover have "atom u \ (bv, b) \ atom ua \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp + ultimately show ?case using eqvt_triple eqvt_at_proj by metis + next + case (4 x1 s' bv b x1a s'a c) + moreover have "atom x1 \ (bv, b) \ atom x1a \ (bv, b)" using x_fresh_b x_fresh_bv fresh_Pair by simp + ultimately show ?case using eqvt_triple eqvt_at_proj by metis + } +qed + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_sb_abbrev :: "s \ bv \ b \ s" ("_[_::=_]\<^sub>s\<^sub>b" [1000,50,50] 1000) + where + "b[bv::=b']\<^sub>s\<^sub>b \ subst_sb b bv b'" + +lemma fresh_subst_sb_if [simp]: + "(j \ (subst_sb A i x )) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" and + "(j \ (subst_branchb B i x )) = ((atom i \ B \ j \ B) \ (j \ x \ (j \ B \ j = atom i)))" and + "(j \ (subst_branchlb C i x )) = ((atom i \ C \ j \ C) \ (j \ x \ (j \ C \ j = atom i)))" +proof (nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct) + case (AS_branch x1 x2 x3) + have " (j \ subst_branchb (AS_branch x1 x2 x3) i x ) = (j \ (AS_branch x1 x2 (subst_sb x3 i x)))" by auto + also have "... = ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1)" using s_branch_s_branch_list.fresh by auto + also have "... = ((atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3) \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" + using subst_branchb.simps(1) s_branch_s_branch_list.fresh(1) fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\_def + v.fresh AS_branch + proof - + have f1: "\cs b. atom (b::bv) \ (cs::char list)" using pure_fresh by auto + + then have "j \ x \ atom i = j \ ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1) = (atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3 \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" + by (metis (full_types) AS_branch.hyps(3)) + then have "j \ x \ ((j \ x3[i::=x]\<^sub>s\<^sub>b \ j \ set [atom x2]) \ j \ x1) = (atom i \ AS_branch x1 x2 x3 \ j \ AS_branch x1 x2 x3 \ j \ x \ (j \ AS_branch x1 x2 x3 \ j = atom i))" + using AS_branch.hyps s_branch_s_branch_list.fresh by metis + moreover + { assume "\ j \ x" + have ?thesis + using f1 AS_branch.hyps(2) AS_branch.hyps(3) by force } + ultimately show ?thesis + by satx + qed + finally show ?case by auto +next + case (AS_cons cs css i x) + show ?case + unfolding subst_branchlb.simps s_branch_s_branch_list.fresh + using AS_cons by auto +next + case (AS_val xx) + then show ?case using subst_sb.simps(1) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_b_def subst_b_v_def by metis +next + case (AS_let x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def + by fastforce +next + case (AS_let2 x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_\_def + by fastforce +next + case (AS_if x1 x2 x3) + then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using + has_subst_b_class.fresh_subst_if subst_b_v_def by metis +next + case (AS_var u t v s) + + have "(((atom i \ s \ j \ s \ j \ x \ (j \ s \ j = atom i)) \ j \ set [atom u]) \ j \ t[i::=x]\<^sub>\\<^sub>b \ j \ v[i::=x]\<^sub>v\<^sub>b) = + (((atom i \ s \ j \ s \ j \ x \ (j \ s \ j = atom i)) \ j \ set [atom u]) \ + ((atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))) \ + ((atom i \ v \ j \ v \ j \ x \ (j \ v \ j = atom i))))" + using has_subst_b_class.fresh_subst_if subst_b_v_def subst_b_\_def by metis + also have "... = (((atom i \ s \ atom i \ set [atom u]) \ atom i \ t \ atom i \ v) \ + (j \ s \ j \ set [atom u]) \ j \ t \ j \ v \ j \ x \ ((j \ s \ j \ set [atom u]) \ j \ t \ j \ v \ j = atom i))" + using u_fresh_b by auto + finally show ?case using subst_sb.simps s_branch_s_branch_list.fresh AS_var + by simp +next + case (AS_assign u v ) + then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh using + has_subst_b_class.fresh_subst_if subst_b_v_def by force +next + case (AS_match v cs) + have " j \ (AS_match v cs)[i::=x]\<^sub>s\<^sub>b = j \ (AS_match (subst_vb v i x) (subst_branchlb cs i x ))" using subst_sb.simps by auto + also have "... = (j \ (subst_vb v i x) \ j \ (subst_branchlb cs i x ))" using s_branch_s_branch_list.fresh by simp + also have "... = (j \ (subst_vb v i x) \ ((atom i \ cs \ j \ cs) \ j \ x \ (j \ cs \ j = atom i)))" using AS_match[of i x] by auto + also have "... = (atom i \ AS_match v cs \ j \ AS_match v cs \ j \ x \ (j \ AS_match v cs \ j = atom i))" + by (metis (no_types) s_branch_s_branch_list.fresh has_subst_b_class.fresh_subst_if subst_b_v_def) (* 93 ms *) + finally show ?case by auto +next + case (AS_while x1 x2) + then show ?case by auto +next + case (AS_seq x1 x2) + then show ?case by auto +next + case (AS_assert x1 x2) + then show ?case unfolding subst_sb.simps s_branch_s_branch_list.fresh + using fresh_at_base has_subst_b_class.fresh_subst_if list.distinct list.set_cases set_ConsD subst_b_e_def + by (metis subst_b_c_def) +qed(auto+) + +lemma + forget_subst_sb[simp]: "atom a \ A \ subst_sb A a x = A" and + forget_subst_branchb [simp]: "atom a \ B \ subst_branchb B a x = B" and + forget_subst_branchlb[simp]: "atom a \ C \ subst_branchlb C a x = C" +proof (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct) + case (AS_let x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_let2 x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_\_def by force +next + case (AS_var x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def using subst_b_\_def + proof - + have f1: "(atom a \ x4 \ atom a \ set [atom x1]) \ atom a \ x2 \ atom a \ x3" + using AS_var.prems s_branch_s_branch_list.fresh by simp + then have "atom a \ x4" + by (metis (no_types) "Nominal-Utils.fresh_star_singleton" AS_var.hyps(1) empty_set fresh_star_def list.simps(15) not_self_fresh) (* 46 ms *) + then show ?thesis + using f1 by (metis AS_var.hyps(3) has_subst_b_class.forget_subst subst_b_\_def subst_b_v_def subst_sb.simps(7)) (* 62 ms *) + qed +next + case (AS_branch x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_cons x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_val x) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_if x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_assign x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_match x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_while x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_seq x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def by force +next + case (AS_assert c s) + then show ?case unfolding subst_sb.simps using + s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.forget_subst subst_b_v_def subst_b_c_def subst_cb.simps by force +qed(auto+) + +lemma subst_sb_id: "subst_sb A a (B_var a) = A" and + subst_branchb_id [simp]: "subst_branchb B a (B_var a) = B" and + subst_branchlb_id: "subst_branchlb C a (B_var a) = C" +proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct) + case (AS_branch x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def + by simp +next + case (AS_cons x1 x2 ) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by simp +next + case (AS_val x) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_if x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_assign x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_match x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_while x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_seq x1 x2) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_let x1 x2 x3) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_e_def has_subst_b_class.subst_id by metis +next + case (AS_let2 x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id by metis +next + case (AS_var x1 x2 x3 x4) + then show ?case using subst_sb.simps s_branch_s_branch_list.fresh subst_b_\_def has_subst_b_class.subst_id subst_b_v_def by metis +next + case (AS_assert c s ) + then show ?case unfolding subst_sb.simps using s_branch_s_branch_list.fresh subst_b_c_def has_subst_b_class.subst_id by metis +qed (auto) + +lemma flip_subst_s: + fixes bv::bv and s::s and cs::branch_s and z::bv + shows "atom bv \ s \ ((bv \ z) \ s) = s[z::=B_var bv]\<^sub>s\<^sub>b" and + "atom bv \ cs \ ((bv \ z) \ cs) = subst_branchb cs z (B_var bv) " and + "atom bv \ css \ ((bv \ z) \ css) = subst_branchlb css z (B_var bv) " + +proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct) + case (AS_branch x1 x2 x3) + hence "((bv \ z) \ x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis + moreover have "((bv \ z) \ x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto + ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto +next + case (AS_cons x1 x2 ) + hence "((bv \ z) \ x1) = subst_branchb x1 z (B_var bv)" using pure_fresh fresh_at_base flip_fresh_fresh s_branch_s_branch_list.fresh(13) by metis + moreover have "((bv \ z) \ x2) = subst_branchlb x2 z (B_var bv)" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_cons s_branch_s_branch_list.fresh by metis + ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_cons by auto +next + case (AS_val x) + then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp +next + case (AS_let x1 x2 x3) + moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto + ultimately show ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def s_branch_s_branch_list.fresh by auto +next + case (AS_let2 x1 x2 x3 x4) + moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto + ultimately show ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\_def by auto +next + case (AS_if x1 x2 x3) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_var x1 x2 x3 x4) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def subst_b_\_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto +next + case (AS_assign x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto +next + case (AS_match x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_while x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_seq x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_assert x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_c_def subst_b_v_def s_branch_s_branch_list.fresh by simp +qed(auto) + +lemma flip_subst_subst_s: + fixes bv::bv and s::s and cs::branch_s and z::bv + shows "atom bv \ s \ ((bv \ z) \ s)[bv::=v]\<^sub>s\<^sub>b = s[z::=v]\<^sub>s\<^sub>b" and + "atom bv \ cs \ subst_branchb ((bv \ z) \ cs) bv v = subst_branchb cs z v" and + "atom bv \ css \ subst_branchlb ((bv \ z) \ css) bv v = subst_branchlb css z v " +proof(nominal_induct s and cs and css rule: s_branch_s_branch_list.strong_induct) + case (AS_branch x1 x2 x3) + hence "((bv \ z) \ x1) = x1" using pure_fresh fresh_at_base flip_fresh_fresh by metis + moreover have "((bv \ z) \ x2) = x2" using fresh_at_base flip_fresh_fresh[of bv x2 z] AS_branch by auto + ultimately show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using s_branch_s_branch_list.fresh(1) AS_branch by auto +next + case (AS_cons x1 x2 ) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_branchlb.simps + using s_branch_s_branch_list.fresh(1) AS_cons by auto + +next + case (AS_val x) + then show ?case unfolding s_branch_s_branch_list.perm_simps subst_branchb.simps using flip_subst subst_b_v_def by simp +next + case (AS_let x1 x2 x3) + moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto + ultimately show ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst_subst subst_b_e_def s_branch_s_branch_list.fresh by force +next + case (AS_let2 x1 x2 x3 x4) + moreover hence "((bv \ z) \ x1) = x1" using fresh_at_base flip_fresh_fresh[of bv x1 z] by auto + ultimately show ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst s_branch_s_branch_list.fresh(5) subst_b_\_def by auto +next + case (AS_if x1 x2 x3) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_var x1 x2 x3 x4) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def subst_b_\_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto +next + case (AS_assign x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh fresh_at_base flip_fresh_fresh[of bv x1 z] by auto +next + case (AS_match x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_while x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_seq x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_v_def s_branch_s_branch_list.fresh by auto +next + case (AS_assert x1 x2) + thus ?case + unfolding s_branch_s_branch_list.perm_simps subst_sb.simps + using flip_subst subst_b_e_def subst_b_c_def s_branch_s_branch_list.fresh by auto +qed(auto) + +instantiation s :: has_subst_b +begin +definition "subst_b = (\s bv b. subst_sb s bv b)" + +instance proof + fix j::atom and i::bv and x::b and t::s + show "j \ subst_b t i x = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_sb_if subst_b_s_def by metis + + fix a::bv and tm::s and x::b + show "atom a \ tm \ subst_b tm a x = tm" using subst_b_s_def forget_subst_sb by metis + + fix a::bv and tm::s + show "subst_b tm a (B_var a) = tm" using subst_b_s_def subst_sb_id by metis + + fix p::perm and x1::bv and v::b and t1::s + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" using subst_b_s_def subst_sb_subst_branchb_subst_branchlb.eqvt by metis + + fix bv::bv and c::s and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + using subst_b_s_def flip_subst_s by metis + + fix bv::bv and c::s and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + using flip_subst_subst_s subst_b_s_def by metis +qed +end + +section \Function Type\ + +nominal_function subst_ft_b :: "fun_typ \ bv \ b \ fun_typ" where + "subst_ft_b ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z (subst_bb b x v) (subst_cb c x v) t[x::=v]\<^sub>\\<^sub>b s[x::=v]\<^sub>s\<^sub>b" + apply(simp add: eqvt_def subst_ft_b_graph_aux_def ) + apply(simp add:fun_typ.strong_exhaust,auto ) + apply(rule_tac y=a and c="(a,b)" in fun_typ.strong_exhaust) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) + done + +nominal_termination (eqvt) by lexicographic_order + +nominal_function subst_ftq_b :: "fun_typ_q \ bv \ b \ fun_typ_q" where + "atom bv \ (x,v) \ subst_ftq_b (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_b ft x v))" +| "subst_ftq_b (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_b ft x v))" + apply(simp add: eqvt_def subst_ftq_b_graph_aux_def ) + apply(simp add:fun_typ_q.strong_exhaust,auto ) + apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust) + by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) +nominal_termination (eqvt) by lexicographic_order + +instantiation fun_typ :: has_subst_b +begin +definition "subst_b = subst_ft_b" + +text \Note: Using just simp in the second apply unpacks and gives a single goal +whereas auto gives 43 non-intuitive goals. These goals are easier to solve +and tedious, however they that it clear if a mistake is made in the definition of the function. +For example, I saw that one of the goals was +going through with metis and the next wasn't. +It turned out the definition of the function itself was wrong\ + +instance proof + fix j::atom and i::bv and x::b and t::fun_typ + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct) + apply(auto simp add: subst_b_fun_typ_def ) + by(metis fresh_subst_if subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def)+ + + fix a::bv and tm::fun_typ and x::b + show "atom a \ tm \ subst_b tm a x = tm" + apply (nominal_induct tm avoiding: a x rule: fun_typ.strong_induct) + apply(simp add: subst_b_fun_typ_def Abs1_eq_iff') + using subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def + forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD + subst_ft_b.simps by metis + + fix a::bv and tm::fun_typ + show "subst_b tm a (B_var a) = tm" + apply (nominal_induct tm rule: fun_typ.strong_induct) + apply(simp add: subst_b_fun_typ_def Abs1_eq_iff',auto) + using subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def + forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD + subst_ft_b.simps + by (metis has_subst_b_class.subst_id)+ + + fix p::perm and x1::bv and v::b and t1::fun_typ + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v) " + apply (nominal_induct t1 avoiding: x1 v rule: fun_typ.strong_induct) + by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps) + + fix bv::bv and c::fun_typ and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (nominal_induct c avoiding: z bv rule: fun_typ.strong_induct) + by(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def) + + fix bv::bv and c::fun_typ and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + apply (nominal_induct c avoiding: bv v z rule: fun_typ.strong_induct) + apply(auto simp add: subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def flip_subst_subst flip_subst) + using subst_b_fun_typ_def Abs1_eq_iff' fun_typ.perm_simps subst_b_b_def subst_b_c_def subst_b_\_def subst_b_s_def flip_subst_subst flip_subst + using flip_subst_s(1) flip_subst_subst_s(1) by auto +qed +end + +instantiation fun_typ_q :: has_subst_b +begin +definition "subst_b = subst_ftq_b" + +instance proof + fix j::atom and i::bv and x::b and t::fun_typ_q + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + apply (nominal_induct t avoiding: i x j rule: fun_typ_q.strong_induct,auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps) + using fresh_subst_if subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def apply metis+ + done + + fix a::bv and t::fun_typ_q and x::b + show "atom a \ t \ subst_b t a x = t" + apply (nominal_induct t avoiding: a x rule: fun_typ_q.strong_induct) + apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') + using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ + + fix p::perm and x1::bv and v::b and t::fun_typ_q + show "p \ subst_b t x1 v = subst_b (p \ t) (p \ x1) (p \ v) " + apply (nominal_induct t avoiding: x1 v rule: fun_typ_q.strong_induct) + by(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') + + fix a::bv and tm::fun_typ_q + show "subst_b tm a (B_var a) = tm" + apply (nominal_induct tm avoiding: a rule: fun_typ_q.strong_induct) + apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') + using subst_id subst_b_b_def subst_b_fun_typ_def subst_b_\_def subst_b_c_def subst_b_s_def + forget_subst fresh_at_base list.set_cases neq_Nil_conv set_ConsD + subst_ft_b.simps by metis+ + + fix bv::bv and c::fun_typ_q and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + apply (nominal_induct c avoiding: z bv rule: fun_typ_q.strong_induct) + apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') + using forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ + + fix bv::bv and c::fun_typ_q and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + apply (nominal_induct c avoiding: z v bv rule: fun_typ_q.strong_induct) + apply(auto simp add: subst_b_fun_typ_q_def subst_ftq_b.simps Abs1_eq_iff') + using flip_subst flip_subst_subst forget_subst subst_b_fun_typ_q_def subst_b_s_def subst_b_\_def subst_b_b_def subst_b_c_def subst_b_fun_typ_def eqvt by metis+ + +qed +end + +section \Contexts\ + +subsection \Immutable Variables\ +nominal_function subst_gb :: "\ \ bv \ b \ \" where + "subst_gb GNil _ _ = GNil" +| "subst_gb ((y,b',c)#\<^sub>\\) bv b = ((y,b'[bv::=b]\<^sub>b\<^sub>b,c[bv::=b]\<^sub>c\<^sub>b)#\<^sub>\ (subst_gb \ bv b))" + apply (simp add: eqvt_def subst_gb_graph_aux_def )+ + apply auto + apply (insert \.exhaust neq_GNil_conv, force) + done +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_gb_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) + where + "g[bv::=b']\<^sub>\\<^sub>b \ subst_gb g bv b'" + +instantiation \ :: has_subst_b +begin +definition "subst_b = subst_gb" + +instance proof + fix j::atom and i::bv and x::b and t::\ + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof(induct t rule: \_induct) + case GNil + then show ?case using fresh_GNil subst_gb.simps fresh_def pure_fresh subst_b_\_def has_subst_b_class.fresh_subst_if fresh_GNil fresh_GCons by metis + next + case (GCons x' b' c' \') + have *: "atom i \ x' " using fresh_at_base by simp + + have "j \ subst_b ((x', b', c') #\<^sub>\ \') i x = j \ ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b) #\<^sub>\ (subst_b \' i x))" using subst_gb.simps subst_b_\_def by auto + also have "... = (j \ ((x', b'[i::=x]\<^sub>b\<^sub>b, c'[i::=x]\<^sub>c\<^sub>b)) \ (j \ (subst_b \' i x)))" using fresh_GCons by auto + also have "... = (((j \ x') \ (j \ b'[i::=x]\<^sub>b\<^sub>b) \ (j \ c'[i::=x]\<^sub>c\<^sub>b)) \ (j \ (subst_b \' i x)))" by auto + also have "... = (((j \ x') \ ((atom i \ b' \ j \ b' \ j \ x \ (j \ b' \ j = atom i))) \ + ((atom i \ c' \ j \ c' \ j \ x \ (j \ c' \ j = atom i))) \ + ((atom i \ \' \ j \ \' \ j \ x \ (j \ \' \ j = atom i)))))" + using fresh_subst_if[of j b' i x] fresh_subst_if[of j c' i x] GCons subst_b_b_def subst_b_c_def by simp + also have "... = ((atom i \ (x', b', c') #\<^sub>\ \' \ j \ (x', b', c') #\<^sub>\ \') \ (j \ x \ (j \ (x', b', c') #\<^sub>\ \' \ j = atom i)))" using * fresh_GCons fresh_prod3 by metis + + finally show ?case by auto + qed + + fix a::bv and tm::\ and x::b + show "atom a \ tm \ subst_b tm a x = tm" + proof (induct tm rule: \_induct) + case GNil + then show ?case using subst_gb.simps subst_b_\_def by auto + next + case (GCons x' b' c' \') + have *:"b'[a::=x]\<^sub>b\<^sub>b = b' \ c'[a::=x]\<^sub>c\<^sub>b = c'" using GCons fresh_GCons[of "atom a"] fresh_prod3[of "atom a"] has_subst_b_class.forget_subst subst_b_b_def subst_b_c_def by metis + have "subst_b ((x', b', c') #\<^sub>\ \') a x = ((x', b'[a::=x]\<^sub>b\<^sub>b, c'[a::=x]\<^sub>c\<^sub>b) #\<^sub>\ (subst_b \' a x))" using subst_b_\_def subst_gb.simps by auto + also have "... = ((x', b', c') #\<^sub>\ \')" using * GCons fresh_GCons[of "atom a"] by auto + finally show ?case using has_subst_b_class.forget_subst fresh_GCons fresh_prod3 GCons subst_b_\_def has_subst_b_class.forget_subst[of a b' x] fresh_prod3[of "atom a"] by argo + qed + + fix a::bv and tm::\ + show "subst_b tm a (B_var a) = tm" + proof(induct tm rule: \_induct) + case GNil + then show ?case using subst_gb.simps subst_b_\_def by auto + next + case (GCons x' b' c' \') + then show ?case using has_subst_b_class.subst_id subst_b_\_def subst_b_b_def subst_b_c_def subst_gb.simps by metis + qed + + fix p::perm and x1::bv and v::b and t1::\ + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + proof (induct tm rule: \_induct) + case GNil + then show ?case using subst_b_\_def subst_gb.simps by simp + next + case (GCons x' b' c' \') + then show ?case using subst_b_\_def subst_gb.simps has_subst_b_class.eqvt by argo + qed + + fix bv::bv and c::\ and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + proof (induct c rule: \_induct) + case GNil + then show ?case using subst_b_\_def subst_gb.simps by auto + next + case (GCons x b c \') + have *:"(bv \ z) \ (x, b, c) = (x, (bv \ z) \ b, (bv \ z) \ c)" using flip_bv_x_cancel by auto + then show ?case + unfolding subst_gb.simps subst_b_\_def permute_\.simps * + using GCons subst_b_\_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto + qed + + fix bv::bv and c::\ and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + proof (induct c rule: \_induct) + case GNil + then show ?case using subst_b_\_def subst_gb.simps by auto + next + case (GCons x b c \') + have *:"(bv \ z) \ (x, b, c) = (x, (bv \ z) \ b, (bv \ z) \ c)" using flip_bv_x_cancel by auto + then show ?case + unfolding subst_gb.simps subst_b_\_def permute_\.simps * + using GCons subst_b_\_def subst_gb.simps flip_subst subst_b_b_def subst_b_c_def fresh_GCons by auto + qed +qed +end + +lemma subst_b_base_for_lit: + "(base_for_lit l)[bv::=b]\<^sub>b\<^sub>b = base_for_lit l" + using base_for_lit.simps l.strong_exhaust + by (metis subst_bb.simps(2) subst_bb.simps(3) subst_bb.simps(6) subst_bb.simps(7)) + +lemma subst_b_lookup: + assumes "Some (b, c) = lookup \ x" + shows " Some (b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) = lookup \[bv::=b']\<^sub>\\<^sub>b x" + using assms by(induct \ rule: \_induct, auto) + +lemma subst_g_b_x_fresh: + fixes x::x and b::b and \::\ and bv::bv + assumes "atom x \ \" + shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" + using subst_b_fresh_x subst_b_\_def assms by metis + +subsection \Mutable Variables\ + +nominal_function subst_db :: "\ \ bv \ b \ \" where + "subst_db []\<^sub>\ _ _ = []\<^sub>\" +| "subst_db ((u,t) #\<^sub>\ \) bv b = ((u,t[bv::=b]\<^sub>\\<^sub>b) #\<^sub>\ (subst_db \ bv b))" + apply (simp add: eqvt_def subst_db_graph_aux_def,auto ) + using list.exhaust delete_aux.elims + using neq_DNil_conv by fastforce +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_db_abbrev :: "\ \ bv \ b \ \" ("_[_::=_]\<^sub>\\<^sub>b" [1000,50,50] 1000) + where + "\[bv::=b]\<^sub>\\<^sub>b \ subst_db \ bv b" + +instantiation \ :: has_subst_b +begin +definition "subst_b = subst_db" + +instance proof + fix j::atom and i::bv and x::b and t::\ + show "j \ subst_b t i x = (atom i \ t \ j \ t \ j \ x \ (j \ t \ j = atom i))" + proof(induct t rule: \_induct) + case DNil + then show ?case using fresh_DNil subst_db.simps fresh_def pure_fresh subst_b_\_def has_subst_b_class.fresh_subst_if fresh_DNil fresh_DCons by metis + next + case (DCons u t \') + have "j \ subst_b ((u, t) #\<^sub>\ \') i x = j \ ((u, t[i::=x]\<^sub>\\<^sub>b) #\<^sub>\ (subst_b \' i x))" using subst_db.simps subst_b_\_def by auto + also have "... = (j \ ((u, t[i::=x]\<^sub>\\<^sub>b)) \ (j \ (subst_b \' i x)))" using fresh_DCons by auto + also have "... = (((j \ u) \ (j \ t[i::=x]\<^sub>\\<^sub>b)) \ (j \ (subst_b \' i x)))" by auto + also have "... = ((j \ u) \ ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i))) \ (atom i \ \' \ j \ \' \ j \ x \ (j \ \' \ j = atom i)))" + using has_subst_b_class.fresh_subst_if[of j t i x] subst_b_\_def DCons subst_b_\_def by auto + also have "... = (atom i \ (u, t) #\<^sub>\ \' \ j \ (u, t) #\<^sub>\ \' \ j \ x \ (j \ (u, t) #\<^sub>\ \' \ j = atom i))" + using DCons subst_db.simps(2) has_subst_b_class.fresh_subst_if fresh_DCons subst_b_\_def pure_fresh fresh_at_base by auto + finally show ?case by auto + qed + + fix a::bv and tm::\ and x::b + show "atom a \ tm \ subst_b tm a x = tm" + proof (induct tm rule: \_induct) + case DNil + then show ?case using subst_db.simps subst_b_\_def by auto + next + case (DCons u t \') + have *:"t[a::=x]\<^sub>\\<^sub>b = t" using DCons fresh_DCons[of "atom a"] fresh_prod2[of "atom a"] has_subst_b_class.forget_subst subst_b_\_def by metis + have "subst_b ((u,t) #\<^sub>\ \') a x = ((u,t[a::=x]\<^sub>\\<^sub>b) #\<^sub>\ (subst_b \' a x))" using subst_b_\_def subst_db.simps by auto + also have "... = ((u, t) #\<^sub>\ \')" using * DCons fresh_DCons[of "atom a"] by auto + finally show ?case using + has_subst_b_class.forget_subst fresh_DCons fresh_prod3 + DCons subst_b_\_def has_subst_b_class.forget_subst[of a t x] fresh_prod3[of "atom a"] by argo + qed + + fix a::bv and tm::\ + show "subst_b tm a (B_var a) = tm" + proof(induct tm rule: \_induct) + case DNil + then show ?case using subst_db.simps subst_b_\_def by auto + next + case (DCons u t \') + then show ?case using has_subst_b_class.subst_id subst_b_\_def subst_b_\_def subst_db.simps by metis + qed + + fix p::perm and x1::bv and v::b and t1::\ + show "p \ subst_b t1 x1 v = subst_b (p \ t1) (p \ x1) (p \ v)" + proof (induct tm rule: \_induct) + case DNil + then show ?case using subst_b_\_def subst_db.simps by simp + next + case (DCons x' b' \') + then show ?case by argo + qed + + fix bv::bv and c::\ and z::bv + show "atom bv \ c \ ((bv \ z) \ c) = c[z::=B_var bv]\<^sub>b" + proof (induct c rule: \_induct) + case DNil + then show ?case using subst_b_\_def subst_db.simps by auto + next + case (DCons u t') + then show ?case + unfolding subst_db.simps subst_b_\_def permute_\.simps + using DCons subst_b_\_def subst_db.simps flip_subst subst_b_\_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp + qed + + fix bv::bv and c::\ and z::bv and v::b + show "atom bv \ c \ ((bv \ z) \ c)[bv::=v]\<^sub>b = c[z::=v]\<^sub>b" + proof (induct c rule: \_induct) + case DNil + then show ?case using subst_b_\_def subst_db.simps by auto + next + case (DCons u t') + then show ?case + unfolding subst_db.simps subst_b_\_def permute_\.simps + using DCons subst_b_\_def subst_db.simps flip_subst subst_b_\_def flip_fresh_fresh fresh_at_base fresh_DCons flip_bv_u_cancel by simp + qed +qed +end + +lemma subst_d_b_member: + assumes "(u, \) \ setD \" + shows "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" + using assms by (induct \,auto) + +lemmas ms_fresh_all = e.fresh s_branch_s_branch_list.fresh \.fresh c.fresh ce.fresh v.fresh l.fresh fresh_at_base opp.fresh pure_fresh ms_fresh + +lemmas fresh_intros[intro] = fresh_GNil x_not_in_b_set x_not_in_u_atoms x_fresh_b u_not_in_x_atoms bv_not_in_x_atoms u_not_in_b_atoms + +lemmas subst_b_simps = subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps subst_bb.simps subst_eb.simps subst_branchb.simps subst_sb.simps + +lemma subst_d_b_x_fresh: + fixes x::x and b::b and \::\ and bv::bv + assumes "atom x \ \" + shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" + using subst_b_fresh_x subst_b_\_def assms by metis + +lemma subst_b_fresh_x: + fixes x::x + shows "atom x \ v \ atom x \ v[bv::=b']\<^sub>v\<^sub>b" and + "atom x \ ce \ atom x \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and + "atom x \ e \ atom x \ e[bv::=b']\<^sub>e\<^sub>b" and + "atom x \ c \ atom x \ c[bv::=b']\<^sub>c\<^sub>b" and + "atom x \ t \ atom x \ t[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ d \ atom x \ d[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ g \ atom x \ g[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ s \ atom x \ s[bv::=b']\<^sub>s\<^sub>b" + using fresh_subst_if x_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\_def subst_b_s_def subst_g_b_x_fresh subst_d_b_x_fresh + by metis+ + +lemma subst_b_fresh_u_cls: + fixes tm::"'a::has_subst_b" and x::u + shows "atom x \ tm = atom x \ tm[bv::=b']\<^sub>b" + using fresh_subst_if[of "atom x" tm bv b'] using u_fresh_b by auto + +lemma subst_g_b_u_fresh: + fixes x::u and b::b and \::\ and bv::bv + assumes "atom x \ \" + shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" + using subst_b_fresh_u_cls subst_b_\_def assms by metis + +lemma subst_d_b_u_fresh: + fixes x::u and b::b and \::\ and bv::bv + assumes "atom x \ \" + shows "atom x \ \[bv::=b]\<^sub>\\<^sub>b" + using subst_b_fresh_u_cls subst_b_\_def assms by metis + +lemma subst_b_fresh_u: + fixes x::u + shows "atom x \ v \ atom x \ v[bv::=b']\<^sub>v\<^sub>b" and + "atom x \ ce \ atom x \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b" and + "atom x \ e \ atom x \ e[bv::=b']\<^sub>e\<^sub>b" and + "atom x \ c \ atom x \ c[bv::=b']\<^sub>c\<^sub>b" and + "atom x \ t \ atom x \ t[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ d \ atom x \ d[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ g \ atom x \ g[bv::=b']\<^sub>\\<^sub>b" and + "atom x \ s \ atom x \ s[bv::=b']\<^sub>s\<^sub>b" + using fresh_subst_if u_fresh_b subst_b_v_def subst_b_ce_def subst_b_e_def subst_b_c_def subst_b_\_def subst_b_s_def subst_g_b_u_fresh subst_d_b_u_fresh + by metis+ + +lemma subst_db_u_fresh: + fixes u::u and b::b and D::\ + assumes "atom u \ D" + shows "atom u \ D[bv::=b]\<^sub>\\<^sub>b" + using assms proof(induct D rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' D') + then show ?case using subst_db.simps fresh_def fresh_DCons fresh_subst_if subst_b_\_def + by (metis fresh_Pair u_not_in_b_atoms) +qed + +lemma flip_bt_subst4: + fixes t::\ and bv::bv + assumes "atom bv \ t" + shows "t[bv'::=b]\<^sub>\\<^sub>b = ((bv' \ bv) \ t)[bv::=b]\<^sub>\\<^sub>b" + using flip_subst_subst[OF assms,of bv' b] + by (simp add: flip_commute subst_b_\_def) + +lemma subst_bt_flip_sym: + fixes t1::\ and t2::\ + assumes "atom bv \ b" and "atom bv \ (bv1, bv2, t1, t2)" and "(bv1 \ bv) \ t1 = (bv2 \ bv) \ t2" + shows " t1[bv1::=b]\<^sub>\\<^sub>b = t2[bv2::=b]\<^sub>\\<^sub>b" + using assms flip_bt_subst4[of bv t1 bv1 b ] flip_bt_subst4 fresh_prod4 fresh_Pair by metis + +end \ No newline at end of file diff --git a/thys/MiniSail/BTVSubstTypingL.thy b/thys/MiniSail/BTVSubstTypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/BTVSubstTypingL.thy @@ -0,0 +1,635 @@ +(*<*) +theory BTVSubstTypingL + imports "HOL-Eisbach.Eisbach_Tools" ContextSubtypingL SubstMethods +begin + (*>*) + +chapter \Basic Type Variable Substitution Lemmas\ +text \Lemmas that show that types are preserved, in some way, +under basic type variable substitution\ + +lemma subst_vv_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "(v'[x::=v]\<^sub>v\<^sub>v)[bv::=b]\<^sub>v\<^sub>b = (v'[bv::=b]\<^sub>v\<^sub>b)[x::=v]\<^sub>v\<^sub>v" + using assms proof(nominal_induct v' rule: v.strong_induct) + case (V_lit x) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_var y) + hence "v[bv::=b]\<^sub>v\<^sub>b=v" using forget_subst subst_b_v_def by metis + then show ?case unfolding subst_vb.simps(2) subst_vv.simps(2) using V_var by auto +next + case (V_pair x1a x2a) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_cons x1a x2a x3) + then show ?case using subst_vb.simps subst_vv.simps by simp +next + case (V_consp x1a x2a x3 x4) + then show ?case using subst_vb.simps subst_vv.simps by simp +qed + +lemma subst_cev_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "(ce[x::=v]\<^sub>v)[bv::=b]\<^sub>c\<^sub>e\<^sub>b = (ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b)[x::=v]\<^sub>v" + using assms apply (nominal_induct ce rule: ce.strong_induct, (simp add: subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps)) + using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps + by (simp add: subst_v_ce_def)+ + +lemma subst_cv_subst_bb_commute: + fixes bv::bv and b::b and x::x and v::v + assumes "atom bv \ v" + shows "c[x::=v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[x::=v]\<^sub>c\<^sub>v" + using assms apply (nominal_induct c rule: c.strong_induct ) + using assms subst_vv_subst_bb_commute subst_ceb.simps subst_cv.simps + subst_v_c_def subst_b_c_def apply auto + using subst_cev_subst_bb_commute subst_v_ce_def by auto+ + +lemma subst_b_c_of: + "(c_of \ z)[bv::=b]\<^sub>c\<^sub>b = c_of (\[bv::=b]\<^sub>\\<^sub>b) z" +proof(nominal_induct \ avoiding: z rule:\.strong_induct) + case (T_refined_type z' b' c') + moreover have "atom bv \ [ z ]\<^sup>v " using fresh_at_base v.fresh by auto + ultimately show ?case using subst_cv_subst_bb_commute[of bv "V_var z" c' z' b] c_of.simps subst_tb.simps by metis +qed + +lemma subst_b_b_of: + "(b_of \)[bv::=b]\<^sub>b\<^sub>b = b_of (\[bv::=b]\<^sub>\\<^sub>b)" + by(nominal_induct \ rule:\.strong_induct, simp add: b_of.simps subst_tb.simps ) + +lemma subst_b_if: + "\ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit ll) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \ = \ z : b_of \ | CE_val (v) == CE_val (V_lit ll) IMP c_of \ z \[bv::=b]\<^sub>\\<^sub>b " + unfolding subst_tb.simps subst_cb.simps subst_ceb.simps subst_vb.simps using subst_b_b_of subst_b_c_of by auto + +lemma subst_b_top_eq: + "\ z : B_unit | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_unit | TRUE \" and "\ z : B_bool | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_bool | TRUE \" and "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[bv::=b]\<^sub>\\<^sub>b" + unfolding subst_tb.simps subst_bb.simps subst_cb.simps by auto + +lemmas subst_b_eq = subst_b_c_of subst_b_b_of subst_b_if subst_b_top_eq + +lemma subst_cx_subst_bb_commute[simp]: + fixes bv::bv and b::b and x::x and v'::c + shows "(v'[x::=V_var y]\<^sub>c\<^sub>v)[bv::=b]\<^sub>c\<^sub>b = (v'[bv::=b]\<^sub>c\<^sub>b)[x::=V_var y]\<^sub>c\<^sub>v" + using subst_cv_subst_bb_commute fresh_at_base v.fresh by auto + +lemma subst_b_infer_b: + fixes l::l and b::b + assumes " \ l \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ l \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms infer_l_form3 infer_l_form4 wf_b_subst infer_l_wf subst_tb.simps base_for_lit.simps subst_tb.simps + subst_b_base_for_lit subst_cb.simps(6) subst_ceb.simps(1) subst_vb.simps(1) subst_vb.simps(2) type_l_eq + by metis + +lemma subst_b_subtype: + fixes s::s and b'::b + assumes "\ ; {|bv|} ; \ \ \1 \ \2" and "\ ; {||} \\<^sub>w\<^sub>f b'" and "B = {|bv|}" + shows "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \ \1[bv::=b']\<^sub>\\<^sub>b \ \2[bv::=b']\<^sub>\\<^sub>b" + using assms proof(nominal_induct "{|bv|}" \ \1 \2 rule:subtype.strong_induct) + case (subtype_baseI x \ \ z c z' c' b) + + hence **: "\ ; {|bv|} ; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \ c'[z'::=V_var x]\<^sub>c\<^sub>v " using validI subst_defs by metis + + have "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \ \ z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ \ z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \" proof(rule Typing.subtype_baseI) + show "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \ z : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ " + using subtype_baseI assms wf_b_subst(4) subst_tb.simps subst_defs by metis + show "\ ; {||} ; \[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \ z' : b[bv::=b']\<^sub>b\<^sub>b | c'[bv::=b']\<^sub>c\<^sub>b \ " + using subtype_baseI assms wf_b_subst(4) subst_tb.simps by metis + show "atom x \(\, {||}::bv fset, \[bv::=b']\<^sub>\\<^sub>b, z , c[bv::=b']\<^sub>c\<^sub>b , z' , c'[bv::=b']\<^sub>c\<^sub>b )" + apply(unfold fresh_prodN,auto simp add: * fresh_prodN fresh_empty_fset) + using subst_b_fresh_x * fresh_prodN \atom x \ c\ \atom x \ c'\ subst_defs subtype_baseI by metis+ + have "\ ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, c[z::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \[bv::=b']\<^sub>\\<^sub>b \ c'[z'::=V_var x]\<^sub>v[bv::=b']\<^sub>c\<^sub>b" + using ** subst_b_valid subst_gb.simps assms subtype_baseI by metis + thus "\ ; {||} ; (x, b[bv::=b']\<^sub>b\<^sub>b, (c[bv::=b']\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\ \[bv::=b']\<^sub>\\<^sub>b \ (c'[bv::=b']\<^sub>c\<^sub>b)[z'::=V_var x]\<^sub>v" + using subst_defs subst_cv_subst_bb_commute by (metis subst_cx_subst_bb_commute) + qed + thus ?case using subtype_baseI subst_tb.simps subst_defs by metis +qed + +lemma b_of_subst_bv: + "(b_of \)[x::=v]\<^sub>b\<^sub>b = b_of (\[x::=v]\<^sub>\\<^sub>b)" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (x,v)" using obtain_fresh_z by metis + thus ?thesis using subst_tv.simps * by auto +qed + +lemma subst_b_infer_v: + fixes v::v and b::b + assumes "\ ; B ; G \ v \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(nominal_induct avoiding: b bv rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b' c x z) + show ?case unfolding subst_b_simps proof + show "\ ; {||} \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_v_varI wf_b_subst by metis + show "Some (b'[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \[bv::=b]\<^sub>\\<^sub>b x" using subst_b_lookup infer_v_varI by metis + show "atom z \ x" using infer_v_varI by auto + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b) " by(fresh_mth add: infer_v_varI subst_b_fresh_x subst_b_\_def fresh_prodN fresh_empty_fset ) + qed +next + case (infer_v_litI \ \ \ l \) + then show ?case using Typing.infer_v_litI subst_b_infer_b + using wf_b_subst1(3) by auto +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + show ?case unfolding subst_b_simps b_of_subst_bv proof + show "atom z \ (v1[bv::=b]\<^sub>v\<^sub>b, v2[bv::=b]\<^sub>v\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x) + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b)" by(fresh_mth add: infer_v_pairI subst_b_fresh_x subst_b_\_def fresh_empty_fset) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ t1[bv::=b]\<^sub>\\<^sub>b" using infer_v_pairI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ t2[bv::=b]\<^sub>\\<^sub>b" using infer_v_pairI by auto + qed +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + show ?case unfolding subst_b_simps b_of_subst_bv proof + show "AF_typedef s dclist \ set \" using infer_v_consI by auto + show "(dc, tc) \ set dclist" using infer_v_consI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b" using infer_v_consI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc" proof - + have "atom bv \ tc" using wfTh_lookup_supp_empty fresh_def infer_v_consI infer_v_wf by fast + moreover have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv::=b]\<^sub>\\<^sub>b" + using subst_b_subtype infer_v_consI by simp + ultimately show ?thesis using forget_subst subst_b_\_def by metis + qed + show "atom z \ v[bv::=b]\<^sub>v\<^sub>b" using infer_v_consI using subst_b_fresh_x subst_b_v_def by metis + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b)" by(fresh_mth add: infer_v_consI subst_b_fresh_x subst_b_\_def fresh_empty_fset) + qed +next + case (infer_v_conspI s bv2 dclist2 \ dc tc \ \ v tv ba z) + + have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) \ \ z : B_app s (ba[bv::=b]\<^sub>b\<^sub>b) | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc (ba[bv::=b]\<^sub>b\<^sub>b) (v[bv::=b]\<^sub>v\<^sub>b) ]\<^sup>c\<^sup>e \" + proof(rule Typing.infer_v_conspI) + show "AF_typedef_poly s bv2 dclist2 \ set \" using infer_v_conspI by auto + show "(dc, tc) \ set dclist2" using infer_v_conspI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b" + using infer_v_conspI subst_tb.simps by metis + + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b" proof - + have "supp tc \ { atom bv2 }" using infer_v_conspI wfTh_poly_lookup_supp wfX_wfY by metis + moreover have "bv2 \ bv" using \atom bv2 \ \\ \\ = {|bv|} \ fresh_at_base fresh_def + using fresh_finsert by fastforce + ultimately have "atom bv \ tc" unfolding fresh_def by auto + hence "tc[bv2::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b = tc[bv2::=ba]\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b" + using subst_tb_commute by metis + moreover have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ tv[bv::=b]\<^sub>\\<^sub>b \ tc[bv2::=ba]\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b" + using infer_v_conspI(7) subst_b_subtype infer_v_conspI by metis + ultimately show ?thesis by auto + qed + show "atom z \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)" + apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset) + using \atom z \ \\ fresh_subst_if subst_b_\_def x_fresh_b apply metis + using \atom z \ v\ fresh_subst_if subst_b_v_def x_fresh_b by metis + show "atom bv2 \ (\, {||}, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b)" + apply(unfold fresh_prodN, intro conjI, auto simp add: infer_v_conspI fresh_empty_fset) + using \atom bv2 \ b\ \atom bv2 \ \\ fresh_subst_if subst_b_\_def apply metis + using \atom bv2 \ b\ \atom bv2 \ v\ fresh_subst_if subst_b_v_def apply metis + using \atom bv2 \ b\ \atom bv2 \ ba\ fresh_subst_if subst_b_b_def by metis + show "\ ; {||} \\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b " + using infer_v_conspI wf_b_subst by metis + qed + thus ?case using subst_vb.simps subst_tb.simps subst_bb.simps by simp + +qed + +lemma subst_b_check_v: + fixes v::v and b::b + assumes "\ ; B ; G \ v \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ (\[bv::=b]\<^sub>\\<^sub>b)" +proof - + obtain \' where "\ ; B ; G \ v \ \' \ \ ; B ; G \ \' \ \" using check_v_elims[OF assms(1)] by metis + thus ?thesis using subst_b_subtype subst_b_infer_v assms + by (metis (no_types) check_v_subtypeI subst_b_infer_v subst_b_subtype) +qed + +lemma subst_vv_subst_vb_switch: + shows "(v'[bv::=b']\<^sub>v\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>v\<^sub>v = v'[x::=v]\<^sub>v\<^sub>v[bv::=b']\<^sub>v\<^sub>b" +proof(nominal_induct v' rule:v.strong_induct) + case (V_lit x) + then show ?case using subst_vv.simps subst_vb.simps by auto +next + case (V_var x) + then show ?case using subst_vv.simps subst_vb.simps by auto +next + case (V_pair x1a x2a) + then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto +next + case (V_cons x1a x2a x3) + then show ?case using subst_vv.simps subst_vb.simps v.fresh by auto +next + case (V_consp x1a x2a x3 x4) + then show ?case using subst_vv.simps subst_vb.simps v.fresh pure_fresh + by (metis forget_subst subst_b_b_def) +qed + +lemma subst_cev_subst_vb_switch: + shows "(ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>e\<^sub>v = (ce[x::=v]\<^sub>c\<^sub>e\<^sub>v)[bv::=b']\<^sub>c\<^sub>e\<^sub>b" + by(nominal_induct ce rule:ce.strong_induct, auto simp add: subst_vv_subst_vb_switch ce.fresh) + +lemma subst_cv_subst_vb_switch: + shows "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" + by(nominal_induct c rule:c.strong_induct, auto simp add: subst_cev_subst_vb_switch c.fresh) + +lemma subst_tv_subst_vb_switch: + shows "(\[bv::=b']\<^sub>\\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>v[bv::=b']\<^sub>\\<^sub>b" +proof(nominal_induct \ avoiding: x v rule:\.strong_induct) + case (T_refined_type z b c ) + hence ceq: "(c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v = c[x::=v]\<^sub>c\<^sub>v[bv::=b']\<^sub>c\<^sub>b" using subst_cv_subst_vb_switch by auto + + moreover have "atom z \ v[bv::=b']\<^sub>v\<^sub>b" using x_fresh_b fresh_subst_if subst_b_v_def T_refined_type by metis + + hence "\ z : b | c \[bv::=b']\<^sub>\\<^sub>b[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \ z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \" + using subst_tv.simps subst_tb.simps T_refined_type fresh_Pair by metis + + moreover have "\ z : b[bv::=b']\<^sub>b\<^sub>b | (c[bv::=b']\<^sub>c\<^sub>b)[x::=v[bv::=b']\<^sub>v\<^sub>b]\<^sub>c\<^sub>v \ = \ z : b | c[x::=v]\<^sub>c\<^sub>v \[bv::=b']\<^sub>\\<^sub>b" + using subst_tv.simps subst_tb.simps ceq \.fresh forget_subst[of "bv" b b'] subst_b_b_def T_refined_type by metis + + ultimately show ?case using subst_tv.simps subst_tb.simps ceq T_refined_type by auto +qed + +lemma subst_tb_triple: + assumes "atom bv \ \'" + shows "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \'[bv'::=b']\<^sub>\\<^sub>b[x'::=v']\<^sub>\\<^sub>v[bv::=b]\<^sub>\\<^sub>b" +proof - + have "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>\\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v = \'[bv'::=b']\<^sub>\\<^sub>b[bv::=b]\<^sub>\\<^sub>b [x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>\\<^sub>v" + using subst_tb_commute \atom bv \ \'\ by auto + also have "... = \'[bv'::=b']\<^sub>\\<^sub>b [x'::=v']\<^sub>\\<^sub>v[bv::=b]\<^sub>\\<^sub>b" + using subst_tv_subst_vb_switch by auto + finally show ?thesis using fresh_subst_if forget_subst by auto +qed + +lemma subst_b_infer_e: + fixes s::s and b::b + assumes "\ ; \ ; B ; G; D \ e \ \" and "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" + shows "\ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b \ (e[bv::=b]\<^sub>e\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(nominal_induct avoiding: b rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \ \ v \) + thus ?case using subst_eb.simps infer_e.intros wf_b_subst subst_db.simps wf_b_subst infer_v_wf subst_b_infer_v + by (metis forget_subst ms_fresh_all(1) wfV_b_fresh) +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + show ?case unfolding subst_b_simps subst_eb.simps proof(rule Typing.infer_e_plusI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_plusI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_plusI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_plusI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_plusI subst_b_simps by force + show "atom z3 \ AE_op Plus (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_plusI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_plusI by auto + qed +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_leqI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_leqI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_leqI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_int | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_leqI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_leqI subst_b_simps by force + show "atom z3 \ AE_op LEq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_leqI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_leqI by auto + qed +next + case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_eqI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_eqI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_eqI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : bb[bv::=b]\<^sub>b\<^sub>b | c1[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_eqI subst_b_simps by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : bb[bv::=b]\<^sub>b\<^sub>b | c2[bv::=b]\<^sub>c\<^sub>b \" using subst_b_infer_v infer_e_eqI subst_b_simps by force + show "atom z3 \ AE_op Eq (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using subst_b_simps infer_e_eqI subst_b_fresh_x subst_b_e_def by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_eqI by auto + show "bb[bv::=b]\<^sub>b\<^sub>b \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed +next + case (infer_e_appI \ \ \ \ \ f x b' c \' s' v \) + show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" using wf_b_subst(10) subst_db.simps infer_e_appI wfX_wfY by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_appI by auto + show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \' s'))) = lookup_fun \ f" using infer_e_appI by auto + (*show "\ ; {||} ; (x, b', c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \'" using infer_e_appI sory*) + have "atom bv \ b'" using \\ \\<^sub>w\<^sub>f \ \ infer_e_appI wfPhi_f_supp fresh_def[of "atom bv" b'] by simp + hence "b' = b'[bv::=b]\<^sub>b\<^sub>b" using subst_b_simps + using has_subst_b_class.forget_subst subst_b_b_def by force + moreover have ceq:"c = c[bv::=b]\<^sub>c\<^sub>b" using subst_b_simps proof - + have "supp c \ {atom x}" using infer_e_appI wfPhi_f_simple_supp_c[OF _ \\ \\<^sub>w\<^sub>f \ \] by simp + hence "atom bv \ c" using + fresh_def[of "atom bv" c] + using fresh_def fresh_finsert insert_absorb + insert_subset ms_fresh_all supp_at_base x_not_in_b_set fresh_prodN by metis + thus ?thesis + using forget_subst subst_b_c_def fresh_def[of "atom bv" c] by metis + qed + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ x : b' | c \" + using subst_b_check_v subst_tb.simps subst_vb.simps infer_e_appI + proof - + have "\ ; {|bv|} ; \ \ v \ \ x : b' | c \" + by (metis \\ = {|bv|}\ \\ ; \ ; \ \ v \ \ x : b' | c \\) (* 0.0 ms *) + then show ?thesis + by (metis (no_types) \\ ; {||} \\<^sub>w\<^sub>f b\ \b' = b'[bv::=b]\<^sub>b\<^sub>b\ subst_b_check_v subst_tb.simps ceq) + qed + show "atom x \ (\, \, {||}::bv fset, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" + apply (fresh_mth add: fresh_prodN subst_g_b_x_fresh infer_e_appI ) + using subst_b_fresh_x infer_e_appI apply metis+ + done + have "supp \' \ { atom x }" using wfPhi_f_simple_supp_t infer_e_appI by auto + hence "atom bv \ \'" using fresh_def fresh_at_base by force + then show "\'[x::=v[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \[bv::=b]\<^sub>\\<^sub>b" using infer_e_appI + forget_subst subst_b_\_def subst_tv_subst_vb_switch subst_defs by metis + qed +next + case (infer_e_appPI \' \ \' \ \' b' f' bv' x' b1 c \' s' v' \1) + + have beq: "b1[bv'::=b']\<^sub>b\<^sub>b[bv::=b]\<^sub>b\<^sub>b = b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" + proof - + have "supp b1 \ { atom bv' }" using wfPhi_f_poly_supp_b infer_e_appPI + using supp_at_base by blast + moreover have "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + ultimately have "atom bv \ b1" using fresh_def fresh_at_base by force + thus ?thesis by simp + qed + + have ceq: "(c[bv'::=b']\<^sub>c\<^sub>b)[bv::=b]\<^sub>c\<^sub>b = c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>c\<^sub>b" proof - + have "supp c \ { atom bv', atom x' }" using wfPhi_f_poly_supp_c infer_e_appPI + using supp_at_base by blast + moreover have "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + moreover have "atom x' \ atom bv" by auto + ultimately have "atom bv \ c" using fresh_def[of "atom bv" c] fresh_at_base by auto + thus ?thesis by simp + qed + + show ?case proof(subst subst_eb.simps, rule Typing.infer_e_appPI) + show "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst subst_db.simps infer_e_appPI wfX_wfY by metis + show "\' \\<^sub>w\<^sub>f \' " using infer_e_appPI by auto + show "Some (AF_fundef f' (AF_fun_typ_some bv' (AF_fun_typ x' b1 c \' s'))) = lookup_fun \' f'" using infer_e_appPI by auto + thus "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \ v'[bv::=b]\<^sub>v\<^sub>b \ \ x' : b1[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b | c[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \" + using subst_b_check_v subst_tb.simps subst_b_simps infer_e_appPI + proof - + have "\' ; {||} ; \'[bv::=b]\<^sub>\\<^sub>b \ v'[bv::=b]\<^sub>v\<^sub>b \ \ x' : b1[bv'::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b | (c[bv'::=b']\<^sub>b)[bv::=b]\<^sub>c\<^sub>b \" + using infer_e_appPI subst_b_check_v subst_tb.simps by metis + thus ?thesis using beq ceq subst_defs by metis + qed + show "atom x' \ \'[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_appPI by auto + show "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \1[bv::=b]\<^sub>\\<^sub>b" proof - (* \'[bv'::=b']\<^sub>\\<^sub>b[x'::=v']\<^sub>\\<^sub>v = \1 *) + + have "supp \' \ { atom x', atom bv' }" using wfPhi_f_poly_supp_t infer_e_appPI by auto + moreover hence "bv \ bv'" using infer_e_appPI fresh_def supp_at_base + by (simp add: fresh_def supp_at_base) + ultimately have "atom bv \ \'" using fresh_def by force + hence "\'[bv'::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b[x'::=v'[bv::=b]\<^sub>v\<^sub>b]\<^sub>v = \'[bv'::=b']\<^sub>b[x'::=v']\<^sub>v[bv::=b]\<^sub>\\<^sub>b" using subst_tb_triple subst_defs by auto + thus ?thesis using infer_e_appPI by metis + qed + show "atom bv' \ (\', \', {||}, \'[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v'[bv::=b]\<^sub>v\<^sub>b, \1[bv::=b]\<^sub>\\<^sub>b)" + unfolding fresh_prodN apply( auto simp add: infer_e_appPI fresh_empty_fset) + using fresh_subst_if subst_b_\_def subst_b_\_def subst_b_b_def subst_b_v_def subst_b_\_def infer_e_appPI by metis+ + show "\' ; {||} \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using infer_e_appPI wf_b_subst by simp + qed +next + case (infer_e_fstI \ \ \ \ \ v z' b1 b2 c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_fstI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_fstI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_fstI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_fstI by force + show "atom z \ AE_fst (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_fstI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_fstI by auto + qed +next + case (infer_e_sndI \ \ \ \ \ v z' b1 b2 c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_sndI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_sndI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_pair b1[bv::=b]\<^sub>b\<^sub>b b2[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_sndI by force + show "atom z \ AE_snd (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_sndI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_sndI by auto + qed +next + case (infer_e_lenI \ \ \ \ \ v z' c z) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_lenI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_lenI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_lenI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z' : B_bitvec | c[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_lenI by force + show "atom z \ AE_len (v[bv::=b]\<^sub>v\<^sub>b)" using infer_e_lenI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_lenI by auto + qed +next + case (infer_e_mvarI \ \ \ \ \ u \) + show ?case proof(subst subst subst_eb.simps, rule Typing.infer_e_mvarI) + show "\ ; {||} \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_e_mvarI wf_b_subst by auto + show "\ \\<^sub>w\<^sub>f \ " using infer_e_mvarI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using infer_e_mvarI using wf_b_subst(10) subst_db.simps infer_e_sndI wfX_wfY + by (metis wf_b_subst(15)) + show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" using infer_e_mvarI subst_db.simps set_insert + subst_d_b_member by simp + qed +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_concatI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst(10) subst_db.simps infer_e_concatI wfX_wfY + by (metis wf_b_subst(15)) + show "\ \\<^sub>w\<^sub>f \ " using infer_e_concatI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_bitvec | c2[bv::=b]\<^sub>c\<^sub>b \" + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_concatI by force + show "atom z3 \ AE_concat (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)" using infer_e_concatI subst_b_fresh_x subst_b_v_def e.fresh by metis + show "atom z3 \ \[bv::=b]\<^sub>\\<^sub>b" using subst_g_b_x_fresh infer_e_concatI by auto + qed +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + show ?case unfolding subst_b_simps proof(rule Typing.infer_e_splitI) + show \ \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wf_b_subst(10) subst_db.simps infer_e_splitI wfX_wfY + by (metis wf_b_subst(15)) + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + show \ \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v1[bv::=b]\<^sub>v\<^sub>b \ \ z1 : B_bitvec | c1[bv::=b]\<^sub>c\<^sub>b \\ + using subst_b_infer_v subst_tb.simps subst_b_simps infer_e_splitI by force + show \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND + [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[bv::=b]\<^sub>v\<^sub>b ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using subst_b_check_v subst_tb.simps subst_b_simps infer_e_splitI + proof - + have "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v2[bv::=b]\<^sub>v\<^sub>b \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \[bv::=b]\<^sub>\\<^sub>b" + using infer_e_splitI.hyps(7) infer_e_splitI.prems(1) infer_e_splitI.prems(2) subst_b_check_v by presburger (* 0.0 ms *) + then show ?thesis + by simp (* 0.0 ms *) + qed + show \atom z1 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + show \atom z1 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + + show \atom z2 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + + show \atom z2 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + show \atom z3 \ AE_split (v1[bv::=b]\<^sub>v\<^sub>b) (v2[bv::=b]\<^sub>v\<^sub>b)\ using infer_e_splitI subst_b_fresh_x subst_b_v_def e.fresh by metis + show \atom z3 \ \[bv::=b]\<^sub>\\<^sub>b\ using subst_g_b_x_fresh infer_e_splitI by auto + qed +qed + +text \This is needed for preservation. When we apply a function "f [b] v" we need to +substitute into the body of the function f replacing type-variable with b\ + +lemma subst_b_c_of_forget: + assumes "atom bv \ const" + shows "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x" + using assms proof(nominal_induct const avoiding: x rule:\.strong_induct) + case (T_refined_type x' b' c') + hence "c_of \ x' : b' | c' \ x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by metis + moreover have "atom bv \ c'[x'::=V_var x]\<^sub>c\<^sub>v" proof - + have "atom bv \ c'" using T_refined_type \.fresh by simp + moreover have "atom bv \ V_var x" using v.fresh by simp + ultimately show ?thesis + using T_refined_type \.fresh subst_b_c_def fresh_subst_if + \_fresh_c fresh_subst_cv_if has_subst_b_class.subst_b_fresh_x ms_fresh_all(37) ms_fresh_all assms by metis + qed + ultimately show ?case using forget_subst subst_b_c_def by metis +qed + +lemma subst_b_check_s: + fixes s::s and b::b and cs::branch_s and css::branch_list and v::v and \::\ + assumes "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" (*and "D = []"*) + shows "\ ; \ ; B ; G; D \ s \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b \ (s[bv::=b]\<^sub>s\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" and + "\ ; \ ; B ; G; D ; tid ; cons ; const ; v \ cs \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b ; tid ; cons ; const ; v[bv::=b]\<^sub>v\<^sub>b \ (subst_branchb cs bv b) \ (\[bv::=b]\<^sub>\\<^sub>b)" and + "\ ; \ ; B ; G; D ; tid ; dclist ; v \ css \ \ \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b; D[bv::=b]\<^sub>\\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \ (subst_branchlb css bv b ) \ (\[bv::=b]\<^sub>\\<^sub>b)" + using assms proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + note facts = wfD_emptyI wfX_wfY wf_b_subst subst_b_subtype subst_b_infer_v + case (check_valI \ \ \ \ \ v \' \) + show ?case + apply(subst subst_sb.simps, rule Typing.check_valI) + using facts check_valI apply metis + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype apply blast + using check_valI subst_b_infer_v wf_b_subst subst_b_subtype by metis +next + case (check_letI x \ \ \ \ \ e \ z s b' c) + show ?case proof(subst subst_sb.simps, rule Typing.check_letI) + + show "atom x \(\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" (*using check_letI subst_b_\_def subst_b_\_def subst_b_fresh_x fresh_prod4 subst_b_c_def sory*) + apply(unfold fresh_prodN,auto) + apply(simp add: check_letI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done + show "atom z \ (x, \, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, e[bv::=b]\<^sub>e\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_letI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_letI fresh_prodN)+ done + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ e[bv::=b]\<^sub>e\<^sub>b \ \ z : b'[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \" + using check_letI subst_b_infer_e subst_tb.simps by metis + have "c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v[bv::=b]\<^sub>c\<^sub>b = (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>c\<^sub>v" + using subst_cv_subst_bb_commute[of bv "V_var x" c z b] fresh_at_base by simp + thus "\ ; \ ; {||} ; ((x, b'[bv::=b]\<^sub>b\<^sub>b , (c[bv::=b]\<^sub>c\<^sub>b)[z::=V_var x]\<^sub>v) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b) ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" + using check_letI subst_gb.simps subst_defs by metis + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof(subst subst_sb.simps, rule Typing.check_assertI) + show "atom x \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, s[bv::=b]\<^sub>s\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_assertI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_assertI fresh_prodN)+ done + + have "\ ; \ ; {||} ; ((x, B_bool, c) #\<^sub>\ \)[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using check_assertI + by metis + thus "\ ; \ ; {||} ; (x, B_bool, c[bv::=b]\<^sub>c\<^sub>b) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using subst_gb.simps by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ c[bv::=b]\<^sub>c\<^sub>b " using subst_b_valid check_assertI by simp + show " \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst2(6) check_assertI by simp + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_consI by simp +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case unfolding subst_branchlb.simps using Typing.check_branch_list_finalI by simp +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + + show ?case unfolding subst_b_simps proof(rule Typing.check_branch_s_branchI) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using check_branch_s_branchI wf_b_subst subst_db.simps by metis + show "\\<^sub>w\<^sub>f \ " using check_branch_s_branchI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using check_branch_s_branchI wf_b_subst by metis + + show "atom x \(\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, tid, cons , const, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)" + apply(unfold fresh_prodN,auto) + apply(simp add: check_branch_s_branchI fresh_empty_fset)+ + apply(metis * subst_b_fresh_x check_branch_s_branchI fresh_prodN)+ + done + show wft:"\ ; {||} ; GNil \\<^sub>w\<^sub>f const" using check_branch_s_branchI by auto + hence "(b_of const) = (b_of const)[bv::=b]\<^sub>b\<^sub>b" + using wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_b_def \.supp + bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset + by (metis b_of_supp) + moreover have "(c_of const x)[bv::=b]\<^sub>c\<^sub>b = c_of const x" + using wft wfT_nil_supp fresh_def[of "atom bv" ] forget_subst subst_b_c_def \.supp + bot.extremum_uniqueI ex_in_conv fresh_def supp_empty_fset subst_b_c_of_forget by metis + ultimately show "\ ; \ ; {||} ; (x, b_of const, CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" + using check_branch_s_branchI subst_gb.simps by auto + qed +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case unfolding subst_b_simps proof(rule Typing.check_ifI) + show \atom z \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, s2[bv::=b]\<^sub>s\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)\ + by(unfold fresh_prodN,auto, auto simp add: check_ifI fresh_empty_fset subst_b_fresh_x ) + have "\ z : B_bool | TRUE \[bv::=b]\<^sub>\\<^sub>b = \ z : B_bool | TRUE \" by auto + thus \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z : B_bool | TRUE \\ using check_ifI subst_b_check_v by metis + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ \ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_true) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \\ + using subst_b_if check_ifI by metis + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \ z : b_of \[bv::=b]\<^sub>\\<^sub>b | CE_val (v[bv::=b]\<^sub>v\<^sub>b) == CE_val (V_lit L_false) IMP c_of \[bv::=b]\<^sub>\\<^sub>b z \\ + using subst_b_if check_ifI by metis + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2 ) + show ?case unfolding subst_b_simps proof (rule Typing.check_let2I) + have "atom x \ b" using x_fresh_b by auto + show \atom x \ (\, \, {||}, G[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, t[bv::=b]\<^sub>\\<^sub>b, s1[bv::=b]\<^sub>s\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b)\ + apply(unfold fresh_prodN, auto, auto simp add: check_let2I fresh_prodN fresh_empty_fset) + apply(metis subst_b_fresh_x check_let2I fresh_prodN)+ + done + + show \ \ ; \ ; {||} ; G[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ t[bv::=b]\<^sub>\\<^sub>b \ using check_let2I subst_tb.simps by auto + show \ \ ; \ ; {||} ; (x, b_of t[bv::=b]\<^sub>\\<^sub>b, c_of t[bv::=b]\<^sub>\\<^sub>b x) #\<^sub>\ G[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b\ + using check_let2I subst_tb.simps subst_gb.simps b_of.simps subst_b_c_of subst_b_b_of by auto + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case unfolding subst_b_simps proof(rule Typing.check_varI) + show "atom u \ (\, \, {||}, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, \'[bv::=b]\<^sub>\\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b) " + by(unfold fresh_prodN,auto simp add: check_varI fresh_empty_fset subst_b_fresh_u ) + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \'[bv::=b]\<^sub>\\<^sub>b" using check_varI subst_b_check_v by auto + show "\ ; \ ; {||} ; (subst_gb \ bv b) ; (u, (\'[bv::=b]\<^sub>\\<^sub>b)) #\<^sub>\ (subst_db \ bv b) \ (s[bv::=b]\<^sub>s\<^sub>b) \ (\[bv::=b]\<^sub>\\<^sub>b)" using check_varI by auto + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + show ?case unfolding subst_b_simps proof( rule Typing.check_assignI) + show "\ \\<^sub>w\<^sub>f \ " using check_assignI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wf_b_subst check_assignI by auto + show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_d_b_member by simp + show " \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_b_check_v by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ \ z : B_unit | TRUE \ \ \'[bv::=b]\<^sub>\\<^sub>b" using check_assignI subst_b_subtype subst_b_simps subst_tb.simps by fastforce + qed +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + show ?case unfolding subst_b_simps proof(rule Typing.check_whileI) + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s1[bv::=b]\<^sub>s\<^sub>b \ \ z : B_bool | TRUE \" using check_whileI by auto + show "\ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \ s2[bv::=b]\<^sub>s\<^sub>b \ \ z : B_unit | TRUE \" using check_whileI by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ \ z : B_unit | TRUE \ \ \'[bv::=b]\<^sub>\\<^sub>b" using subst_b_subtype check_whileI by fastforce + qed +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + then show ?case unfolding subst_sb.simps using check_seqI Typing.check_seqI subst_b_eq by metis +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case unfolding subst_b_simps proof(rule Typing.check_caseI) + show \ \ ; \ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b ; tid ; dclist ; v[bv::=b]\<^sub>v\<^sub>b \ subst_branchlb cs bv b \ \[bv::=b]\<^sub>\\<^sub>b\ using check_caseI by auto + show \AF_typedef tid dclist \ set \\ using check_caseI by auto + show \\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ v[bv::=b]\<^sub>v\<^sub>b \ \ z : B_id tid | TRUE \\ using check_caseI subst_b_check_v subst_b_simps subst_tb.simps subst_b_simps + proof - + have "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[bv::=b]\<^sub>\\<^sub>b" using subst_b_eq by auto + then show ?thesis + by (metis (no_types) check_caseI.hyps(4) check_caseI.prems(1) check_caseI.prems(2) subst_b_check_v) (* 31 ms *) + qed + show \ \\<^sub>w\<^sub>f \ \ using check_caseI by auto + qed +qed + +end diff --git a/thys/MiniSail/ContextSubtypingL.thy b/thys/MiniSail/ContextSubtypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/ContextSubtypingL.thy @@ -0,0 +1,1068 @@ +(*<*) +theory ContextSubtypingL + imports TypingL "HOL-Eisbach.Eisbach_Tools" SubstMethods +begin + (*>*) + +declare freshers[simp del] + +chapter \Context Subtyping Lemmas\ + +text \Lemmas allowing us to replace the type of a variable in the context with a subtype +and have the judgement remain valid. Also known as narrowing.\ + +section \Replace or exchange type of variable in a context\ + +text \ Because the G-context is extended by the statements like let, we will need a generalised +substitution lemma for statements. +For this we setup a function that replaces in G (rig) for a particular x the constraint for it. +We also define a well-formedness relation for RIGs that ensures that each new constraint +implies the old one\ + +nominal_function replace_in_g_many :: "\ \ (x*c) list \ \" where + "replace_in_g_many G xcs = List.foldr (\(x,c) G. G[x \ c]) xcs G" + by(auto,simp add: eqvt_def replace_in_g_many_graph_aux_def) +nominal_termination (eqvt) by lexicographic_order + +inductive replace_in_g_subtyped :: "\ \ \ \ \ \ (x*c) list \ \ \ bool" (" _ ; _ \ _ \ _ \ \ _" [100,50,50] 50) where + replace_in_g_subtyped_nilI: "\; \ \ G \ [] \ \ G" +| replace_in_g_subtyped_consI: "\ + Some (b,c') = lookup G x ; + \; \; G \\<^sub>w\<^sub>f c ; + \; \; G[x\c] \ c' ; + \; \ \ G[x\c] \ xcs \ \ G'; x \ fst ` set xcs \ \ + \; \ \ G \ (x,c)#xcs \ \ G'" +equivariance replace_in_g_subtyped +nominal_inductive replace_in_g_subtyped . + +inductive_cases replace_in_g_subtyped_elims[elim!]: + "\; \ \ G \ [] \ \ G'" + "\; \ \ ((x,b,c)#\<^sub>\\ G) \ acs \ \ ((x,b,c)#\<^sub>\G')" + "\; \ \ G' \ (x,c)# acs \ \ G" + +lemma rigs_atom_dom_eq: + assumes "\; \ \ G \ xcs \ \ G'" + shows "atom_dom G = atom_dom G'" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI G) + then show ?case by simp +next + case (replace_in_g_subtyped_consI b c' G x \ \ c xcs G') + then show ?case using rig_dom_eq atom_dom.simps dom.simps by simp +qed + +lemma replace_in_g_wfG: + assumes "\; \ \ G \ xcs \ \ G'" and "wfG \ \ G " + shows "wfG \ \ G'" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c' G x \ c xcs G') + then show ?case using valid_g_wf by auto +qed + +lemma wfD_rig_single: + fixes \::\ and x::x and c::c and G::\ + assumes "\; \; G \\<^sub>w\<^sub>f \ " and "wfG \ \ (G[x\c])" + shows "\; \; G[x\c] \\<^sub>w\<^sub>f \" +proof(cases "atom x \ atom_dom G") + case False + hence "(G[x\c]) = G" using assms replace_in_g_forget wfX_wfY by metis + then show ?thesis using assms by auto +next + case True + then obtain G1 G2 b c' where *: "G=G1@(x,b,c')#\<^sub>\G2" using split_G by fastforce + hence **: "(G[x\c]) = G1@(x,b,c)#\<^sub>\G2" using replace_in_g_inside wfD_wf assms wfD_wf by metis + + hence "wfG \ \ ((x,b,c)#\<^sub>\G2)" using wfG_suffix assms by auto + hence "\; \; (x, b, TRUE) #\<^sub>\ G2 \\<^sub>w\<^sub>f c" using wfG_elim2 by auto + + thus ?thesis using wf_replace_inside1 assms * ** + by (simp add: wf_replace_inside2(6)) +qed + +lemma wfD_rig: + assumes "\; \ \ G \ xcs \ \ G'" and "wfD \ \ G \" + shows "wfD \ \ G' \" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c' G x \ c xcs G') + then show ?case using wfD_rig_single valid.simps wfC_wf by auto +qed + +lemma replace_in_g_fresh: + fixes x::x + assumes "\; \ \ \ \ xcs \ \ \'" and "wfG \ \ \" and "wfG \ \ \'" and "atom x \ \" + shows "atom x \ \'" + using wfG_dom_supp assms fresh_def rigs_atom_dom_eq by metis + +lemma replace_in_g_fresh1: + fixes x::x + assumes "\; \ \ \ \ xcs \ \ \'" and "wfG \ \ \" and "atom x \ \" + shows "atom x \ \'" +proof - + have "wfG \ \ \'" using replace_in_g_wfG assms by auto + thus ?thesis using assms replace_in_g_fresh by metis +qed + +text \ Wellscoping for an eXchange list\ + +inductive wsX:: "\ \ (x*c) list \ bool" where + wsX_NilI: "wsX G []" +| wsX_ConsI: "\ wsX G xcs ; atom x \ atom_dom G ; x \ fst ` set xcs \ \ wsX G ((x,c)#xcs)" +equivariance wsX +nominal_inductive wsX . + +lemma wsX_if1: + assumes "wsX G xcs" + shows "(( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs)" + using assms by(induct rule: wsX.induct,force+ ) + +lemma wsX_if2: + assumes "(( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs)" + shows "wsX G xcs" + using assms proof(induct xcs) + case Nil + then show ?case using wsX_NilI by fast +next + case (Cons a xcs) + then obtain x and c where xc: "a=(x,c)" by force + have " wsX G xcs" proof - + have "distinct (map fst xcs)" using Cons by force + moreover have " atom ` fst ` set xcs \ atom_dom G" using Cons by simp + ultimately show ?thesis using Cons by fast + qed + moreover have "atom x \ atom_dom G" using Cons xc + by simp + moreover have "x \ fst ` set xcs" using Cons xc + by simp + ultimately show ?case using wsX_ConsI xc by blast +qed + +lemma wsX_iff: + "wsX G xcs = ((( atom ` fst ` set xcs) \ atom_dom G) \ List.distinct (List.map fst xcs))" + using wsX_if1 wsX_if2 by meson + +inductive_cases wsX_elims[elim!]: + "wsX G []" + "wsX G ((x,c)#xcs)" + +lemma wsX_cons: + assumes "wsX \ xcs" and "x \ fst ` set xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) ((x, c2) # xcs)" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast + then have "wsX ((x, b, c1) #\<^sub>\ xbc #\<^sub>\ \) xcs" + by (simp add: Un_commute subset_Un_eq wsX_if2) + then show ?case by (simp add: GCons.prems(2) wsX_ConsI) +qed + +lemma wsX_cons2: + assumes "wsX \ xcs" and "x \ fst ` set xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) xcs" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2) +qed + +lemma wsX_cons3: + assumes "wsX \ xcs" + shows "wsX ((x, b, c1) #\<^sub>\ \) xcs" + using assms proof(induct \) + case GNil + then show ?case using atom_dom.simps wsX_iff by auto +next + case (GCons xbc \) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + then have "atom ` fst ` set xcs \ atom_dom (xbc #\<^sub>\ \) \ distinct (map fst xcs)" + using GCons.prems(1) wsX_iff by blast then show ?case by (simp add: Un_commute subset_Un_eq wsX_if2) +qed + +lemma wsX_fresh: + assumes "wsX G xcs" and "atom x \ G" and "wfG \ \ G " + shows "x \ fst ` set xcs" +proof - + have "atom x \ atom_dom G" using assms + using fresh_def wfG_dom_supp by auto + thus ?thesis using wsX_iff assms by blast +qed + +lemma replace_in_g_dist: + assumes "x' \ x" + shows "replace_in_g ((x, b,c) #\<^sub>\ G) x' c'' = ((x, b,c) #\<^sub>\ (replace_in_g G x' c''))" using replace_in_g.simps assms by presburger + +lemma wfG_replace_inside_rig: + fixes c''::c + assumes \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ + shows "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G[x'\c'']" +proof(rule wfG_consI) + + have "wfG \ \ G " using wfG_cons assms by auto + + show *:"\; \ \\<^sub>w\<^sub>f G[x'\c'']" using assms by auto + show "atom x \ G[x'\c'']" using replace_in_g_fresh_single[OF *] assms wfG_elims assms by metis + show **:"\; \ \\<^sub>w\<^sub>f b " using wfG_elim2 assms by auto + show "\; \; (x, b, TRUE) #\<^sub>\ G[x'\c''] \\<^sub>w\<^sub>f c " + proof(cases "atom x' \ atom_dom G") + case True + hence "G = G[x'\c'']" using replace_in_g_forget \wfG \ \ G\ by auto + thus ?thesis using assms wfG_wfC by auto + next + case False + then obtain G1 G2 b' c' where **:"G=G1@(x',b',c')#\<^sub>\G2" + using split_G by fastforce + hence ***: "(G[x'\c'']) = G1@(x',b',c'')#\<^sub>\G2" + using replace_in_g_inside \wfG \ \ G \ by metis + hence "\; \; (x, b, TRUE) #\<^sub>\ G1@(x',b',c')#\<^sub>\G2 \\<^sub>w\<^sub>f c" using * ** assms wfG_wfC by auto + hence "\; \; (x, b, TRUE) #\<^sub>\ G1@(x',b',c'')#\<^sub>\G2 \\<^sub>w\<^sub>f c" using * *** wf_replace_inside assms + by (metis "**" append_g.simps(2) wfG_elim2 wfG_suffix) + thus ?thesis using ** * *** by auto + qed +qed + +lemma replace_in_g_valid_weakening: + assumes "\; \; \[x'\c''] \ c'" and "x' \ x" and "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \[x'\c'']" + shows "\; \; ((x, b,c) #\<^sub>\ \)[x'\ c''] \ c'" + apply(subst replace_in_g_dist,simp add: assms,rule valid_weakening) + using assms by auto+ + +lemma replace_in_g_subtyped_cons: + assumes "replace_in_g_subtyped \ \ G xcs G'" and "wfG \ \ ((x,b,c)#\<^sub>\G)" + shows "x \ fst ` set xcs \ replace_in_g_subtyped \ \ ((x,b,c)#\<^sub>\G) xcs ((x,b,c)#\<^sub>\G')" + using assms proof(induct rule: replace_in_g_subtyped.induct) + case (replace_in_g_subtyped_nilI G) + then show ?case + by (simp add: replace_in_g_subtyped.replace_in_g_subtyped_nilI) +next + case (replace_in_g_subtyped_consI b' c' G x' \ \ c'' xcs' G') + hence "\; \ \\<^sub>w\<^sub>f G[x'\c'']" using valid.simps wfC_wf by auto + + show ?case proof(rule replace_in_g_subtyped.replace_in_g_subtyped_consI) + show " Some (b', c') = lookup ((x, b,c) #\<^sub>\ G) x'" using lookup.simps + fst_conv image_iff \_set_intros surj_pair replace_in_g_subtyped_consI by force + show wbc: " \; \; (x, b, c) #\<^sub>\ G \\<^sub>w\<^sub>f c'' " using wf_weakening \ \; \; G \\<^sub>w\<^sub>f c''\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ by fastforce + have "x' \ x" using replace_in_g_subtyped_consI by auto + have wbc1: "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G[x'\c'']" proof - + have "(x, b, c) #\<^sub>\ G[x'\c''] = ((x, b, c) #\<^sub>\ G)[x'\c'']" using \x' \ x\ using replace_in_g.simps by auto + thus ?thesis using wfG_replace_inside_rig \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ \\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ G \ by fastforce + qed + show *: "\; \; replace_in_g ((x, b,c) #\<^sub>\ G) x' c'' \ c'" + proof - + have "\; \; G[x'\c''] \ c'" using replace_in_g_subtyped_consI by auto + thus ?thesis using replace_in_g_valid_weakening wbc1 \x'\x\ by auto + qed + + show "replace_in_g_subtyped \ \ (replace_in_g ((x, b,c) #\<^sub>\ G) x' c'') xcs' ((x, b,c) #\<^sub>\ G')" + using replace_in_g_subtyped_consI wbc1 by auto + show "x' \ fst ` set xcs'" + using replace_in_g_subtyped_consI by linarith + qed +qed + +lemma replace_in_g_split: + fixes G::\ + assumes "\ = replace_in_g \' x c" and "\' = G'@(x,b,c')#\<^sub>\G" and "wfG \ \ \'" + shows "\ = G'@(x,b,c)#\<^sub>\G" + using assms proof(induct G' arbitrary: G \ \' rule: \_induct) + case GNil + then show ?case by simp +next + case (GCons x1 b1 c1 \1) + hence "x1 \ x" + using wfG_cons_fresh2[of \ \ x1 b1 c1 \1 x b ] + using GCons.prems(2) GCons.prems(3) append_g.simps(2) by auto + moreover hence *: "\; \ \\<^sub>w\<^sub>f (\1 @ (x, b, c') #\<^sub>\ G)" using GCons append_g.simps wfG_elims by metis + moreover hence "replace_in_g (\1 @ (x, b, c') #\<^sub>\ G) x c = \1 @ (x, b, c) #\<^sub>\ G" using GCons replace_in_g_inside[OF *, of c] by auto + + ultimately show ?case using replace_in_g.simps(2)[of x1 b1 c1 " \1 @ (x, b, c') #\<^sub>\ G" x c] GCons + by (simp add: GCons.prems(1) GCons.prems(2)) +qed + +lemma replace_in_g_subtyped_split0: + fixes G::\ + assumes "replace_in_g_subtyped \ \ \'[(x,c)] \" and "\' = G'@(x,b,c')#\<^sub>\G" and "wfG \ \ \'" + shows "\ = G'@(x,b,c)#\<^sub>\G" +proof - + have "\ = replace_in_g \' x c" using assms replace_in_g_subtyped.simps + by (metis Pair_inject list.distinct(1) list.inject) + thus ?thesis using assms replace_in_g_split by blast +qed + +lemma replace_in_g_subtyped_split: + assumes "Some (b, c') = lookup G x" and "\; \; replace_in_g G x c \ c'" and "wfG \ \ G " + shows "\ \ \'. G = \'@(x,b,c')#\<^sub>\\ \ \; \; \'@(x,b,c)#\<^sub>\\ \ c'" +proof - + obtain \ and \' where "G = \'@(x,b,c')#\<^sub>\\" using assms lookup_split by blast + moreover hence "replace_in_g G x c = \'@(x,b,c)#\<^sub>\\" using replace_in_g_split assms by blast + ultimately show ?thesis by (metis assms(2)) +qed + +section \Validity and Subtyping\ + +lemma wfC_replace_in_g: + fixes c::c and c0::c + assumes "\; \; \'@(x,b,c0')#\<^sub>\\ \\<^sub>w\<^sub>f c" and "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c0" + shows "\; \; \' @ (x, b, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f c" + using wf_replace_inside1(2) assms by auto + + +lemma ctx_subtype_valid: + assumes "\; \; \'@(x,b,c0')#\<^sub>\\ \ c" and + "\; \; \'@(x,b,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@(x,b,c0)#\<^sub>\\ \ c" +proof(rule validI) + show "\; \; \' @ (x, b, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f c" proof - + have "\; \; \'@(x,b,c0')#\<^sub>\\ \\<^sub>w\<^sub>f c" using valid.simps assms by auto + moreover have "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c0" proof - + have "wfG \ \ (\'@(x,b,c0)#\<^sub>\\)" using assms valid.simps wfC_wf by auto + hence "wfG \ \ ((x,b,c0)#\<^sub>\\)" using wfG_suffix by auto + thus ?thesis using wfG_wfC by auto + qed + ultimately show ?thesis using assms wfC_replace_in_g by auto + qed + + show "\i. wfI \ (\' @ (x, b, c0) #\<^sub>\ \) i \ is_satis_g i (\' @ (x, b, c0) #\<^sub>\ \) \ is_satis i c" proof(rule,rule) + fix i + assume * : "wfI \ (\' @ (x, b, c0) #\<^sub>\ \) i \ is_satis_g i (\' @ (x, b, c0) #\<^sub>\ \) " + + hence "is_satis_g i (\'@(x, b, c0) #\<^sub>\ \) \ wfI \ (\'@(x, b, c0) #\<^sub>\ \) i" using is_satis_g_append wfI_suffix by metis + moreover hence "is_satis i c0'" using valid.simps assms by presburger + + moreover have "is_satis_g i \'" using is_satis_g_append * by simp + ultimately have "is_satis_g i (\' @ (x, b, c0') #\<^sub>\ \) " using is_satis_g_append by simp + moreover have "wfI \ (\' @ (x, b, c0') #\<^sub>\ \) i" using wfI_def wfI_suffix * wfI_def wfI_replace_inside by metis + ultimately show "is_satis i c" using assms valid.simps by metis + qed +qed + +lemma ctx_subtype_subtype: + fixes \::\ + shows "\; \; G \ t1 \ t2 \ G = \'@(x,b0,c0')#\<^sub>\\ \ \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0' \ \; \; \'@(x,b0,c0)#\<^sub>\\ \ t1 \ t2" +proof(nominal_induct avoiding: c0 rule: subtype.strong_induct) + + case (subtype_baseI x' \ \ \'' z c z' c' b) + let ?\c0 = "\'@(x,b0,c0)#\<^sub>\\" + have wb1: "wfG \ \ ?\c0" using valid.simps wfC_wf subtype_baseI by metis + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ z : b | c \ \ using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ z' : b | c' \ \ using wfT_replace_inside2[OF _ wb1] subtype_baseI by metis + have "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \" using fresh_prodN subtype_baseI fresh_replace_inside wb1 subtype_wf wfX_wfY by metis + thus \atom x' \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, z, c , z' , c' )\ using subtype_baseI fresh_prodN by metis + have "\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v " proof(rule ctx_subtype_valid) + show 1: \\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0') #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v \ + using subtype_baseI append_g.simps subst_defs by metis + have *:"\; \ \\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ " proof(rule wfG_replace_inside2) + show "\; \ \\<^sub>w\<^sub>f ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0') #\<^sub>\ \" + using * valid_wf_all wfC_wf 1 append_g.simps by metis + show "\; \ \\<^sub>w\<^sub>f (x, b0, c0) #\<^sub>\ \" using wfG_suffix wb1 by auto + qed + moreover have "toSet (\' @ (x, b0, c0) #\<^sub>\ \) \ toSet (((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \)" using toSet.simps append_g.simps by auto + ultimately show \\; \; ((x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \') @ (x, b0, c0) #\<^sub>\ \ \ c0' \ using valid_weakening subtype_baseI * by blast + qed + thus \\; \; (x', b, c[z::=V_var x']\<^sub>v) #\<^sub>\ \' @ (x, b0, c0) #\<^sub>\ \ \ c'[z'::=V_var x']\<^sub>v \ using append_g.simps subst_defs by simp + qed +qed + +lemma ctx_subtype_subtype_rig: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and "\; \; \' \ t1 \ t2" + shows "\; \; \ \ t1 \ t2" +proof - + have wf: "wfG \ \ \'" using subtype_g_wf assms by auto + obtain b and c0' where "Some (b, c0') = lookup \' x \ (\; \; replace_in_g \' x c0 \ c0')" using + replace_in_g_subtyped.simps[of \ \ \' "[(x, c0)]" \] assms(1) + + by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair) + moreover then obtain G and G' where *: "\' = G'@(x,b,c0')#\<^sub>\G \ \; \; G'@(x,b,c0)#\<^sub>\G \ c0'" + using replace_in_g_subtyped_split[of b c0' \' x \ \ c0] wf by metis + + ultimately show ?thesis using ctx_subtype_subtype + assms(1) assms(2) replace_in_g_subtyped_split0 subtype_g_wf + by (metis (no_types, lifting) local.wf replace_in_g_split) +qed + +text \ We now prove versions of the @{text "ctx_subtype"} lemmas above using @{text "replace_in_g"}. First we do case where +the replace is just for a single variable (indicated by suffix rig) and then the general case for +multiple replacements (indicated by suffix rigs)\ + +lemma ctx_subtype_subtype_rigs: + assumes "replace_in_g_subtyped \ \ \' xcs \" and "\; \; \' \ t1 \ t2" + shows "\; \; \ \ t1 \ t2" + using assms proof(induct xcs arbitrary: \ \' ) + case Nil + moreover have "\' = \" using replace_in_g_subtyped_nilI + using calculation(1) by blast + ultimately show ?case by auto +next + case (Cons a xcs) + + then obtain x and c where "a=(x,c)" by fastforce + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons + by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + + hence "\; \; (replace_in_g \' x c) \ t1 \ t2" + using ctx_subtype_subtype_rig * assms Cons.prems(2) by auto + + moreover have "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using Cons + using bc by blast + + ultimately show ?case using Cons by blast +qed + +lemma replace_in_g_inside_valid: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and "wfG \ \ \'" + shows "\b c0' G G'. \' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" +proof - + obtain b and c0' where bc: "Some (b, c0') = lookup \' x \ \; \; replace_in_g \' x c0 \ c0'" using + replace_in_g_subtyped.simps[of \ \ \' "[(x, c0)]" \] assms(1) + by (metis fst_conv list.inject list.set_intros(1) list.simps(15) not_Cons_self2 old.prod.exhaust prod.inject set_ConsD surj_pair) + then obtain G and G' where *: "\' = G'@(x,b,c0')#\<^sub>\G \ \; \; G'@(x,b,c0)#\<^sub>\G \ c0'" using replace_in_g_subtyped_split[of b c0' \' x \ \ c0] assms + by metis + thus ?thesis using replace_in_g_inside bc + using assms(1) assms(2) by blast +qed + +lemma replace_in_g_valid: + assumes "\; \ \ G \ xcs \ \ G'" and "\; \; G \ c " + shows \\; \; G' \ c \ + using assms proof(induct rule: replace_in_g_subtyped.inducts) + case (replace_in_g_subtyped_nilI \ \ G) + then show ?case by auto +next + case (replace_in_g_subtyped_consI b c1 G x \ \ c2 xcs G') + hence "\; \; G[x\c2] \ c" + by (metis ctx_subtype_valid replace_in_g_split replace_in_g_subtyped_split valid_g_wf) + then show ?case using replace_in_g_subtyped_consI by auto +qed + +section \Literals\ + +section \Values\ + +lemma lookup_inside_unique_b[simp]: + assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b0,c0)#\<^sub>\\)" and "\ ; B \\<^sub>w\<^sub>f (\'@(x,b0,c0')#\<^sub>\\)" + and "Some (b, c) = lookup (\' @ (x, b0, c0') #\<^sub>\ \) y" and "Some (b0,c0) = lookup (\'@((x,b0,c0))#\<^sub>\\) x" and "x=y" + shows "b = b0" + by (metis assms(2) assms(3) assms(5) lookup_inside_wf old.prod.exhaust option.inject prod.inject) + +lemma ctx_subtype_v_aux: + fixes v::v + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" + using assms proof(nominal_induct "\'@((x,b0,c0')#\<^sub>\\)" v t1 avoiding: c0 rule: infer_v.strong_induct) + case (infer_v_varI \ \ b c xa z) + have wf:\ \; \ \\<^sub>w\<^sub>f \' @ (x, b0, c0) #\<^sub>\ \ \ using wfG_inside_valid2 infer_v_varI by metis + have xf1:\atom z \ xa\ using infer_v_varI by metis + have xf2: \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)\ apply( fresh_mth add: infer_v_varI ) + using fresh_def infer_v_varI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show ?case proof (cases "x=xa") + case True + moreover have "b = b0" using infer_v_varI True by simp + moreover hence \Some (b, c0) = lookup (\' @ (x, b0, c0) #\<^sub>\ \) xa\ using lookup_inside_wf[OF wf] infer_v_varI True by auto + ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by metis + next + case False + moreover hence \Some (b, c) = lookup (\' @ (x, b0, c0) #\<^sub>\ \) xa\ using lookup_inside2 infer_v_varI by metis + ultimately show ?thesis using wf xf1 xf2 Typing.infer_v_varI by simp + qed +next + case (infer_v_litI \ \ l \) + thus ?case using Typing.infer_v_litI wfG_inside_valid2 by simp +next + case (infer_v_pairI z v1 v2 \ \ t1' t2' c0) + show ?case proof + show "atom z \ (v1, v2)" using infer_v_pairI fresh_Pair by simp + show "atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)" apply( fresh_mth add: infer_v_pairI ) + using fresh_def infer_v_pairI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ t1'" using infer_v_pairI by simp + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ t2'" using infer_v_pairI by simp + qed +next + case (infer_v_consI s dclist \ dc tc \ v tv z) + show ?case proof + show \AF_typedef s dclist \ set \\ using infer_v_consI by auto + show \(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ tv\ using infer_v_consI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ tv \ tc\ using infer_v_consI ctx_subtype_subtype by auto + show \atom z \ v\ using infer_v_consI by auto + show \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \)\ apply( fresh_mth add: infer_v_consI ) + using fresh_def infer_v_consI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + qed +next + case (infer_v_conspI s bv dclist \ dc tc \ v tv b z) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ tv\ using infer_v_conspI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ tv \ tc[bv::=b]\<^sub>\\<^sub>b\ using infer_v_conspI ctx_subtype_subtype by auto + show \atom z \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, v, b)\ apply( fresh_mth add: infer_v_conspI ) + using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show \atom bv \ (\, \, \' @ (x, b0, c0) #\<^sub>\ \, v, b)\ apply( fresh_mth add: infer_v_conspI ) + using fresh_def infer_v_conspI wfG_supp fresh_append_g fresh_GCons fresh_prodN by metis+ + show \ \; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + qed +qed + +lemma ctx_subtype_v: + fixes v::v + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\t2. \; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2 \ \; \; \'@((x,b0,c0)#\<^sub>\\) \ t2 \ t1" +proof - + + have "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1 " using ctx_subtype_v_aux assms by auto + moreover hence "\; \; \'@((x,b0,c0)#\<^sub>\\) \ t1 \ t1" using subtype_reflI2 infer_v_wf by simp + ultimately show ?thesis by auto +qed + +lemma ctx_subtype_v_eq: + fixes v::v + assumes + "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and + " \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" +proof - + obtain t1' where "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1'" using ctx_subtype_v assms by metis + moreover have "replace_in_g (\'@((x,b0,c0')#\<^sub>\\)) x c0 = \'@((x,b0,c0)#\<^sub>\\)" using replace_in_g_inside infer_v_wf assms by metis + ultimately show ?thesis using infer_v_uniqueness_rig assms by metis +qed + +lemma ctx_subtype_check_v_eq: + assumes "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t1" and " \; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t1" +proof - + obtain t2 where t2: "\; \; \'@((x,b0,c0')#\<^sub>\\) \ v \ t2 \ \; \; \'@((x,b0,c0')#\<^sub>\\) \ t2 \ t1" + using check_v_elims assms by blast + hence t3: "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2" + using assms ctx_subtype_v_eq by blast + + have "\; \; \'@((x,b0,c0)#\<^sub>\\) \ v \ t2" using t3 by auto + moreover have " \; \; \'@((x,b0,c0)#\<^sub>\\) \ t2 \ t1" proof - + + have " \; \; \'@((x,b0,c0')#\<^sub>\\) \ t2 \ t1" using t2 by auto + thus ?thesis using subtype_trans + using assms(2) ctx_subtype_subtype by blast + qed + ultimately show ?thesis using check_v.intros by presburger +qed + +text \ Basically the same as @{text "ctx_subtype_v_eq"} but in a different form\ +lemma ctx_subtype_v_rig_eq: + fixes v::v + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" +proof - + obtain b and c0' and G and G' where "\' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" + using assms replace_in_g_inside_valid infer_v_wf by metis + thus ?thesis using ctx_subtype_v_eq[of \ \ G' x b c0' G v t1 c0] assms by simp +qed + +lemma ctx_subtype_v_rigs_eq: + fixes v::v + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" + using assms proof(induct xcs arbitrary: \ \' t1 ) + case Nil + then show ?case by auto +next + case (Cons a xcs) + then obtain x and c where "a=(x,c)" by fastforce + + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " + using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + hence t2: "\; \; (replace_in_g \' x c) \ v \ t1 " using ctx_subtype_v_rig_eq[OF * Cons(3)] by blast + moreover have **: "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using bc by auto + ultimately have t2': "\; \; \ \ v \ t1" using Cons by blast + thus ?case by blast +qed + +lemma ctx_subtype_check_v_rigs_eq: + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\; \; \' \ v \ t1" + shows "\; \; \ \ v \ t1" +proof - + obtain t2 where "\; \; \' \ v \ t2 \ \; \; \' \ t2 \ t1" using check_v_elims assms by fast + hence "\; \; \ \ v \ t2 \ \; \; \ \ t2 \ t1" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs + using assms(1) by blast + thus ?thesis + using check_v_subtypeI by blast +qed + +section \Expressions\ + +lemma valid_wfC: + fixes c0::c + assumes "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\; \; (x, b0, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c0" + using wfG_elim2 valid.simps wfG_suffix + using assms valid_g_wf by metis + +lemma ctx_subtype_e_eq: + fixes G::\ + assumes + "\ ; \ ; \ ; G ; \ \ e \ t1" and "G = \'@((x,b0,c0')#\<^sub>\\)" + "\; \; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "\ ; \ ; \ ; \'@((x,b0,c0)#\<^sub>\\) ; \ \ e \ t1" + using assms proof(nominal_induct t1 avoiding: c0 rule: infer_e.strong_induct) + case (infer_e_valI \ \ \'' \ \ v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_valI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_valI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \\ using infer_e_valI ctx_subtype_v_eq by auto + qed +next + case (infer_e_plusI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_plusI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_plusI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_int | c1 \\ using infer_e_plusI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_int | c2 \\ using infer_e_plusI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op Plus v1 v2\ using infer_e_plusI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_plusI fresh_replace_inside infer_v_wf by metis + qed +next + case (infer_e_leqI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_leqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_leqI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_int | c1 \\ using infer_e_leqI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_int | c2 \\ using infer_e_leqI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op LEq v1 v2\ using infer_e_leqI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_leqI fresh_replace_inside infer_v_wf by metis + qed +next + case (infer_e_eqI \ \ \'' \ \ v1 z1 bb c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_eqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_eqI by auto + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : bb | c1 \\ using infer_e_eqI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : bb | c2 \\ using infer_e_eqI ctx_subtype_v_eq by auto + show \atom z3 \ AE_op Eq v1 v2\ using infer_e_eqI by auto + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using * infer_e_eqI fresh_replace_inside infer_v_wf by metis + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed +next + case (infer_e_appI \ \ \'' \ \ f x' b c \' s' v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_appI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \' s'))) = lookup_fun \ f\ using infer_e_appI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ x' : b | c \\ using infer_e_appI ctx_subtype_check_v_eq by auto + thus \atom x' \ (\, \, \, \' @ (x, b0, c0) #\<^sub>\ \, \, v, \)\ + using infer_e_appI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 x'] infer_v_wf by auto + show \\'[x'::=v]\<^sub>v = \\ using infer_e_appI by auto + qed +next + case (infer_e_appPI \ \ \1 \ \ b' f bv x1 b c \' s' v \) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_appPI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using infer_e_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b c \' s'))) = lookup_fun \ f\ using infer_e_appPI by auto + show \\; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ x1 : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \\ using infer_e_appPI ctx_subtype_check_v_eq subst_defs by auto + thus \atom x1 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 x1 ] infer_v_wf infer_e_appPI by auto + show \\'[bv::=b']\<^sub>b[x1::=v]\<^sub>v = \\ using infer_e_appPI by auto + have "atom bv \ \' @ (x, b0, c0') #\<^sub>\ \" using infer_e_appPI by metis + hence "atom bv \ \' @ (x, b0, c0) #\<^sub>\ \" + unfolding fresh_append_g fresh_GCons fresh_prod3 using \atom bv \ c0 \ fresh_append_g by metis + thus \atom bv \ (\, \, \, \' @ (x, b0, c0) #\<^sub>\ \, \, b', v, \)\ using infer_e_appPI by auto + qed +next + case (infer_e_fstI \ \ \'' \ \ v z' b1 b2 c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_fstI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_fstI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_pair b1 b2 | c \\ using infer_e_fstI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_fstI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_fst v\ using infer_e_fstI by auto + qed +next + case (infer_e_sndI \ \ \'' \ \ v z' b1 b2 c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_sndI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_sndI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_pair b1 b2 | c \\ using infer_e_sndI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_sndI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_snd v\ using infer_e_sndI by auto + qed +next + case (infer_e_lenI \ \ \'' \ \ v z' c z) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_lenI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_lenI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v \ \ z' : B_bitvec | c \\ using infer_e_lenI ctx_subtype_v_eq by auto + thus \atom z \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_lenI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z] infer_v_wf by auto + show \atom z \ AE_len v\ using infer_e_lenI by auto + qed +next + case (infer_e_mvarI \ \ \'' \ \ u \) + show ?case proof + show "\; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \" using wf_replace_inside2(6) valid_wfC infer_e_mvarI by auto + thus "\; \ \\<^sub>w\<^sub>f \' @ (x, b0, c0) #\<^sub>\ \" using infer_e_mvarI fresh_replace_inside wfD_wf by blast + show "\ \\<^sub>w\<^sub>f \ " using infer_e_mvarI by auto + show "(u, \) \ setD \" using infer_e_mvarI by auto + qed +next + case (infer_e_concatI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + show ?case proof + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_concatI by auto + thus \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using infer_e_concatI fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z3] infer_v_wf wfX_wfY by metis + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_concatI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_bitvec | c1 \\ using infer_e_concatI ctx_subtype_v_eq by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v2 \ \ z2 : B_bitvec | c2 \\ using infer_e_concatI ctx_subtype_v_eq by auto + show \atom z3 \ AE_concat v1 v2\ using infer_e_concatI by auto + qed +next + case (infer_e_splitI \ \ \'' \ \ v1 z1 c1 v2 z2 z3) + show ?case proof + show *:\ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wf_replace_inside2(6) valid_wfC infer_e_splitI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + show \ \; \; \' @ (x, b0, c0) #\<^sub>\ \ \ v1 \ \ z1 : B_bitvec | c1 \\ using infer_e_splitI ctx_subtype_v_eq by auto + show \\; \; \' @ + (x, b0, c0) #\<^sub>\ + \ \ v2 \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND + [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using infer_e_splitI ctx_subtype_check_v_eq by auto + + show \atom z1 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 z1] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z2 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z3 \ \' @ (x, b0, c0) #\<^sub>\ \\ using fresh_replace_inside[of \ \ \' x b0 c0' \ c0 ] infer_e_splitI infer_v_wf wfX_wfY * by metis + show \atom z1 \ AE_split v1 v2\ using infer_e_splitI by auto + show \atom z2 \ AE_split v1 v2\ using infer_e_splitI by auto + show \atom z3 \ AE_split v1 v2\ using infer_e_splitI by auto + qed +qed + +lemma ctx_subtype_e_rig_eq: + assumes "replace_in_g_subtyped \ \ \' [(x,c0)] \" and + "\ ; \ ; \ ; \' ; \ \ e \ t1" + shows "\ ; \ ; \ ; \ ; \ \ e \ t1" +proof - + obtain b and c0' and G and G' where "\' = G' @ (x,b,c0')#\<^sub>\G \ \ = G' @ (x,b,c0)#\<^sub>\G \ \; \; G'@ (x,b,c0)#\<^sub>\G \ c0'" + using assms replace_in_g_inside_valid infer_e_wf by meson + thus ?thesis + using assms ctx_subtype_e_eq by presburger +qed + +lemma ctx_subtype_e_rigs_eq: + assumes "replace_in_g_subtyped \ \ \' xcs \" and + "\ ; \ ; \ ; \'; \ \ e \ t1" + shows "\ ; \ ; \ ; \ ; \ \ e \ t1" + using assms proof(induct xcs arbitrary: \ \' t1 ) + case Nil + moreover have "\' = \" using replace_in_g_subtyped_nilI + using calculation(1) by blast + moreover have "\;\;\ \ t1 \ t1" using subtype_reflI2 Nil infer_e_t_wf by blast + ultimately show ?case by blast +next + case (Cons a xcs) + + then obtain x and c where "a=(x,c)" by fastforce + then obtain b and c' where bc: "Some (b, c') = lookup \' x \ + replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \ \ \; \; \' \\<^sub>w\<^sub>f c \ + x \ fst ` set xcs \ \; \; (replace_in_g \' x c) \ c' " using replace_in_g_subtyped_elims(3)[of \ \ \' x c xcs \] Cons + by (metis valid.simps) + + hence *: "replace_in_g_subtyped \ \ \' [(x,c)] (replace_in_g \' x c)" using replace_in_g_subtyped_consI + by (meson image_iff list.distinct(1) list.set_cases replace_in_g_subtyped_nilI) + hence t2: "\ ; \ ; \ ; (replace_in_g \' x c) ; \ \ e \ t1 " using ctx_subtype_e_rig_eq[OF * Cons(3)] by blast + moreover have **: "replace_in_g_subtyped \ \ (replace_in_g \' x c) xcs \" using bc by auto + ultimately have t2': "\ ; \ ; \ ; \ ; \ \ e \ t1" using Cons by blast + thus ?case by blast +qed + +section \Statements\ + +lemma ctx_subtype_s_rigs: + fixes c0::c and s::s and G'::\ and xcs :: "(x*c) list" and css::branch_list + shows + "check_s \ \ \ G \ s t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_s \ \ \ G' \ s t1" and + "check_branch_s \ \ \ G \ tid cons const v cs t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_branch_s \ \ \ G' \ tid cons const v cs t1" + "check_branch_list \ \ \ G \ tid dclist v css t1 \ wsX G xcs \ replace_in_g_subtyped \ \ G xcs G' \ check_branch_list \ \ \ G' \ tid dclist v css t1" +proof(induction arbitrary: xcs G' and xcs G' and xcs G' rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \) + hence *:"\; \; G' \ v \ \' \ \; \; G' \ \' \ \" using ctx_subtype_v_rigs_eq ctx_subtype_subtype_rigs + by (meson check_v.simps) + show ?case proof + show \\; \; G' \\<^sub>w\<^sub>f \\ using check_valI wfD_rig by auto + show \\ \\<^sub>w\<^sub>f \ \ using check_valI by auto + show \\; \; G' \ v \ \'\ using * by auto + show \\; \; G' \ \' \ \\ using * by auto + qed +next + case (check_letI x \ \ \ \ \ e \ z' s b' c') + + show ?case proof + have wfG: "\; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f G'" using infer_e_wf check_letI replace_in_g_wfG using infer_e_wf(2) by (auto simp add: freshers) + hence "atom x \ G'" using check_letI replace_in_g_fresh replace_in_g_wfG by auto + thus "atom x \ (\, \, \, G', \, e, \)" using check_letI by auto + have "atom z' \ G'" apply(rule replace_in_g_fresh[OF check_letI(7)]) + using replace_in_g_wfG check_letI fresh_prodN infer_e_wf by metis+ + thus "atom z' \ (x, \, \, \, G', \, e, \, s)" using check_letI fresh_prodN by metis + + show "\ ; \ ; \ ; G' ; \ \ e \ \ z' : b' | c' \" + using check_letI ctx_subtype_e_rigs_eq by blast + show "\ ; \ ; \ ; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G' ; \ \ s \ \" + proof(rule check_letI(5)) + have vld: "\;\;((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) \ c'[z'::=V_var x]\<^sub>c\<^sub>v" proof - + have "wfG \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)" using check_letI check_s_wf by metis + hence "wfC \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) (c'[z'::=V_var x]\<^sub>c\<^sub>v)" using wfC_refl subst_defs by auto + thus ?thesis using valid_reflI[of \ \ x b' "c'[z'::=V_var x]\<^sub>v" \ " c'[z'::=V_var x]\<^sub>v"] subst_defs by auto + qed + have xf: "x \ fst ` set xcs" proof - + have "atom ` fst ` set xcs \ atom_dom \" using check_letI wsX_iff by meson + moreover have "wfG \ \ \" using infer_e_wf check_letI by metis + ultimately show ?thesis using fresh_def check_letI wfG_dom_supp + using wsX_fresh by auto + qed + show "replace_in_g_subtyped \ \ ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) ((x, c'[z'::=V_var x]\<^sub>v) # xcs) ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G')" proof - + + have "Some (b', c'[z'::=V_var x]\<^sub>v) = lookup ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x" by auto + + moreover have "\; \; replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) \ c'[z'::=V_var x]\<^sub>v" proof - + have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) = ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)" + using replace_in_g.simps by presburger + thus ?thesis using vld subst_defs by auto + qed + + moreover have " replace_in_g_subtyped \ \ (replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G'))" proof - + have "wfG \ \ ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \))" using check_letI check_s_wf by metis + hence "replace_in_g_subtyped \ \ ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \)) xcs ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ G'))" + using check_letI replace_in_g_subtyped_cons xf by meson + moreover have "replace_in_g ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>v) = ( ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \))" + using replace_in_g.simps by presburger + ultimately show ?thesis by argo + qed + moreover have "\; \; (x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>v " using vld subst_defs by auto + ultimately show ?thesis using replace_in_g_subtyped_consI xf replace_in_g.simps(2) by metis + qed + + show " wsX ((x, b', c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \) ((x, c'[z'::=V_var x]\<^sub>v) # xcs)" + using check_letI xf subst_defs by (simp add: wsX_cons) + qed + qed + +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using Typing.check_branch_list_consI by auto +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using Typing.check_branch_list_finalI by auto +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + + have wfcons: "wfG \ \ ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" using check_s_wf check_branch_s_branchI + by meson + hence wf: "wfG \ \ \" using wfG_cons by metis + + moreover have "atom x \ (const, G',v)" proof - + have "atom x \ G'" using check_branch_s_branchI wf replace_in_g_fresh + wfG_dom_supp replace_in_g_wfG by simp + thus ?thesis using check_branch_s_branchI fresh_prodN by simp + qed + + moreover have st: "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ G' ; \ \ s \ \ " proof - + have " wsX ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) xcs" using check_branch_s_branchI wsX_cons2 wsX_fresh wf by force + moreover have "replace_in_g_subtyped \ \ ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) xcs ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ G' )" + using replace_in_g_subtyped_cons wsX_fresh wf check_branch_s_branchI wfcons by auto + thus ?thesis using check_branch_s_branchI calculation by meson + qed + moreover have wft: " wfT \ \ G' \ " using + check_branch_s_branchI ctx_subtype_subtype_rigs subtype_reflI2 subtype_wf by metis + moreover have "wfD \ \ G' \" using check_branch_s_branchI wfD_rig by presburger + ultimately show ?case using + Typing.check_branch_s_branchI + using check_branch_s_branchI.hyps by simp + +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + hence wf:"wfG \ \ \" using check_s_wf by presburger + show ?case proof(rule check_s_check_branch_s_check_branch_list.check_ifI) + show \atom z \ (\, \, \, G', \, v, s1, s2, \)\ using fresh_prodN replace_in_g_fresh1 wf check_ifI by auto + show \\; \; G' \ v \ \ z : B_bool | TRUE \\ using ctx_subtype_check_v_rigs_eq check_ifI by presburger + show \ \ ; \ ; \ ; G' ; \ \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; G' ; \ \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + + case (check_let2I x P \ \ G \ t s1 \ s2 ) + show ?case proof + have "wfG P \ G" using check_let2I check_s_wf by metis + show *: "P ; \ ; \ ; G' ; \ \ s1 \t" using check_let2I by blast + show "atom x \ (P, \, \, G', \, t, s1, \)" proof - + have "wfG P \ G'" using check_s_wf * by blast + hence "atom_dom G = atom_dom G'" using check_let2I rigs_atom_dom_eq by presburger + moreover have "atom x \ G" using check_let2I by auto + moreover have "wfG P \ G" using check_s_wf * replace_in_g_wfG check_let2I by simp + ultimately have "atom x \ G'" using wfG_dom_supp fresh_def \wfG P \ G'\ by metis + thus ?thesis using check_let2I by auto + qed + show "P ; \ ; \ ;(x, b_of t, c_of t x) #\<^sub>\ G' ; \ \ s2 \ \ " proof - + have "wsX ((x, b_of t, c_of t x) #\<^sub>\ G) xcs" using check_let2I wsX_cons2 wsX_fresh \wfG P \ G\ by simp + moreover have "replace_in_g_subtyped P \ ((x, b_of t, c_of t x) #\<^sub>\ G) xcs ((x, b_of t, c_of t x) #\<^sub>\ G')" proof(rule replace_in_g_subtyped_cons ) + show "replace_in_g_subtyped P \ G xcs G'" using check_let2I by auto + have "atom x \ G" using check_let2I by auto + moreover have "wfT P \ G t" using check_let2I check_s_wf by metis + + moreover have "atom x \ t" using check_let2I check_s_wf wfT_supp by auto + ultimately show "wfG P \ ((x, b_of t, c_of t x) #\<^sub>\ G)" using wfT_wf_cons b_of_c_of_eq[of x t] by auto + show "x \ fst ` set xcs" using check_let2I wsX_fresh \wfG P \ G\ by simp + qed + ultimately show ?thesis using check_let2I by presburger + qed + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + have "atom u \ G'" unfolding fresh_def + apply(rule u_not_in_g , rule replace_in_g_wfG) + using check_v_wf check_varI by simp+ + thus \atom u \ (\, \, \, G', \, \', v, \)\ unfolding fresh_prodN using check_varI by simp + show \\; \; G' \ v \ \'\ using ctx_subtype_check_v_rigs_eq check_varI by auto + show \ \ ; \ ; \ ; G' ; (u, \') #\<^sub>\ \ \ s \ \\ using check_varI by auto + qed +next + case (check_assignI P \ \ G \ u \ v z \') + show ?case proof + show \P \\<^sub>w\<^sub>f \ \ using check_assignI by auto + show \P ; \ ; G' \\<^sub>w\<^sub>f \ \ using check_assignI wfD_rig by auto + show \(u, \) \ setD \\ using check_assignI by auto + show \P ; \ ; G' \ v \ \\ using ctx_subtype_check_v_rigs_eq check_assignI by auto + show \P ; \ ; G' \ \ z : B_unit | TRUE \ \ \'\ using ctx_subtype_subtype_rigs check_assignI by auto + qed +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using Typing.check_whileI + by (meson ctx_subtype_subtype_rigs) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case + using check_s_check_branch_s_check_branch_list.check_seqI by blast +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case proof + show " \ ; \ ; \ ; G' ; \ ; tid ; dclist ; v \ cs \ \" using check_caseI ctx_subtype_check_v_rigs_eq by auto + show "AF_typedef tid dclist \ set \" using check_caseI by auto + show "\; \; G' \ v \ \ z : B_id tid | TRUE \" using check_caseI ctx_subtype_check_v_rigs_eq by auto + show "\\<^sub>w\<^sub>f \ " using check_caseI by auto + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + have wfG: "\; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f G'" using check_s_wf check_assertI replace_in_g_wfG wfX_wfY by metis + hence "atom x \ G'" using check_assertI replace_in_g_fresh replace_in_g_wfG by auto + thus \atom x \ (\, \, \, G', \, c, \, s)\ using check_assertI fresh_prodN by auto + show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ G' ; \ \ s \ \\ proof(rule check_assertI(5) ) + show "wsX ((x, B_bool, c) #\<^sub>\ \) xcs" using check_assertI wsX_cons3 by simp + show "\; \ \ (x, B_bool, c) #\<^sub>\ \ \ xcs \ \ (x, B_bool, c) #\<^sub>\ G'" proof(rule replace_in_g_subtyped_cons) + show \ \; \ \ \ \ xcs \ \ G'\ using check_assertI by auto + show \ \; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \ \ using check_assertI check_s_wf by metis + thus \x \ fst ` set xcs\ using check_assertI wsX_fresh wfG_elims wfX_wfY by metis + qed + qed + show \\; \; G' \ c \ using check_assertI replace_in_g_valid by auto + show \ \; \; G' \\<^sub>w\<^sub>f \ \ using check_assertI wfD_rig by auto + qed +qed + +lemma replace_in_g_subtyped_empty: + assumes "wfG \ \ (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + shows "replace_in_g_subtyped \ \ (replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" +proof - + have "replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) = (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using replace_in_g.simps by auto + next + case (GCons x1 b1 c1 \1) + have "x \ fst ` toSet ((x1,b1,c1)#\<^sub>\\1)" using GCons wfG_inside_fresh atom_dom.simps dom.simps toSet.simps append_g.simps by fast + hence "x1 \ x" using assms wfG_inside_fresh GCons by force + hence "((x1,b1,c1) #\<^sub>\ (\1 @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \))[x\c'[z'::=V_var x]\<^sub>c\<^sub>v] = (x1,b1,c1) #\<^sub>\ (\1 @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using replace_in_g.simps GCons wfG_elims append_g.simps by metis + thus ?case using append_g.simps by simp + qed + thus ?thesis using replace_in_g_subtyped_nilI by presburger +qed + +lemma ctx_subtype_s: + fixes s::s + assumes "\ ; \ ; \ ; \'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\) ; \ \ s \ \" and + "\; \; \ \ \ z' : b | c' \ \ \ z : b | c \" and + "atom x \ (z,z',c,c')" + shows "\ ; \ ; \ ; \'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\ ; \ \ s \ \" +proof - + + have wf: "wfG \ \ (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\))" using check_s_wf assms by meson + hence *:"x \ fst ` toSet \'" using wfG_inside_fresh by force + have "wfG \ \ ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)" using wf wfG_suffix by metis + hence xfg: "atom x \ \" using wfG_elims by metis + have "x \ z'" using assms fresh_at_base fresh_prod4 by metis + hence a2: "atom x \ c'" using assms fresh_prod4 by metis + + have "atom x \ (z', c', z, c, \)" proof - + have "x \ z" using assms using assms fresh_at_base fresh_prod4 by metis + hence a1 : "atom x \ c" using assms subtype_wf subtype_wf assms wfT_fresh_c xfg by meson + thus ?thesis using a1 a2 \atom x \ (z,z',c,c')\ fresh_prod4 fresh_Pair xfg by simp + qed + hence wc1:" \; \; (x, b, c'[z'::=V_var x]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var x]\<^sub>v" + using subtype_valid assms fresh_prodN by metis + + have vld: "\;\ ; (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ c[z::=V_var x]\<^sub>c\<^sub>v" proof - + + have "toSet ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ toSet (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" by auto + moreover have "wfG \ \ (\'@(x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" proof - + have *:"wfT \ \ \ (\ z' : b | c' \)" using subtype_wf assms by meson + moreover have "atom x \ (c',\)" using xfg a2 by simp + ultimately have "wfG \ \ ((x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" using wfT_wf_cons_flip freshers by blast + thus ?thesis using wfG_replace_inside2 check_s_wf assms by metis + qed + ultimately show ?thesis using wc1 valid_weakening subst_defs by metis + qed + hence wbc: "\; \; \' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using valid.simps by auto + have wbc1: "\; \; (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v" using wc1 valid.simps subst_defs by auto + have "wsX (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)]" proof + show "wsX (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) []" using wsX_NilI by auto + show "atom x \ atom_dom (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" by simp + show "x \ fst ` set []" by auto + qed + moreover have "replace_in_g_subtyped \ \ (\'@((x,b,c[z::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)) [(x, c'[z'::=V_var x]\<^sub>c\<^sub>v)] (\'@(x,b,c'[z'::=V_var x]\<^sub>c\<^sub>v)#\<^sub>\\)" proof + show "Some (b, c[z::=V_var x]\<^sub>c\<^sub>v) = lookup (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x" using lookup_inside* by auto + show "\; \; replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v) \ c[z::=V_var x]\<^sub>c\<^sub>v" using vld replace_in_g_split wf by metis + show "replace_in_g_subtyped \ \ (replace_in_g (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) x (c'[z'::=V_var x]\<^sub>c\<^sub>v)) [] (\' @ (x, b, c'[z'::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + using replace_in_g_subtyped_empty wf by presburger + show "x \ fst ` set []" by auto + show "\; \; \' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=V_var x]\<^sub>c\<^sub>v" + proof(rule wf_weakening) + show \\; \; (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c'[z'::=[ x ]\<^sup>v]\<^sub>c\<^sub>v \ using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis + show \\; \ \\<^sub>w\<^sub>f \' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using wfC_cons_switch[OF wbc1] wf_weakening(6) check_s_wf assms toSet.simps by metis + show \toSet ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ toSet (\' @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)\ using append_g.simps toSet.simps by auto + qed + qed + ultimately show ?thesis using ctx_subtype_s_rigs(1)[OF assms(1)] by presburger +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/Examples.thy b/thys/MiniSail/Examples.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Examples.thy @@ -0,0 +1,162 @@ +(*<*) +theory Examples +imports SubstMethods fuzzyrule.fuzzyrule +begin +(*>*) + +subsection \Examples from thesis\ + +text \ +let x = false in +let y = 0b11 in +let z = 42 in +let w = (true,(z,y)) in () +\ + +method fresh_mth_aux uses add = ( + (match conclusion in "atom z \ GNil" for z \ \ simp add: fresh_GNil[of "atom z"] add\) + | (match conclusion in "atom z \ DNil" for z \ \ simp add: fresh_DNil[of "atom z" ] add\) + | (match conclusion in "atom z \ Nil" for z \ \ simp add: fresh_Nil[of "atom z" ] add\) + | (match conclusion in "atom z \ {||}" for z \ \ simp add: fresh_empty_fset[of "atom z" ] add\) + | (match conclusion in "atom (z::x) \ ((x::x,b::b,c::c)#\<^sub>\G)" for z x b c G \ \simp add: fresh_GCons[of "atom z" "(x,b,c)" G] add\) + | (match conclusion in "atom z \ (b::b)" for z b \ \ simp add: b.fresh[of "atom z" ] add\) + | (match conclusion in "atom z \ (c::c)" for z c \ \ simp add: c.fresh[of "atom z" ] add\) + | (match conclusion in "atom z \ (i::int)" for z i \ \ simp add: pure_fresh add\) +(* tbc delta and types + | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *)*) +) + +method fresh_mth_basic uses add = ( + (unfold fresh_prodN, intro conjI)?, + (fresh_mth_aux add: add)+) + +method wf_ctx uses add = ( + (rule wfPhi_emptyI | rule wfTh_emptyI| rule wfD_emptyI| rule wfG_nilI | rule wfG_cons2I )+ +) + +method wfb uses add = ( + (rule wfB_boolI | rule wfB_intI | rule wfB_unitI | rule wfB_pairI )+ +) + +method wfc uses add = ( + (rule wfC_trueI | rule wfC_falseI | rule wfC_notI | rule wfC_eqI )+ +) + +method wfv uses add = ( + (rule wfV_varI | fuzzy_rule wfV_litI | rule wfV_pairI )+ +) + +method wfce uses add = ( + (rule wfCE_valI | rule wfCE_fstI | rule wfCE_sndI | rule wfCE_eqI )+ +) + +method wfe uses add = ( + (rule wfE_valI | rule wfE_fstI | rule wfE_sndI )+ +) + +method wfs uses add = ( + (rule wfS_letI | rule wfS_valI )+ +) + +method wf_all uses add = ( + (wfb | wfv | wfc | wfce | wfe | wfs | auto simp add: add | wf_ctx | fresh_mth_basic add: add )+ +) + +method wf_all2 uses add = ( + (wfb | wfv | wfc | wfce | wfe | wfs | wf_ctx )+ +) + +method wf_let uses add = ( + (rule wfS_letI, rule wfE_valI, wf_ctx add: add ), + (rule wfV_litI,rule wfG_nilI, rule wfTh_emptyI)) + +(* +rule wfS_letI, rule wfE_valI, rule wfPhi_emptyI , rule wfTh_emptyI, rule wfD_emptyI, rule wfG_nilI, rule wfTh_emptyI), + (rule wfV_litI,rule wfG_nilI, rule wfTh_emptyI) +*) + +lemma + shows "[] ; [] ; {||} ; GNil ; []\<^sub>\ \\<^sub>w\<^sub>f LET x = [ [ L_false ]\<^sup>v ]\<^sup>e IN [[ L_unit ]\<^sup>v]\<^sup>s : B_unit" + apply wf_let + apply(rule wfS_valI,wf_ctx) + by(fuzzy_rule wfV_litI, wf_all) + +lemma + assumes " atom z \ x" + shows "[] ; [] ; {||} ; GNil ; []\<^sub>\ \\<^sub>w\<^sub>f LET x = [ [ L_false ]\<^sup>v ]\<^sup>e IN + LET z = [ [ L_num 42]\<^sup>v ]\<^sup>e IN [[ L_unit ]\<^sup>v]\<^sup>s : B_unit" + + + apply wf_let + apply(rule wfS_letI) + + apply(rule wfE_valI, wf_ctx,auto,wf_ctx) + by(wf_all add: assms) + + +text \ +val not_bool : (x : bool [ T ]) -> { z : bool | ~ (z = x) } +function not_bool(x) = if x then F else T + +val eq_int : ( x : int x int [ T ]) -> { z : bool | z = (fst x = snd x ) } +function eq_int ( x ) = + let x1 = fst x in + let x2 = snd x in + let x3 = x1 = x2 in x3 + +val neq_int : ( ( x : int x int [ T ]) -> { z : bool | ~ ( z = (fst x = snd x )) } +function neq_int (x) = not_bool(eq_int(x)) +\ + + +lemma + assumes "atom z \ x" + shows "wfFT [] [] {||} (AF_fun_typ (x::x) B_bool C_true (\ z : B_bool | C_not (C_eq (CE_val (V_var z)) (CE_val (V_var x))) \) + (AS_if (V_var x) [[L_false]\<^sup>v]\<^sup>s [[L_true]\<^sup>v]\<^sup>s))" + apply(rule wfFTI,wf_all) + apply (simp add: supp_at_base)+ + apply(rule wfTI) + by(wf_all add: assms) + +lemma + assumes "atom z \ x" + shows "wfFT [] [] {||} (AF_fun_typ (x::x) (B_pair B_int B_int) C_true (\ z : B_bool | (C_eq (CE_val (V_var z)) (CE_op Eq (CE_fst (CE_val (V_var x))) (CE_fst (CE_val (V_var x))))) \) + (AS_let x1 (AE_fst ((V_var x))) ( + AS_let x2 (AE_snd ((V_var x))) ( + AS_let x3 (AE_op Eq ((V_var x1)) ((V_var x2))) (AS_val (V_var x3))))))" + + apply(rule wfFTI,wf_all2) + defer 1 + + + apply simp + apply(rule wfTI) + apply(fresh_mth_basic ) + defer 1 + apply(wf_all ) + using assms apply simp + apply(wf_all ) + using assms apply simp + apply(wf_all ) + using assms apply simp + apply(wf_all ) + defer 1 + apply(wf_all ) + using assms apply simp +(*,wf_ctx,rule wfC_eqI,rule wfCE_valI,rule wfV_varI, (wf_ctx|auto|wf1)+)*) + apply(fresh_mth_basic ) + apply((wf_ctx|auto|wf1)+) + apply(rule wfCE_eqI,rule wfCE_fstI, rule wfCE_valI,rule wfV_varI) + apply((wf_ctx|auto|wf1)+) + apply(fresh_mth_basic add: assms)+ + apply((wf_ctx|auto|wf1)+) + using assms apply simp + apply(rule wfCE_fstI ,rule wfCE_valI,rule wfV_varI) + apply((wf_ctx|auto|wf1)+) + apply(fresh_mth_basic add: assms)+ + apply((wf_ctx|auto|wf1)+) + using assms apply simp + by((wf_ctx|auto|wf1)+) + + +end \ No newline at end of file diff --git a/thys/MiniSail/IVSubst.thy b/thys/MiniSail/IVSubst.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/IVSubst.thy @@ -0,0 +1,1446 @@ +(*<*) +theory IVSubst + imports Syntax +begin + (*>*) + +chapter \Immutable Variable Substitution\ + +text \Substitution involving immutable variables. We define a class and instances for all +of the term forms\ + +section \Class\ + +class has_subst_v = fs + + fixes subst_v :: "'a::fs \ x \ v \ 'a::fs" ("_[_::=_]\<^sub>v" [1000,50,50] 1000) + assumes fresh_subst_v_if: "y \ (subst_v a x v) \ (atom x \ a \ y \ a) \ (y \ v \ (y \ a \ y = atom x))" + and forget_subst_v[simp]: "atom x \ a \ subst_v a x v = a" + and subst_v_id[simp]: "subst_v a x (V_var x) = a" + and eqvt[simp,eqvt]: "(p::perm) \ (subst_v a x v) = (subst_v (p \ a) (p \x) (p \v))" + and flip_subst_v[simp]: "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + and subst_v_simple_commute[simp]: "atom x \ c \(c[z::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = c[z::=b]\<^sub>v" +begin + +lemma subst_v_flip_eq_one: + fixes z1::x and z2::x and x1::x and x2::x + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + and "atom x1 \ (z1,z2,c1,c2)" + shows "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (c2[z2::=[x1]\<^sup>v]\<^sub>v)" +proof - + have "(c1[z1::=[x1]\<^sup>v]\<^sub>v) = (x1 \ z1) \ c1" using assms flip_subst_v by auto + moreover have "(c2[z2::=[x1]\<^sup>v]\<^sub>v) = (x1 \ z2) \ c2" using assms flip_subst_v by auto + ultimately show ?thesis using Abs1_eq_iff_all(3)[of z1 c1 z2 c2 z1] assms + by (metis Abs1_eq_iff_fresh(3) flip_commute) +qed + +lemma subst_v_flip_eq_two: + fixes z1::x and z2::x and x1::x and x2::x + assumes "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + shows "(c1[z1::=b]\<^sub>v) = (c2[z2::=b]\<^sub>v)" +proof - + obtain x::x where *:"atom x \ (z1,z2,c1,c2)" using obtain_fresh by metis + hence "(c1[z1::=[x]\<^sup>v]\<^sub>v) = (c2[z2::=[x]\<^sup>v]\<^sub>v)" using subst_v_flip_eq_one[OF assms, of x] by metis + hence "(c1[z1::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v = (c2[z2::=[x]\<^sup>v]\<^sub>v)[x::=b]\<^sub>v" by auto + thus ?thesis using subst_v_simple_commute * fresh_prod4 by metis +qed + +lemma subst_v_flip_eq_three: + assumes "[[atom z1]]lst. c1 = [[atom z1']]lst. c1'" and "atom x \ c1" and "atom x' \ (x,z1,z1', c1, c1')" + shows "(x \ x') \ (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1'[z1'::=[x']\<^sup>v]\<^sub>v" +proof - + have "atom x' \ c1[z1::=[x]\<^sup>v]\<^sub>v" using assms fresh_subst_v_if by simp + hence "(x \ x') \ (c1[z1::=[x]\<^sup>v]\<^sub>v) = c1[z1::=[x]\<^sup>v]\<^sub>v[x::=[x']\<^sup>v]\<^sub>v" using flip_subst_v[of x' "c1[z1::=[x]\<^sup>v]\<^sub>v" x] flip_commute by metis + also have "... = c1[z1::=[x']\<^sup>v]\<^sub>v" using subst_v_simple_commute fresh_prod4 assms by auto + also have "... = c1'[z1'::=[x']\<^sup>v]\<^sub>v" using subst_v_flip_eq_one[of z1 c1 z1' c1' x'] using assms by auto + finally show ?thesis by auto +qed + +end + +section \Values\ + +nominal_function + subst_vv :: "v \ x \ v \ v" where + "subst_vv (V_lit l) x v = V_lit l" +| "subst_vv (V_var y) x v = (if x = y then v else V_var y)" +| "subst_vv (V_cons tyid c v') x v = V_cons tyid c (subst_vv v' x v)" +| "subst_vv (V_consp tyid c b v') x v = V_consp tyid c b (subst_vv v' x v)" +| "subst_vv (V_pair v1 v2) x v = V_pair (subst_vv v1 x v ) (subst_vv v2 x v )" + by(auto simp: eqvt_def subst_vv_graph_aux_def, metis v.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_vv_abbrev :: "v \ x \ v \ v" ("_[_::=_]\<^sub>v\<^sub>v" [1000,50,50] 1000) + where + "v[x::=v']\<^sub>v\<^sub>v \ subst_vv v x v'" + +lemma fresh_subst_vv_if [simp]: + "j \ t[i::=x]\<^sub>v\<^sub>v = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using supp_l_empty apply (induct t rule: v.induct,auto simp add: subst_vv.simps fresh_def, auto) + by (simp add: supp_at_base |metis b.supp supp_b_empty )+ + +lemma forget_subst_vv [simp]: "atom a \ tm \ tm[a::=x]\<^sub>v\<^sub>v = tm" + by (induct tm rule: v.induct) (simp_all add: fresh_at_base) + +lemma subst_vv_id [simp]: "tm[a::=V_var a]\<^sub>v\<^sub>v = tm" + by (induct tm rule: v.induct) simp_all + +lemma subst_vv_commute [simp]: + "atom j \ tm \ tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>v\<^sub>v " + by (induct tm rule: v.induct) (auto simp: fresh_Pair) + +lemma subst_vv_commute_full [simp]: + "atom j \ t \ atom i \ u \ i \ j \ tm[i::=t]\<^sub>v\<^sub>v[j::=u]\<^sub>v\<^sub>v = tm[j::=u]\<^sub>v\<^sub>v[i::=t]\<^sub>v\<^sub>v" + by (induct tm rule: v.induct) auto + +lemma subst_vv_var_flip[simp]: + fixes v::v + assumes "atom y \ v" + shows "(y \ x) \ v = v [x::=V_var y]\<^sub>v\<^sub>v" + using assms apply(induct v rule:v.induct) + apply auto + using l.fresh l.perm_simps l.strong_exhaust supp_l_empty permute_pure permute_list.simps fresh_def flip_fresh_fresh apply fastforce + using permute_pure apply blast+ + done + +instantiation v :: has_subst_v +begin + +definition + "subst_v = subst_vv" + +instance proof + fix j::atom and i::x and x::v and t::v + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_vv_if[of j t i x] subst_v_v_def by metis + + fix a::x and tm::v and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_vv subst_v_v_def by simp + + fix a::x and tm::v + show "subst_v tm a (V_var a) = tm" using subst_vv_id subst_v_v_def by simp + + fix p::perm and x1::x and v::v and t1::v + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_v_v_def by simp + + fix x::x and c::v and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_v_v_def by simp + + fix x::x and c::v and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_v_def by simp +qed + +end + +section \Expressions\ + +nominal_function subst_ev :: "e \ x \ v \ e" where + "subst_ev ( (AE_val v') ) x v = ( (AE_val (subst_vv v' x v)) )" +| "subst_ev ( (AE_app f v') ) x v = ( (AE_app f (subst_vv v' x v )) )" +| "subst_ev ( (AE_appP f b v') ) x v = ( (AE_appP f b (subst_vv v' x v )) )" +| "subst_ev ( (AE_op opp v1 v2) ) x v = ( (AE_op opp (subst_vv v1 x v ) (subst_vv v2 x v )) )" +| "subst_ev [#1 v']\<^sup>e x v = [#1 (subst_vv v' x v )]\<^sup>e" +| "subst_ev [#2 v']\<^sup>e x v = [#2 (subst_vv v' x v )]\<^sup>e" +| "subst_ev ( (AE_mvar u)) x v = AE_mvar u" +| "subst_ev [| v' |]\<^sup>e x v = [| (subst_vv v' x v ) |]\<^sup>e" +| "subst_ev ( AE_concat v1 v2) x v = AE_concat (subst_vv v1 x v ) (subst_vv v2 x v )" +| "subst_ev ( AE_split v1 v2) x v = AE_split (subst_vv v1 x v ) (subst_vv v2 x v )" + by(simp add: eqvt_def subst_ev_graph_aux_def,auto)(meson e.strong_exhaust) + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_ev_abbrev :: "e \ x \ v \ e" ("_[_::=_]\<^sub>e\<^sub>v" [1000,50,50] 500) + where + "e[x::=v']\<^sub>e\<^sub>v \ subst_ev e x v' " + +lemma size_subst_ev [simp]: "size ( subst_ev A i x) = size A" + apply (nominal_induct A avoiding: i x rule: e.strong_induct) + by auto + +lemma forget_subst_ev [simp]: "atom a \ A \ subst_ev A a x = A" + apply (nominal_induct A avoiding: a x rule: e.strong_induct) + by (auto simp: fresh_at_base) + +lemma subst_ev_id [simp]: "subst_ev A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: e.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_ev_if [simp]: + "j \ (subst_ev A i x ) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" + apply (induct A rule: e.induct) + unfolding subst_ev.simps fresh_subst_vv_if apply auto+ + using pure_fresh fresh_opp_all apply metis+ + done + +lemma subst_ev_commute [simp]: + "atom j \ A \ (A[i::=t]\<^sub>e\<^sub>v)[j::=u]\<^sub>e\<^sub>v = A[i::=t[j::=u]\<^sub>v\<^sub>v]\<^sub>e\<^sub>v" + by (nominal_induct A avoiding: i j t u rule: e.strong_induct) (auto simp: fresh_at_base) + +lemma subst_ev_var_flip[simp]: + fixes e::e and y::x and x::x + assumes "atom y \ e" + shows "(y \ x) \ e = e [x::=V_var y]\<^sub>e\<^sub>v" + using assms apply(nominal_induct e rule:e.strong_induct) + apply (simp add: subst_v_v_def) + apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip) + apply (metis (mono_tags, lifting) b.eq_iff b.perm_simps e.fresh e.perm_simps flip_b_id subst_ev.simps subst_vv_var_flip) + subgoal for x + apply (rule_tac y=x in opp.strong_exhaust) + using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+ + using subst_vv_var_flip flip_def by (simp add: flip_def permute_pure)+ + +lemma subst_ev_flip: + fixes e::e and ea::e and c::x + assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" + shows "e[x::=v']\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>e\<^sub>v" +proof - + have "e[x::=v']\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>e\<^sub>v)[c::=v']\<^sub>e\<^sub>v" using subst_ev_commute assms by simp + also have "... = ((c \ x) \ e)[c::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + also have "... = ((c \ xa) \ ea)[c::=v']\<^sub>e\<^sub>v" using assms flip_commute by metis + also have "... = ea[xa::=v']\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + finally show ?thesis by auto +qed + +lemma subst_ev_var[simp]: + "(AE_val (V_var x))[x::=[z]\<^sup>v]\<^sub>e\<^sub>v = AE_val (V_var z)" + by auto + +instantiation e :: has_subst_v +begin + +definition + "subst_v = subst_ev" + +instance proof + fix j::atom and i::x and x::v and t::e + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_ev_if[of j t i x] subst_v_e_def by metis + + fix a::x and tm::e and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_ev subst_v_e_def by simp + + fix a::x and tm::e + show "subst_v tm a (V_var a) = tm" using subst_ev_id subst_v_e_def by simp + + fix p::perm and x1::x and v::v and t1::e + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_ev_commute subst_v_e_def by simp + + fix x::x and c::e and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_v_e_def by simp + + fix x::x and c::e and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_e_def by simp +qed +end + +lemma subst_ev_commute_full: + fixes e::e and w::v and v::v + assumes "atom z \ v" and "atom x \ w" and "x \ z" + shows "subst_ev (e[z::=w]\<^sub>e\<^sub>v) x v = subst_ev (e[x::=v]\<^sub>e\<^sub>v) z w" + using assms by(nominal_induct e rule: e.strong_induct,simp+) + +lemma subst_ev_v_flip1[simp]: + fixes e::e + assumes "atom z1 \ (z,e)" and "atom z1' \ (z,e)" + shows"(z1 \ z1') \ e[z::=v]\<^sub>e\<^sub>v = e[z::= ((z1 \ z1') \ v)]\<^sub>e\<^sub>v" + using assms proof(nominal_induct e rule:e.strong_induct) +qed (simp add: flip_def fresh_Pair swap_fresh_fresh)+ + +section \Expressions in Constraints\ + +nominal_function subst_cev :: "ce \ x \ v \ ce" where + "subst_cev ( (CE_val v') ) x v = ( (CE_val (subst_vv v' x v )) )" +| "subst_cev ( (CE_op opp v1 v2) ) x v = ( (CE_op opp (subst_cev v1 x v ) (subst_cev v2 x v )) )" +| "subst_cev ( (CE_fst v')) x v = CE_fst (subst_cev v' x v )" +| "subst_cev ( (CE_snd v')) x v = CE_snd (subst_cev v' x v )" +| "subst_cev ( (CE_len v')) x v = CE_len (subst_cev v' x v )" +| "subst_cev ( CE_concat v1 v2) x v = CE_concat (subst_cev v1 x v ) (subst_cev v2 x v )" + apply (simp add: eqvt_def subst_cev_graph_aux_def,auto) + by (meson ce.strong_exhaust) + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_cev_abbrev :: "ce \ x \ v \ ce" ("_[_::=_]\<^sub>c\<^sub>e\<^sub>v" [1000,50,50] 500) + where + "e[x::=v']\<^sub>c\<^sub>e\<^sub>v \ subst_cev e x v'" + +lemma size_subst_cev [simp]: "size ( subst_cev A i x ) = size A" + by (nominal_induct A avoiding: i x rule: ce.strong_induct,auto) + +lemma forget_subst_cev [simp]: "atom a \ A \ subst_cev A a x = A" + by (nominal_induct A avoiding: a x rule: ce.strong_induct, auto simp: fresh_at_base) + +lemma subst_cev_id [simp]: "subst_cev A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: ce.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_cev_if [simp]: + "j \ (subst_cev A i x ) = ((atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i)))" +proof(nominal_induct A avoiding: i x rule: ce.strong_induct) + case (CE_op opp v1 v2) + then show ?case using fresh_subst_vv_if subst_ev.simps e.supp pure_fresh opp.fresh + fresh_e_opp + using fresh_opp_all by auto +qed(auto)+ + +lemma subst_cev_commute [simp]: + "atom j \ A \ (subst_cev (subst_cev A i t ) j u) = subst_cev A i (subst_vv t j u )" + by (nominal_induct A avoiding: i j t u rule: ce.strong_induct) (auto simp: fresh_at_base) + +lemma subst_cev_var_flip[simp]: + fixes e::ce and y::x and x::x + assumes "atom y \ e" + shows "(y \ x) \ e = e [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v" + using assms proof(nominal_induct e rule:ce.strong_induct) + case (CE_val v) + then show ?case using subst_vv_var_flip by auto +next + case (CE_op opp v1 v2) + hence yf: "atom y \ v1 \ atom y \ v2" using ce.fresh by blast + have " (y \ x) \ (CE_op opp v1 v2 ) = CE_op ((y \ x) \ opp) ( (y \ x) \ v1 ) ( (y \ x) \ v2)" + using opp.perm_simps ce.perm_simps permute_pure ce.fresh opp.strong_exhaust by presburger + also have "... = CE_op ((y \ x) \ opp) (v1[x::=V_var y]\<^sub>c\<^sub>e\<^sub>v) (v2 [x::=V_var y]\<^sub>c\<^sub>e\<^sub>v)" using yf + by (simp add: CE_op.hyps(1) CE_op.hyps(2)) + finally show ?case using subst_cev.simps opp.perm_simps opp.strong_exhaust + by (metis (full_types)) +qed( (auto simp add: permute_pure subst_vv_var_flip)+) + +lemma subst_cev_flip: + fixes e::ce and ea::ce and c::x + assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" + shows "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" +proof - + have "e[x::=v']\<^sub>c\<^sub>e\<^sub>v = (e[x::=V_var c]\<^sub>c\<^sub>e\<^sub>v)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_commute assms by simp + also have "... = ((c \ x) \ e)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + also have "... = ((c \ xa) \ ea)[c::=v']\<^sub>c\<^sub>e\<^sub>v" using assms flip_commute by metis + also have "... = ea[xa::=v']\<^sub>c\<^sub>e\<^sub>v" using subst_ev_var_flip assms by simp + finally show ?thesis by auto +qed + +lemma subst_cev_var[simp]: + fixes z::x and x::x + shows "[[x]\<^sup>v]\<^sup>c\<^sup>e [x::=[z]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v = [[z]\<^sup>v]\<^sup>c\<^sup>e" + by auto + +instantiation ce :: has_subst_v +begin + +definition + "subst_v = subst_cev" + +instance proof + fix j::atom and i::x and x::v and t::ce + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_cev_if[of j t i x] subst_v_ce_def by metis + + fix a::x and tm::ce and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_cev subst_v_ce_def by simp + + fix a::x and tm::ce + show "subst_v tm a (V_var a) = tm" using subst_cev_id subst_v_ce_def by simp + + fix p::perm and x1::x and v::v and t1::ce + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_cev_commute subst_v_ce_def by simp + + fix x::x and c::ce and z::x + show "atom x \ c \ ((x \ z) \ c) = c [z::=V_var x]\<^sub>v" + using subst_v_ce_def by simp + + fix x::x and c::ce and z::x + show "atom x \ c \ c [z::=V_var x]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_v_ce_def by simp +qed + +end + +lemma subst_cev_commute_full: + fixes e::ce and w::v and v::v + assumes "atom z \ v" and "atom x \ w" and "x \ z" + shows "subst_cev (e[z::=w]\<^sub>c\<^sub>e\<^sub>v) x v = subst_cev (e[x::=v]\<^sub>c\<^sub>e\<^sub>v) z w " + using assms by(nominal_induct e rule: ce.strong_induct,simp+) + + +lemma subst_cev_v_flip1[simp]: + fixes e::ce + assumes "atom z1 \ (z,e)" and "atom z1' \ (z,e)" + shows"(z1 \ z1') \ e[z::=v]\<^sub>c\<^sub>e\<^sub>v = e[z::= ((z1 \ z1') \ v)]\<^sub>c\<^sub>e\<^sub>v" + using assms apply(nominal_induct e rule:ce.strong_induct) + by (simp add: flip_def fresh_Pair swap_fresh_fresh)+ + +section \Constraints\ + +nominal_function subst_cv :: "c \ x \ v \ c" where + "subst_cv (C_true) x v = C_true" +| "subst_cv (C_false) x v = C_false" +| "subst_cv (C_conj c1 c2) x v = C_conj (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (C_disj c1 c2) x v = C_disj (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (C_imp c1 c2) x v = C_imp (subst_cv c1 x v ) (subst_cv c2 x v )" +| "subst_cv (e1 == e2) x v = ((subst_cev e1 x v ) == (subst_cev e2 x v ))" +| "subst_cv (C_not c) x v = C_not (subst_cv c x v )" + apply (simp add: eqvt_def subst_cv_graph_aux_def,auto) + using c.strong_exhaust by metis +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_cv_abbrev :: "c \ x \ v \ c" ("_[_::=_]\<^sub>c\<^sub>v" [1000,50,50] 1000) + where + "c[x::=v']\<^sub>c\<^sub>v \ subst_cv c x v'" + +lemma size_subst_cv [simp]: "size ( subst_cv A i x ) = size A" + by (nominal_induct A avoiding: i x rule: c.strong_induct,auto) + +lemma forget_subst_cv [simp]: "atom a \ A \ subst_cv A a x = A" + by (nominal_induct A avoiding: a x rule: c.strong_induct, auto simp: fresh_at_base) + +lemma subst_cv_id [simp]: "subst_cv A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: c.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_cv_if [simp]: + "j \ (subst_cv A i x ) \ (atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i))" + by (nominal_induct A avoiding: i x rule: c.strong_induct, (auto simp add: pure_fresh)+) + +lemma subst_cv_commute [simp]: + "atom j \ A \ (subst_cv (subst_cv A i t ) j u ) = subst_cv A i (subst_vv t j u )" + by (nominal_induct A avoiding: i j t u rule: c.strong_induct) (auto simp: fresh_at_base) + +lemma let_s_size [simp]: "size s \ size (AS_let x e s)" + apply (nominal_induct s avoiding: e x rule: s_branch_s_branch_list.strong_induct(1)) + apply auto + done + +lemma subst_cv_var_flip[simp]: + fixes c::c + assumes "atom y \ c" + shows "(y \ x) \ c = c[x::=V_var y]\<^sub>c\<^sub>v" + using assms by(nominal_induct c rule:c.strong_induct,(simp add: flip_subst_v subst_v_ce_def)+) + +instantiation c :: has_subst_v +begin + +definition + "subst_v = subst_cv" + +instance proof + fix j::atom and i::x and x::v and t::c + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_cv_if[of j t i x] subst_v_c_def by metis + + fix a::x and tm::c and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_cv subst_v_c_def by simp + + fix a::x and tm::c + show "subst_v tm a (V_var a) = tm" using subst_cv_id subst_v_c_def by simp + + fix p::perm and x1::x and v::v and t1::c + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_cv_commute subst_v_c_def by simp + + fix x::x and c::c and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_cv_var_flip subst_v_c_def by simp + + fix x::x and c::c and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_cv_var_flip subst_v_c_def by simp +qed + +end + +lemma subst_cv_var_flip1[simp]: + fixes c::c + assumes "atom y \ c" + shows "(x \ y) \ c = c[x::=V_var y]\<^sub>c\<^sub>v" + using subst_cv_var_flip flip_commute + by (metis assms) + +lemma subst_cv_v_flip3[simp]: + fixes c::c + assumes "atom z1 \ c" and "atom z1' \ c" + shows"(z1 \ z1') \ c[z::=[z1]\<^sup>v]\<^sub>c\<^sub>v = c[z::=[z1']\<^sup>v]\<^sub>c\<^sub>v" +proof - + consider "z1' = z" | "z1 = z" | "atom z1 \ z \ atom z1' \ z" by force + then show ?thesis proof(cases) + case 1 + then show ?thesis using 1 assms by auto + next + case 2 + then show ?thesis using 2 assms by auto + next + case 3 + then show ?thesis using assms by auto + qed +qed + +lemma subst_cv_v_flip[simp]: + fixes c::c + assumes "atom x \ c" + shows "((x \ z) \ c)[x::=v]\<^sub>c\<^sub>v = c [z::=v]\<^sub>c\<^sub>v" + using assms subst_v_c_def by auto + +lemma subst_cv_commute_full: + fixes c::c + assumes "atom z \ v" and "atom x \ w" and "x\z" + shows "(c[z::=w]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v = (c[x::=v]\<^sub>c\<^sub>v)[z::=w]\<^sub>c\<^sub>v" + using assms proof(nominal_induct c rule: c.strong_induct) + case (C_eq e1 e2) + then show ?case using subst_cev_commute_full by simp +qed(force+) + +lemma subst_cv_eq[simp]: + assumes "atom z1 \ e1" + shows "(CE_val (V_var z1) == e1 )[z1::=[x]\<^sup>v]\<^sub>c\<^sub>v = (CE_val (V_var x) == e1 )" (is "?A = ?B") +proof - + have "?A = (((CE_val (V_var z1))[z1::=[x]\<^sup>v]\<^sub>c\<^sub>e\<^sub>v) == e1)" using subst_cv.simps assms by simp + thus ?thesis by simp +qed + +section \Variable Context\ + +text \The idea of this substitution is to remove x from the context. We really want to add the condition +that x is fresh in v but this causes problems with proofs.\ + +nominal_function subst_gv :: "\ \ x \ v \ \" where + "subst_gv GNil x v = GNil" +| "subst_gv ((y,b,c) #\<^sub>\ \) x v = (if x = y then \ else ((y,b,c[x::=v]\<^sub>c\<^sub>v)#\<^sub>\ (subst_gv \ x v )))" +proof(goal_cases) + case 1 + then show ?case by(simp add: eqvt_def subst_gv_graph_aux_def ) +next + case (3 P x) + then show ?case by (metis neq_GNil_conv prod_cases3) +qed(fast+) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_gv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "g[x::=v]\<^sub>\\<^sub>v \ subst_gv g x v" + +lemma size_subst_gv [simp]: "size ( subst_gv G i x ) \ size G" + by (induct G,auto) + +lemma forget_subst_gv [simp]: "atom a \ G \ subst_gv G a x = G" + apply (induct G ,auto) + using fresh_GCons fresh_PairD(1) not_self_fresh apply blast + apply (simp add: fresh_GCons)+ + done + +lemma fresh_subst_gv: "atom a \ G \ atom a \ v \ atom a \ subst_gv G x v" +proof(induct G) + case GNil + then show ?case by auto +next + case (GCons xbc G) + obtain x' and b' and c' where xbc: "xbc = (x',b',c')" using prod_cases3 by blast + show ?case proof(cases "x=x'") + case True + have "atom a \ G" using GCons fresh_GCons by blast + thus ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc True by presburger + next + case False + then show ?thesis using subst_gv.simps(2)[of x' b' c' G] GCons xbc False fresh_GCons by simp + qed +qed + +lemma subst_gv_flip: + fixes x::x and xa::x and z::x and c::c and b::b and \::\ + assumes "atom xa \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ \" and "atom x \ \" and "atom x \ (z, c)" and "atom xa \ (z, c)" + shows "(x \ xa) \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) = (xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \" +proof - + have "(x \ xa) \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) = (( (x \ xa) \ x, b, (x \ xa) \ c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" + using subst Cons_eqvt flip_fresh_fresh using G_cons_flip by simp + also have "... = ((xa, b, (x \ xa) \ c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" using assms by fastforce + also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ ((x \ xa) \ \))" using assms subst_cv_var_flip by fastforce + also have "... = ((xa, b, c[z::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \)" using assms flip_fresh_fresh by blast + finally show ?thesis by simp +qed + +section \Types\ + +nominal_function subst_tv :: "\ \ x \ v \ \" where + "atom z \ (x,v) \ subst_tv \ z : b | c \ x v = \ z : b | c[x::=v]\<^sub>c\<^sub>v \" + apply (simp add: eqvt_def subst_tv_graph_aux_def ) + apply auto + subgoal for P a aa b + apply(rule_tac y=a and c="(aa,b)" in \.strong_exhaust) + by (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) +proof - + fix z :: x and c :: c and za :: x and xa :: x and va :: v and ca :: c and cb :: x + assume a1: "atom za \ va" and a2: "atom z \ va" and a3: "\cb. atom cb \ c \ atom cb \ ca \ cb \ z \ cb \ za \ c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" + assume a4: "atom cb \ c" and a5: "atom cb \ ca" and a6: "cb \ z" and a7: "cb \ za" and "atom cb \ va" and a8: "za \ xa" and a9: "z \ xa" + assume a10:"cb \ xa" + note assms = a10 a9 a8 a7 a6 a5 a4 a3 a2 a1 + + have "c[z::=V_var cb]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v" using assms by auto + hence "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by simp + moreover have "c[z::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v" using subst_cv_commute_full[of z va xa "V_var cb" ] assms fresh_def v.supp by fastforce + moreover have "ca[za::=V_var cb]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" + using subst_cv_commute_full[of za va xa "V_var cb" ] assms fresh_def v.supp by fastforce + + ultimately show "c[xa::=va]\<^sub>c\<^sub>v[z::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=va]\<^sub>c\<^sub>v[za::=V_var cb]\<^sub>c\<^sub>v" by simp +qed + +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_tv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "t[x::=v]\<^sub>\\<^sub>v \ subst_tv t x v" + +lemma size_subst_tv [simp]: "size ( subst_tv A i x ) = size A" +proof (nominal_induct A avoiding: i x rule: \.strong_induct) + case (T_refined_type x' b' c') + then show ?case by auto +qed + +lemma forget_subst_tv [simp]: "atom a \ A \ subst_tv A a x = A" + apply (nominal_induct A avoiding: a x rule: \.strong_induct) + apply(auto simp: fresh_at_base) + done + +lemma subst_tv_id [simp]: "subst_tv A a (V_var a) = A" + by (nominal_induct A avoiding: a rule: \.strong_induct) (auto simp: fresh_at_base) + +lemma fresh_subst_tv_if [simp]: + "j \ (subst_tv A i x ) \ (atom i \ A \ j \ A) \ (j \ x \ (j \ A \ j = atom i))" + apply (nominal_induct A avoiding: i x rule: \.strong_induct) + using fresh_def supp_b_empty x_fresh_b by auto + +lemma subst_tv_commute [simp]: + "atom y \ \ \ (\[x::= t]\<^sub>\\<^sub>v)[y::=v]\<^sub>\\<^sub>v = \[x::= t[y::=v]\<^sub>v\<^sub>v]\<^sub>\\<^sub>v " + by (nominal_induct \ avoiding: x y t v rule: \.strong_induct) (auto simp: fresh_at_base) + +lemma subst_tv_var_flip [simp]: + fixes x::x and xa::x and \::\ + assumes "atom xa \ \" + shows "(x \ xa) \ \ = \[x::=V_var xa]\<^sub>\\<^sub>v" +proof - + obtain z::x and b and c where zbc: "atom z \ (x,xa, V_var xa) \ \ = \ z : b | c \" + using obtain_fresh_z by (metis prod.inject subst_tv.cases) + hence "atom xa \ supp c - { atom z }" using \.supp[of z b c] fresh_def supp_b_empty assms + by auto + moreover have "xa \ z" using zbc fresh_prod3 by force + ultimately have xaf: "atom xa \ c" using fresh_def by auto + have "(x \ xa) \ \ = \ z : b | (x \ xa) \ c \" + by (metis \.perm_simps empty_iff flip_at_base_simps(3) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2) fresh_def not_self_fresh supp_b_empty v.fresh(2) zbc) + also have "... = \ z : b | c[x::=V_var xa]\<^sub>c\<^sub>v \" using subst_cv_v_flip xaf + by (metis permute_flip_cancel permute_flip_cancel2 subst_cv_var_flip) + finally show ?thesis using subst_tv.simps zbc + using fresh_PairD(1) not_self_fresh by force +qed + +instantiation \ :: has_subst_v +begin + +definition + "subst_v = subst_tv" + +instance proof + fix j::atom and i::x and x::v and t::\ + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + + proof(nominal_induct t avoiding: i x rule:\.strong_induct) + case (T_refined_type z b c) + hence " j \ \ z : b | c \[i::=x]\<^sub>v = j \ \ z : b | c[i::=x]\<^sub>c\<^sub>v \" using subst_tv.simps subst_v_\_def fresh_Pair by simp + also have "... = (atom i \ \ z : b | c \ \ j \ \ z : b | c \ \ j \ x \ (j \ \ z : b | c \ \ j = atom i))" + unfolding \.fresh using subst_v_c_def fresh_subst_v_if + using T_refined_type.hyps(1) T_refined_type.hyps(2) x_fresh_b by auto + finally show ?case by auto + qed + + fix a::x and tm::\ and x::v + show "atom a \ tm \ subst_v tm a x = tm" + apply(nominal_induct tm avoiding: a x rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix a::x and tm::\ + show "subst_v tm a (V_var a) = tm" + apply(nominal_induct tm avoiding: a rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix p::perm and x1::x and v::v and t1::\ + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + apply(nominal_induct tm avoiding: a x rule:\.strong_induct) + using subst_v_c_def forget_subst_v subst_tv.simps subst_v_\_def fresh_Pair by simp + + fix x::x and c::\ and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: z x rule:\.strong_induct) + using subst_v_c_def flip_subst_v subst_tv.simps subst_v_\_def fresh_Pair by auto + + fix x::x and c::\ and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: x v z rule:\.strong_induct) + using subst_v_c_def subst_tv.simps subst_v_\_def fresh_Pair + by (metis flip_commute subst_tv_commute subst_tv_var_flip subst_v_\_def subst_vv.simps(2)) +qed + +end + +lemma subst_tv_commute_full: + fixes c::\ + assumes "atom z \ v" and "atom x \ w" and "x\z" + shows "(c[z::=w]\<^sub>\\<^sub>v)[x::=v]\<^sub>\\<^sub>v = (c[x::=v]\<^sub>\\<^sub>v)[z::=w]\<^sub>\\<^sub>v" + using assms proof(nominal_induct c avoiding: x v z w rule: \.strong_induct) + case (T_refined_type x1a x2a x3a) + then show ?case using subst_cv_commute_full by simp +qed + +lemma type_eq_subst_eq: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = \ z2 : b2 | c2 \" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" + using subst_v_flip_eq_two[of z1 c1 z2 c2 v] \.eq_iff assms subst_v_c_def by simp + +text \Extract constraint from a type. We cannot just project out the constraint as this would +mean alpha-equivalent types give different answers \ + +nominal_function c_of :: "\ \ x \ c" where + "atom z \ x \ c_of (T_refined_type z b c) x = c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" +proof(goal_cases) + case 1 + then show ?case using eqvt_def c_of_graph_aux_def by force +next + case (2 x y) + then show ?case using eqvt_def c_of_graph_aux_def by force +next + case (3 P x) + then obtain x1::\ and x2::x where *:"x = (x1,x2)" by force + obtain z' and b' and c' where "x1 = \ z' : b' | c' \ \ atom z' \ x2" using obtain_fresh_z by metis + then show ?case using 3 * by auto +next + case (4 z1 x1 b1 c1 z2 x2 b2 c2) + then show ?case using subst_v_flip_eq_two \.eq_iff by (metis prod.inject type_eq_subst_eq) +qed + +nominal_termination (eqvt) by lexicographic_order + +lemma c_of_eq: + shows "c_of \ x : b | c \ x = c" +proof(nominal_induct "\ x : b | c \" avoiding: x rule: \.strong_induct) + case (T_refined_type x' c') + moreover hence "c_of \ x' : b | c' \ x = c'[x'::=V_var x]\<^sub>c\<^sub>v" using c_of.simps by auto + moreover have "\ x' : b | c' \ = \ x : b | c \" using T_refined_type \.eq_iff by metis + moreover have "c'[x'::=V_var x]\<^sub>c\<^sub>v = c" using T_refined_type Abs1_eq_iff flip_subst_v subst_v_c_def + by (metis subst_cv_id) + ultimately show ?case by auto +qed + +lemma obtain_fresh_z_c_of: + fixes t::"'b::fs" + obtains z where "atom z \ t \ \ = \ z : b_of \ | c_of \ z \" +proof - + obtain z and c where "atom z \ t \ \ = \ z : b_of \ | c \" using obtain_fresh_z2 by metis + moreover hence "c = c_of \ z" using c_of.simps using c_of_eq by metis + ultimately show ?thesis + using that by auto +qed + +lemma c_of_fresh: + fixes x::x + assumes "atom x \ (t,z)" + shows "atom x \ c_of t z" +proof - + obtain z' and c' where z:"t = \ z' : b_of t | c' \ \ atom z' \ (x,z)" using obtain_fresh_z_c_of by metis + hence *:"c_of t z = c'[z'::=V_var z]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair by metis + have "(atom x \ c' \ atom x \ set [atom z']) \ atom x \ b_of t" using \.fresh assms z fresh_Pair by metis + hence "atom x \ c'" using fresh_Pair z fresh_at_base(2) by fastforce + moreover have "atom x \ V_var z" using assms fresh_Pair v.fresh by metis + ultimately show ?thesis using assms fresh_subst_v_if[of "atom x" c' z' "V_var z"] subst_v_c_def * by metis +qed + +lemma c_of_switch: + fixes z::x + assumes "atom z \ t" + shows "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c_of t x" +proof - + obtain z' and c' where z:"t = \ z' : b_of t | c' \ \ atom z' \ (x,z)" using obtain_fresh_z_c_of by metis + hence "(atom z \ c' \ atom z \ set [atom z']) \ atom z \ b_of t" using \.fresh[of "atom z" z' "b_of t" c'] assms by metis + moreover have " atom z \ set [atom z']" using z fresh_Pair by force + ultimately have **:"atom z \ c'" using fresh_Pair z fresh_at_base(2) by metis + + have "(c_of t z)[z::=V_var x]\<^sub>c\<^sub>v = c'[z'::=V_var z]\<^sub>c\<^sub>v[z::=V_var x]\<^sub>c\<^sub>v" using c_of.simps fresh_Pair z by metis + also have "... = c'[z'::=V_var x]\<^sub>c\<^sub>v" using subst_v_simple_commute subst_v_c_def assms c_of.simps z ** by metis + finally show ?thesis using c_of.simps[of z' x "b_of t" c'] fresh_Pair z by metis +qed + +lemma type_eq_subst_eq1: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" and "atom z1 \ c2" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and " c1 = (z1 \ z2) \ c2" +proof - + show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast + show "b1=b2" using \.eq_iff assms by blast + have "z1 = z2 \ c1 = c2 \ z1 \ z2 \ c1 = (z1 \ z2) \ c2 \ atom z1 \ c2" + using \.eq_iff Abs1_eq_iff[of z1 c1 z2 c2] assms by blast + thus "c1 = (z1 \ z2) \ c2" by auto +qed + +lemma type_eq_subst_eq2: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" + shows "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" and "b1=b2" and "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" +proof - + show "c1[z1::=v]\<^sub>c\<^sub>v = c2[z2::=v]\<^sub>c\<^sub>v" using type_eq_subst_eq assms by blast + show "b1=b2" using \.eq_iff assms by blast + show "[[atom z1]]lst. c1 = [[atom z2]]lst. c2" + using \.eq_iff assms by auto +qed + +lemma type_eq_subst_eq3: + fixes v::v and c1::c + assumes "\ z1 : b1 | c1 \ = (\ z2 : b2 | c2 \)" and "atom z1 \ c2" + shows "c1 = c2[z2::=V_var z1]\<^sub>c\<^sub>v" and "b1=b2" + using type_eq_subst_eq1 assms subst_v_c_def + by (metis subst_cv_var_flip)+ + +lemma type_eq_flip: + assumes "atom x \ c" + shows "\ z : b | c \ = \ x : b | (x \ z ) \ c \" + using \.eq_iff Abs1_eq_iff assms + by (metis (no_types, lifting) flip_fresh_fresh) + +lemma c_of_true: + "c_of \ z' : B_bool | TRUE \ x = C_true" +proof(nominal_induct "\ z' : B_bool | TRUE \" avoiding: x rule:\.strong_induct) + case (T_refined_type x1a x3a) + hence "\ z' : B_bool | TRUE \ = \ x1a : B_bool | x3a \" using \.eq_iff by metis + then show ?case using subst_cv.simps c_of.simps T_refined_type + type_eq_subst_eq3 + by (metis type_eq_subst_eq) +qed + +lemma type_eq_subst: + assumes "atom x \ c" + shows "\ z : b | c \ = \ x : b | c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v \" + using \.eq_iff Abs1_eq_iff assms + using subst_cv_var_flip type_eq_flip by auto + +lemma type_e_subst_fresh: + fixes x::x and z::x + assumes "atom z \ (x,v)" and "atom x \ e" + shows "\ z : b | CE_val (V_var z) == e \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == e \" + using assms subst_tv.simps subst_cv.simps forget_subst_cev by simp + +lemma type_v_subst_fresh: + fixes x::x and z::x + assumes "atom z \ (x,v)" and "atom x \ v'" + shows "\ z : b | CE_val (V_var z) == CE_val v' \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == CE_val v' \" + using assms subst_tv.simps subst_cv.simps by simp + +lemma subst_tbase_eq: + "b_of \ = b_of \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z and b and c where zbc: "\ = \ z:b|c\ \ atom z \ (x,v)" using \.exhaust + by (metis prod.inject subst_tv.cases) + hence "b_of \ z:b|c\ = b_of \ z:b|c\[x::=v]\<^sub>\\<^sub>v" using subst_tv.simps by simp + thus ?thesis using zbc by blast +qed + +lemma subst_tv_if: + assumes "atom z1 \ (x,v)" and "atom z' \ (x,v)" + shows "\ z1 : b | CE_val (v'[x::=v]\<^sub>v\<^sub>v) == CE_val (V_lit l) IMP (c'[x::=v]\<^sub>c\<^sub>v)[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \ = + \ z1 : b | CE_val v' == CE_val (V_lit l) IMP c'[z'::=[z1]\<^sup>v]\<^sub>c\<^sub>v \[x::=v]\<^sub>\\<^sub>v" + using subst_cv_commute_full[of z' v x "V_var z1" c'] subst_tv.simps subst_vv.simps(1) subst_ev.simps subst_cv.simps assms + by simp + +lemma subst_tv_tid: + assumes "atom za \ (x,v)" + shows "\ za : B_id tid | TRUE \ = \ za : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" + using assms subst_tv.simps subst_cv.simps by presburger + + +lemma b_of_subst: + "b_of (\[x::=v]\<^sub>\\<^sub>v) = b_of \" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (x,v)" using obtain_fresh_z by metis + thus ?thesis using subst_tv.simps * by auto +qed + +lemma subst_tv_flip: + assumes "\'[x::=v]\<^sub>\\<^sub>v = \" and "atom x \ (v,\)" and "atom x' \ (v,\)" + shows "((x' \ x) \ \')[x'::=v]\<^sub>\\<^sub>v = \" +proof - + have "(x' \ x) \ v = v \ (x' \ x) \ \ = \" using assms flip_fresh_fresh by auto + thus ?thesis using subst_tv.eqvt[of "(x' \ x)" \' x v ] assms by auto +qed + +lemma subst_cv_true: + "\ z : B_id tid | TRUE \ = \ z : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain za::x where "atom za \ (x,v)" using obtain_fresh by auto + hence "\ z : B_id tid | TRUE \ = \ za: B_id tid | TRUE \" using \.eq_iff Abs1_eq_iff by fastforce + moreover have "\ za : B_id tid | TRUE \ = \ za : B_id tid | TRUE \[x::=v]\<^sub>\\<^sub>v" + using subst_cv.simps subst_tv.simps by (simp add: \atom za \ (x, v)\) + ultimately show ?thesis by argo +qed + +lemma t_eq_supp: + assumes "(\ z : b | c \) = (\ z1 : b1 | c1 \)" + shows "supp c - { atom z } = supp c1 - { atom z1 }" +proof - + have "supp c - { atom z } \ supp b = supp c1 - { atom z1 } \ supp b1" using \.supp assms + by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty) + moreover have "supp b = supp b1" using assms \.eq_iff by simp + moreover have "atom z1 \ supp b1 \ atom z \ supp b" using supp_b_empty by simp + ultimately show ?thesis + by (metis \.eq_iff \.supp assms b.supp(1) list.set(1) list.set(2) sup_bot.right_neutral) +qed + +lemma fresh_t_eq: + fixes x::x + assumes "(\ z : b | c \) = (\ zz : b | cc \)" and "atom x \ c" and "x \ zz" + shows "atom x \ cc" +proof - + have "supp c - { atom z } \ supp b = supp cc - { atom zz } \ supp b" using \.supp assms + by (metis list.set(1) list.simps(15) sup_bot.right_neutral supp_b_empty) + moreover have "atom x \ supp c" using assms fresh_def by blast + ultimately have "atom x \ supp cc - { atom zz } \ supp b" by force + hence "atom x \ supp cc" using assms by simp + thus ?thesis using fresh_def by auto +qed + +section \Mutable Variable Context\ + +nominal_function subst_dv :: "\ \ x \ v \ \" where + "subst_dv DNil x v = DNil" +| "subst_dv ((u,t) #\<^sub>\ \) x v = ((u,t[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ (subst_dv \ x v ))" + apply (simp add: eqvt_def subst_dv_graph_aux_def,auto ) + using delete_aux.elims by (metis \.exhaust surj_pair) +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_dv_abbrev :: "\ \ x \ v \ \" ("_[_::=_]\<^sub>\\<^sub>v" [1000,50,50] 1000) + where + "\[x::=v]\<^sub>\\<^sub>v \ subst_dv \ x v " + +nominal_function dmap :: "(u*\ \ u*\) \ \ \ \" where + "dmap f DNil = DNil" +| "dmap f ((u,t)#\<^sub>\\) = (f (u,t) #\<^sub>\ (dmap f \ ))" + apply (simp add: eqvt_def dmap_graph_aux_def,auto ) + using delete_aux.elims by (metis \.exhaust surj_pair) +nominal_termination (eqvt) by lexicographic_order + +lemma subst_dv_iff: + "\[x::=v]\<^sub>\\<^sub>v = dmap (\(u,t). (u, t[x::=v]\<^sub>\\<^sub>v)) \" + by(induct \, auto) + +lemma size_subst_dv [simp]: "size ( subst_dv G i x) \ size G" + by (induct G,auto) + +lemma forget_subst_dv [simp]: "atom a \ G \ subst_dv G a x = G" + apply (induct G ,auto) + using fresh_DCons fresh_PairD(1) not_self_fresh apply fastforce + apply (simp add: fresh_DCons)+ + done + +lemma subst_dv_member: + assumes "(u,\) \ setD \" + shows "(u, \[x::=v]\<^sub>\\<^sub>v) \ setD (\[x::=v]\<^sub>\\<^sub>v)" + using assms by(induct \ rule: \_induct,auto) + +lemma fresh_subst_dv: + fixes x::x + assumes "atom xa \ \" and "atom xa \ v" + shows "atom xa \\[x::=v]\<^sub>\\<^sub>v" + using assms proof(induct \ rule:\_induct) + case DNil + then show ?case by auto +next + case (DCons u t \) + then show ?case using subst_dv.simps subst_v_\_def fresh_DCons fresh_Pair by simp +qed + +lemma fresh_subst_dv_if: + fixes j::atom and i::x and x::v and t::\ + assumes "j \ t \ j \ x" + shows "(j \ subst_dv t i x)" + using assms proof(induct t rule: \_induct) + case DNil + then show ?case using subst_gv.simps fresh_GNil by auto +next + case (DCons u' t' D') + then show ?case unfolding subst_dv.simps using fresh_DCons fresh_subst_tv_if fresh_Pair by metis +qed + +section \Statements\ + +text \ Using ideas from proofs at top of AFP/Launchbury/Substitution.thy. + Subproofs borrowed from there; hence the apply style proofs. \ + +nominal_function (default "case_sum (\x. Inl undefined) (case_sum (\x. Inl undefined) (\x. Inr undefined))") + subst_sv :: "s \ x \ v \ s" + and subst_branchv :: "branch_s \ x \ v \ branch_s" + and subst_branchlv :: "branch_list \ x \ v \ branch_list" where + "subst_sv ( (AS_val v') ) x v = (AS_val (subst_vv v' x v ))" +| "atom y \ (x,v) \ subst_sv (AS_let y e s) x v = (AS_let y (e[x::=v]\<^sub>e\<^sub>v) (subst_sv s x v ))" +| "atom y \ (x,v) \ subst_sv (AS_let2 y t s1 s2) x v = (AS_let2 y (t[x::=v]\<^sub>\\<^sub>v) (subst_sv s1 x v ) (subst_sv s2 x v ))" +| " subst_sv (AS_match v' cs) x v = AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v )" +| "subst_sv (AS_assign y v') x v = AS_assign y (subst_vv v' x v )" +| "subst_sv ( (AS_if v' s1 s2) ) x v = (AS_if (subst_vv v' x v ) (subst_sv s1 x v ) (subst_sv s2 x v ) )" +| "atom u \ (x,v) \ subst_sv (AS_var u \ v' s) x v = AS_var u (subst_tv \ x v ) (subst_vv v' x v ) (subst_sv s x v ) " +| "subst_sv (AS_while s1 s2) x v = AS_while (subst_sv s1 x v ) (subst_sv s2 x v )" +| "subst_sv (AS_seq s1 s2) x v = AS_seq (subst_sv s1 x v ) (subst_sv s2 x v )" +| "subst_sv (AS_assert c s) x v = AS_assert (subst_cv c x v) (subst_sv s x v)" +| "atom x1 \ (x,v) \ subst_branchv (AS_branch dc x1 s1 ) x v = AS_branch dc x1 (subst_sv s1 x v )" + +| "subst_branchlv (AS_final cs) x v = AS_final (subst_branchv cs x v )" +| "subst_branchlv (AS_cons cs css) x v = AS_cons (subst_branchv cs x v ) (subst_branchlv css x v )" + apply (auto,simp add: eqvt_def subst_sv_subst_branchv_subst_branchlv_graph_aux_def ) +proof(goal_cases) + + have eqvt_at_proj: "\ s xa va . eqvt_at subst_sv_subst_branchv_subst_branchlv_sumC (Inl (s, xa, va)) \ + eqvt_at (\a. projl (subst_sv_subst_branchv_subst_branchlv_sumC (Inl a))) (s, xa, va)" + apply(simp add: eqvt_at_def) + apply(rule) + apply(subst Projl_permute) + apply(thin_tac _)+ + apply (simp add: subst_sv_subst_branchv_subst_branchlv_sumC_def) + apply (simp add: THE_default_def) + apply (case_tac "Ex1 (subst_sv_subst_branchv_subst_branchlv_graph (Inl (s,xa,va)))") + apply simp + apply(auto)[1] + apply (erule_tac x="x" in allE) + apply simp + apply(cases rule: subst_sv_subst_branchv_subst_branchlv_graph.cases) + apply(assumption) + apply(rule_tac x="Sum_Type.projl x" in exI,clarify,rule the1_equality,blast,simp (no_asm) only: sum.sel)+ + apply blast + + + apply(simp)+ + done + + { + case (1 P x') + then show ?case proof(cases x') + case (Inl a) thus P + proof(cases a) + case (fields aa bb cc) + thus P using Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + next + case (Inr b) thus P + proof(cases b) + case (Inl a) thus P proof(cases a) + case (fields aa bb cc) + then show ?thesis using Inr Inl 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + next + case Inr2: (Inr b) thus P proof(cases b) + case (fields aa bb cc) + then show ?thesis using Inr Inr2 1 s_branch_s_branch_list.strong_exhaust fresh_star_insert by metis + qed + qed + qed + next + case (2 y s ya xa va sa c) + thus ?case using eqvt_triple eqvt_at_proj by blast + next + case (3 y s2 ya xa va s1a s2a c) + thus ?case using eqvt_triple eqvt_at_proj by blast + next + case (4 u xa va s ua sa c) + moreover have "atom u \ (xa, va) \ atom ua \ (xa, va)" + using fresh_Pair u_fresh_xv by auto + ultimately show ?case using eqvt_triple[of u xa va ua s sa] subst_sv_def eqvt_at_proj by metis + next + case (5 x1 s1 x1a xa va s1a c) + thus ?case using eqvt_triple eqvt_at_proj by blast + } +qed +nominal_termination (eqvt) by lexicographic_order + +abbreviation + subst_sv_abbrev :: "s \ x \ v \ s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000) + where + "s[x::=v]\<^sub>s\<^sub>v \ subst_sv s x v" + +abbreviation + subst_branchv_abbrev :: "branch_s \ x \ v \ branch_s" ("_[_::=_]\<^sub>s\<^sub>v" [1000,50,50] 1000) + where + "s[x::=v]\<^sub>s\<^sub>v \ subst_branchv s x v" + +lemma size_subst_sv [simp]: "size (subst_sv A i x ) = size A" and "size (subst_branchv B i x ) = size B" and "size (subst_branchlv C i x ) = size C" + by(nominal_induct A and B and C avoiding: i x rule: s_branch_s_branch_list.strong_induct,auto) + +lemma forget_subst_sv [simp]: shows "atom a \ A \ subst_sv A a x = A" and "atom a \ B \ subst_branchv B a x = B" and "atom a \ C \ subst_branchlv C a x = C" + by (nominal_induct A and B and C avoiding: a x rule: s_branch_s_branch_list.strong_induct,auto simp: fresh_at_base) + +lemma subst_sv_id [simp]: "subst_sv A a (V_var a) = A" and "subst_branchv B a (V_var a) = B" and "subst_branchlv C a (V_var a) = C" +proof(nominal_induct A and B and C avoiding: a rule: s_branch_s_branch_list.strong_induct) + case (AS_let x option e s) + then show ?case + by (metis (no_types, lifting) fresh_Pair not_None_eq subst_ev_id subst_sv.simps(2) subst_sv.simps(3) subst_tv_id v.fresh(2)) +next + case (AS_match v branch_s) + then show ?case using fresh_Pair not_None_eq subst_ev_id subst_sv.simps subst_sv.simps subst_tv_id v.fresh subst_vv_id + by metis +qed(auto)+ + +lemma fresh_subst_sv_if_rl: + shows + "(atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x)) \ j \ (subst_sv s x v )" and + "(atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x)) \ j \ (subst_branchv cs x v)" and + "(atom x \ css \ j \ css) \ (j \ v \ (j \ css \ j = atom x)) \ j \ (subst_branchlv css x v )" + apply(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct) + using pure_fresh by force+ + +lemma fresh_subst_sv_if_lr: + shows "j \ (subst_sv s x v) \ (atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x))" and + "j \ (subst_branchv cs x v) \ (atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x))" and + "j \ (subst_branchlv css x v ) \ (atom x \ css \ j \ css) \ (j \ v \ (j \ css \ j = atom x))" +proof(nominal_induct s and cs and css avoiding: v x rule: s_branch_s_branch_list.strong_induct) + case (AS_branch list x s ) + then show ?case using s_branch_s_branch_list.fresh fresh_Pair list.distinct(1) list.set_cases pure_fresh set_ConsD subst_branchv.simps by metis +next + case (AS_let y e s') + thus ?case proof(cases "atom x \ (AS_let y e s')") + case True + hence "subst_sv (AS_let y e s') x v = (AS_let y e s')" using forget_subst_sv by simp + hence "j \ (AS_let y e s')" using AS_let by argo + then show ?thesis using True by blast + next + case False + have "subst_sv (AS_let y e s') x v = AS_let y (e[x::=v]\<^sub>e\<^sub>v) (s'[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps(2) AS_let by force + hence "((j \ s'[x::=v]\<^sub>s\<^sub>v \ j \ set [atom y]) \ j \ None \ j \ e[x::=v]\<^sub>e\<^sub>v)" using s_branch_s_branch_list.fresh AS_let + by (simp add: fresh_None) + then show ?thesis using AS_let fresh_None fresh_subst_ev_if list.discI list.set_cases s_branch_s_branch_list.fresh set_ConsD + by metis + qed +next + case (AS_let2 y \ s1 s2) + thus ?case proof(cases "atom x \ (AS_let2 y \ s1 s2)") + case True + hence "subst_sv (AS_let2 y \ s1 s2) x v = (AS_let2 y \ s1 s2)" using forget_subst_sv by simp + hence "j \ (AS_let2 y \ s1 s2)" using AS_let2 by argo + then show ?thesis using True by blast + next + case False + have "subst_sv (AS_let2 y \ s1 s2) x v = AS_let2 y (\[x::=v]\<^sub>\\<^sub>v) (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps AS_let2 by force + then show ?thesis using AS_let2 + fresh_subst_tv_if list.discI list.set_cases s_branch_s_branch_list.fresh(4) set_ConsD by auto + qed +qed(auto)+ + +lemma fresh_subst_sv_if[simp]: + fixes x::x and v::v + shows "j \ (subst_sv s x v) \ (atom x \ s \ j \ s) \ (j \ v \ (j \ s \ j = atom x))" and + "j \ (subst_branchv cs x v) \ (atom x \ cs \ j \ cs) \ (j \ v \ (j \ cs \ j = atom x))" + using fresh_subst_sv_if_lr fresh_subst_sv_if_rl by metis+ + +lemma subst_sv_commute [simp]: + fixes A::s and t::v and j::x and i::x + shows "atom j \ A \ (subst_sv (subst_sv A i t) j u ) = subst_sv A i (subst_vv t j u )" and + "atom j \ B \ (subst_branchv (subst_branchv B i t ) j u ) = subst_branchv B i (subst_vv t j u )" and + "atom j \ C \ (subst_branchlv (subst_branchlv C i t) j u ) = subst_branchlv C i (subst_vv t j u ) " + apply(nominal_induct A and B and C avoiding: i j t u rule: s_branch_s_branch_list.strong_induct) + by(auto simp: fresh_at_base) + +lemma c_eq_perm: + assumes "( (atom z) \ (atom z') ) \ c = c'" and "atom z' \ c" + shows "\ z : b | c \ = \ z' : b | c' \" + using \.eq_iff Abs1_eq_iff(3) + by (metis Nominal2_Base.swap_commute assms(1) assms(2) flip_def swap_fresh_fresh) + +lemma subst_sv_flip: + fixes s::s and sa::s and v'::v + assumes "atom c \ (s, sa)" and "atom c \ (v',x, xa, s, sa)" "atom x \ v'" and "atom xa \ v'" and "(x \ c) \ s = (xa \ c) \ sa" + shows "s[x::=v']\<^sub>s\<^sub>v = sa[xa::=v']\<^sub>s\<^sub>v" +proof - + have "atom x \ (s[x::=v']\<^sub>s\<^sub>v)" and xafr: "atom xa \ (sa[xa::=v']\<^sub>s\<^sub>v)" + and "atom c \ ( s[x::=v']\<^sub>s\<^sub>v, sa[xa::=v']\<^sub>s\<^sub>v)" using assms using fresh_subst_sv_if assms by( blast+ ,force) + + hence "s[x::=v']\<^sub>s\<^sub>v = (x \ c) \ (s[x::=v']\<^sub>s\<^sub>v)" by (simp add: flip_fresh_fresh fresh_Pair) + also have " ... = ((x \ c) \ s)[ ((x \ c) \ x) ::= ((x \ c) \ v') ]\<^sub>s\<^sub>v" using subst_sv_subst_branchv_subst_branchlv.eqvt by blast + also have "... = ((xa \ c) \ sa)[ ((x \ c) \ x) ::= ((x \ c) \ v') ]\<^sub>s\<^sub>v" using assms by presburger + also have "... = ((xa \ c) \ sa)[ ((xa \ c) \ xa) ::= ((xa \ c) \ v') ]\<^sub>s\<^sub>v" using assms + by (metis flip_at_simps(1) flip_fresh_fresh fresh_PairD(1)) + also have "... = (xa \ c) \ (sa[xa::=v']\<^sub>s\<^sub>v)" using subst_sv_subst_branchv_subst_branchlv.eqvt by presburger + also have "... = sa[xa::=v']\<^sub>s\<^sub>v" using xafr assms by (simp add: flip_fresh_fresh fresh_Pair) + finally show ?thesis by simp +qed + +lemma if_type_eq: + fixes \::\ and v::v and z1::x + assumes "atom z1' \ (v, ca, (x, b, c) #\<^sub>\ \, (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" and "atom z1 \ v" + and "atom z1 \ (za,ca)" and "atom z1' \ (za,ca)" + shows "(\ z1' : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v \) = \ z1 : ba | CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v \" +proof - + have "atom z1' \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" using assms fresh_prod4 by blast + moreover hence "(CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v) = (z1' \ z1) \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )" + proof - + have "(z1' \ z1) \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ) = ( (z1' \ z1) \ (CE_val v == CE_val (V_lit ll)) IMP ((z1' \ z1) \ ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" + by auto + also have "... = ((CE_val v == CE_val (V_lit ll)) IMP ((z1' \ z1) \ ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v ))" + using \atom z1 \ v\ assms + by (metis (mono_tags) \atom z1' \ (CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v )\ c.fresh(6) c.fresh(7) ce.fresh(1) flip_at_simps(2) flip_fresh_fresh fresh_at_base_permute_iff fresh_def supp_l_empty v.fresh(1)) + also have "... = ((CE_val v == CE_val (V_lit ll)) IMP (ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v ))" + using assms by fastforce + finally show ?thesis by auto + qed + ultimately show ?thesis + using \.eq_iff Abs1_eq_iff(3)[of z1' "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1']\<^sup>v]\<^sub>c\<^sub>v" + z1 "CE_val v == CE_val (V_lit ll) IMP ca[za::=[z1]\<^sup>v]\<^sub>c\<^sub>v"] by blast +qed + +lemma subst_sv_var_flip: + fixes x::x and s::s and z::x + shows "atom x \ s \ ((x \ z) \ s) = s[z::=[x]\<^sup>v]\<^sub>s\<^sub>v" and + "atom x \ cs \ ((x \ z) \ cs) = subst_branchv cs z [x]\<^sup>v" and + "atom x \ css \ ((x \ z) \ css) = subst_branchlv css z [x]\<^sup>v" + apply(nominal_induct s and cs and css avoiding: z x rule: s_branch_s_branch_list.strong_induct) + using [[simproc del: alpha_lst]] + apply (auto ) (* This unpacks subst, perm *) + using subst_tv_var_flip flip_fresh_fresh v.fresh s_branch_s_branch_list.fresh + subst_v_\_def subst_v_v_def subst_vv_var_flip subst_v_e_def subst_ev_var_flip pure_fresh apply auto + defer 1 (* Sometimes defering hard goals to the end makes it easier to finish *) + using x_fresh_u apply blast (* Next two involve u and flipping with x *) + defer 1 + using x_fresh_u apply blast + defer 1 + using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh + apply (simp add: subst_v_c_def) + using x_fresh_u Abs1_eq_iff'(3) flip_fresh_fresh + by (simp add: flip_fresh_fresh) + +instantiation s :: has_subst_v +begin + +definition + "subst_v = subst_sv" + +instance proof + fix j::atom and i::x and x::v and t::s + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + using fresh_subst_sv_if subst_v_s_def by auto + + fix a::x and tm::s and x::v + show "atom a \ tm \ subst_v tm a x = tm" + using forget_subst_sv subst_v_s_def by simp + + fix a::x and tm::s + show "subst_v tm a (V_var a) = tm" using subst_sv_id subst_v_s_def by simp + + fix p::perm and x1::x and v::v and t1::s + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + using subst_sv_commute subst_v_s_def by simp + + fix x::x and c::s and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + using subst_sv_var_flip subst_v_s_def by simp + + fix x::x and c::s and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + using subst_sv_var_flip subst_v_s_def by simp +qed +end + +section \Type Definition\ + +nominal_function subst_ft_v :: "fun_typ \ x \ v \ fun_typ" where + "atom z \ (x,v) \ subst_ft_v ( AF_fun_typ z b c t (s::s)) x v = AF_fun_typ z b c[x::=v]\<^sub>c\<^sub>v t[x::=v]\<^sub>\\<^sub>v s[x::=v]\<^sub>s\<^sub>v" + apply(simp add: eqvt_def subst_ft_v_graph_aux_def ) + apply(simp add:fun_typ.strong_exhaust ) + apply(auto) + apply(rule_tac y=a and c="(aa,b)" in fun_typ.strong_exhaust) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) + +proof(goal_cases) + case (1 z xa va c t s za ca ta sa cb) + hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v" + by (metis flip_commute subst_cv_var_flip) + hence "c[z::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v = ca[za::=[ cb ]\<^sup>v]\<^sub>c\<^sub>v[xa::=va]\<^sub>c\<^sub>v" by auto + then show ?case using subst_cv_commute atom_eq_iff fresh_atom fresh_atom_at_base subst_cv_commute_full v.fresh + using 1 subst_cv_var_flip flip_commute by metis +next + case (2 z xa va c t s za ca ta sa cb) + hence "t[z::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v" by metis + hence "t[z::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v[xa::=va]\<^sub>\\<^sub>v = ta[za::=[ cb ]\<^sup>v]\<^sub>\\<^sub>v[xa::=va]\<^sub>\\<^sub>v" by auto + then show ?case using subst_tv_commute_full 2 + by (metis atom_eq_iff fresh_atom fresh_atom_at_base v.fresh(2)) +qed + +nominal_termination (eqvt) by lexicographic_order + +nominal_function subst_ftq_v :: "fun_typ_q \ x \ v \ fun_typ_q" where + "atom bv \ (x,v) \ subst_ftq_v (AF_fun_typ_some bv ft) x v = (AF_fun_typ_some bv (subst_ft_v ft x v))" +| "subst_ftq_v (AF_fun_typ_none ft) x v = (AF_fun_typ_none (subst_ft_v ft x v))" + apply(simp add: eqvt_def subst_ftq_v_graph_aux_def ) + apply(simp add:fun_typ_q.strong_exhaust ) + apply(auto) + apply(rule_tac y=a and c="(aa,b)" in fun_typ_q.strong_exhaust) + apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base) +proof(goal_cases) + case (1 bv ft bva fta xa va c) + then show ?case using subst_ft_v.simps by (simp add: flip_fresh_fresh) +qed +nominal_termination (eqvt) by lexicographic_order + +lemma size_subst_ft[simp]: "size (subst_ft_v A x v) = size A" + by(nominal_induct A avoiding: x v rule: fun_typ.strong_induct,auto) + +lemma forget_subst_ft [simp]: shows "atom x \ A \ subst_ft_v A x a = A" + by (nominal_induct A avoiding: a x rule: fun_typ.strong_induct,auto simp: fresh_at_base) + +lemma subst_ft_id [simp]: "subst_ft_v A a (V_var a) = A" + by(nominal_induct A avoiding: a rule: fun_typ.strong_induct,auto) + +instantiation fun_typ :: has_subst_v +begin + +definition + "subst_v = subst_ft_v" + +instance proof + + fix j::atom and i::x and x::v and t::fun_typ + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + apply(nominal_induct t avoiding: i x rule:fun_typ.strong_induct) + apply(simp only: subst_v_fun_typ_def subst_ft_v.simps ) + using fun_typ.fresh fresh_subst_v_if apply simp + by auto + + fix a::x and tm::fun_typ and x::v + show "atom a \ tm \ subst_v tm a x = tm" + proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix a::x and tm::fun_typ + show "subst_v tm a (V_var a) = tm" + proof(nominal_induct tm avoiding: a x rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix p::perm and x1::x and v::v and t1::fun_typ + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + proof(nominal_induct t1 avoiding: x1 v rule:fun_typ.strong_induct) + case (AF_fun_typ x1a x2a x3a x4a x5a) + then show ?case unfolding subst_ft_v.simps subst_v_fun_typ_def fun_typ.fresh using forget_subst_v subst_ft_v.simps subst_v_c_def forget_subst_sv subst_v_\_def by fastforce + qed + + fix x::x and c::fun_typ and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: x z rule:fun_typ.strong_induct) + by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\_def subst_v_fun_typ_def) + + fix x::x and c::fun_typ and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: z x v rule:fun_typ.strong_induct) + apply auto + by (auto simp add: subst_v_c_def subst_v_s_def subst_v_\_def subst_v_fun_typ_def ) +qed +end + +instantiation fun_typ_q :: has_subst_v +begin + +definition + "subst_v = subst_ftq_v" + +instance proof + fix j::atom and i::x and x::v and t::fun_typ_q + show "(j \ subst_v t i x) = ((atom i \ t \ j \ t) \ (j \ x \ (j \ t \ j = atom i)))" + apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto) + apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + by (metis (no_types) fresh_subst_v_if subst_v_fun_typ_def)+ + + fix i::x and t::fun_typ_q and x::v + show "atom i \ t \ subst_v t i x = t" + apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix i::x and t::fun_typ_q + show "subst_v t i (V_var i) = t" using subst_cv_id subst_v_fun_typ_def + apply(nominal_induct t avoiding: i x rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix p::perm and x1::x and v::v and t1::fun_typ_q + show "p \ subst_v t1 x1 v = subst_v (p \ t1) (p \ x1) (p \ v)" + apply(nominal_induct t1 avoiding: v x1 rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix x::x and c::fun_typ_q and z::x + show "atom x \ c \ ((x \ z) \ c) = c[z::=[x]\<^sup>v]\<^sub>v" + apply(nominal_induct c avoiding: x z rule:fun_typ_q.strong_induct,auto) + by(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + + fix x::x and c::fun_typ_q and z::x + show "atom x \ c \ c[z::=[x]\<^sup>v]\<^sub>v[x::=v]\<^sub>v = c[z::=v]\<^sub>v" + apply(nominal_induct c avoiding: z x v rule:fun_typ_q.strong_induct,auto) + apply(auto simp add: subst_v_fun_typ_def subst_v_s_def subst_v_\_def subst_v_fun_typ_q_def fresh_subst_v_if ) + by (metis subst_v_fun_typ_def flip_bv_x_cancel subst_ft_v.eqvt subst_v_simple_commute v.perm_simps )+ +qed + +end + +section \Variable Context\ + +lemma subst_dv_fst_eq: + "fst ` setD (\[x::=v]\<^sub>\\<^sub>v) = fst ` setD \" + by(induct \ rule: \_induct,simp,force) + +lemma subst_gv_member_iff: + fixes x'::x and x::x and v::v and c'::c + assumes "(x',b',c') \ toSet \" and "atom x \ atom_dom \" + shows "(x',b',c'[x::=v]\<^sub>c\<^sub>v) \ toSet \[x::=v]\<^sub>\\<^sub>v" +proof - + have "x' \ x" using assms fresh_dom_free2 by metis + then show ?thesis using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto + next + case (GCons x1 b1 c1 \') + show ?case proof(cases "(x',b',c') = (x1,b1,c1)") + case True + hence "((x1, b1, c1) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v = ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v))" using subst_gv.simps \x'\x\ by auto + then show ?thesis using True by auto + next + case False + have "x1\x" using fresh_def fresh_GCons fresh_Pair supp_at_base GCons fresh_dom_free2 by auto + hence "(x', b', c') \ toSet \'" using GCons False toSet.simps by auto + moreover have "atom x \ atom_dom \'" using fresh_GCons GCons dom.simps toSet.simps by simp + ultimately have "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \ toSet \'[x::=v]\<^sub>\\<^sub>v" using GCons by auto + hence "(x', b', c'[x::=v]\<^sub>c\<^sub>v) \ toSet ((x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v))" by auto + then show ?thesis using subst_gv.simps \x1\x\ by auto + qed + qed +qed + +lemma fresh_subst_gv_if: + fixes j::atom and i::x and x::v and t::\ + assumes "j \ t \ j \ x" + shows "(j \ subst_gv t i x)" + using assms proof(induct t rule: \_induct) + case GNil + then show ?case using subst_gv.simps fresh_GNil by auto +next + case (GCons x' b' c' \') + then show ?case unfolding subst_gv.simps using fresh_GCons fresh_subst_cv_if by auto +qed + +section \Lookup\ + +lemma set_GConsD: "y \ toSet (x #\<^sub>\ xs) \ y=x \ y \ toSet xs" + by auto + +lemma subst_g_assoc_cons: + assumes "x \ x'" + shows "(((x', b', c') #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v @ G) = ((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ ((\'[x::=v]\<^sub>\\<^sub>v) @ G))" + using subst_gv.simps append_g.simps assms by auto + +end \ No newline at end of file diff --git a/thys/MiniSail/IVSubstTypingL.thy b/thys/MiniSail/IVSubstTypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/IVSubstTypingL.thy @@ -0,0 +1,1534 @@ +(*<*) +theory IVSubstTypingL + imports SubstMethods ContextSubtypingL +begin + (*>*) + +chapter \Immutable Variable Substitution Lemmas\ +text \Lemmas that show that types are preserved, in some way, +under immutable variable substitution\ + +section \Proof Methods\ + +method subst_mth = (metis subst_g_inside infer_e_wf infer_v_wf infer_v_wf) + +method subst_tuple_mth uses add = ( + (unfold fresh_prodN), (simp add: add )+, + (rule,metis fresh_z_subst_g add fresh_Pair ), + (metis fresh_subst_dv add fresh_Pair ) ) + +section \Prelude\ + +lemma subst_top_eq: + "\ z : b | TRUE \ = \ z : b | TRUE \[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z'::x and c' where zeq: "\ z : b | TRUE \ = \ z' : b | c' \ \ atom z' \ (x,v)" using obtain_fresh_z2 b_of.simps by metis + hence "\ z' : b | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z' : b | TRUE \" using subst_tv.simps subst_cv.simps by metis + moreover have "c' = C_true" using \.eq_iff Abs1_eq_iff(3) c.fresh flip_fresh_fresh by (metis zeq) + ultimately show ?thesis using zeq by metis +qed + +lemma wfD_subst: + fixes \\<^sub>1::\ and v::v and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and "wfD \ \ (\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)) \" and "b_of \\<^sub>1=b\<^sub>1" + shows " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v" +proof - + have "\ ; \ ; \\\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_v_wf assms by auto + moreover have "(\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\))[x::=v]\<^sub>\\<^sub>v = \'[x::=v]\<^sub>\\<^sub>v @ \" using subst_g_inside wfD_wf assms by metis + ultimately show ?thesis using wf_subst assms by metis +qed + +lemma subst_v_c_of: + assumes "atom xa \ (v,x)" + shows "c_of t[x::=v]\<^sub>\\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" + using assms proof(nominal_induct t avoiding: x v xa rule:\.strong_induct) + case (T_refined_type z' b' c') + then have " c_of \ z' : b' | c' \[x::=v]\<^sub>\\<^sub>v xa = c_of \ z' : b' | c'[x::=v]\<^sub>c\<^sub>v \ xa" + using subst_tv.simps fresh_Pair by metis + also have "... = c'[x::=v]\<^sub>c\<^sub>v [z'::=V_var xa]\<^sub>c\<^sub>v" using c_of.simps T_refined_type by metis + also have "... = c' [z'::=V_var xa]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v" + using subst_cv_commute_full[of z' v x "V_var xa" c'] subst_v_c_def T_refined_type fresh_Pair fresh_at_base v.fresh fresh_x_neq by metis + finally show ?case using c_of.simps T_refined_type by metis +qed + +section \Context\ + +lemma subst_lookup: + assumes "Some (b,c) = lookup (\'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\\)) y" and "x \ y" and "wfG \ \ (\'@((x,b\<^sub>1,c\<^sub>1)#\<^sub>\\))" + shows "\d. Some (b,d) = lookup ((\'[x::=v]\<^sub>\\<^sub>v)@\) y" + using assms proof(induct \' rule: \_induct) + case GNil + hence "Some (b,c) = lookup \ y" by (simp add: assms(1)) + then show ?case using subst_gv.simps by auto +next + case (GCons x1 b1 c1 \1) + show ?case proof(cases "x1 = x") + case True + hence "atom x \ (\1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\ \)" using GCons wfG_elims(2) + append_g.simps by metis + moreover have "atom x \ atom_dom (\1 @ (x, b\<^sub>1, c\<^sub>1) #\<^sub>\ \)" by simp + ultimately show ?thesis + using forget_subst_gv not_GCons_self2 subst_gv.simps append_g.simps + by (metis GCons.prems(3) True wfG_cons_fresh2 ) + next + case False + hence "((x1,b1,c1) #\<^sub>\ \1)[x::=v]\<^sub>\\<^sub>v = (x1,b1,c1[x::=v]\<^sub>c\<^sub>v)#\<^sub>\\1[x::=v]\<^sub>\\<^sub>v" using subst_gv.simps by auto + then show ?thesis proof(cases "x1=y") + case True + then show ?thesis using GCons using lookup.simps + by (metis \((x1, b1, c1) #\<^sub>\ \1)[x::=v]\<^sub>\\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ \1[x::=v]\<^sub>\\<^sub>v\ append_g.simps fst_conv option.inject) + next + case False + then show ?thesis using GCons using lookup.simps + using \((x1, b1, c1) #\<^sub>\ \1)[x::=v]\<^sub>\\<^sub>v = (x1, b1, c1[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ \1[x::=v]\<^sub>\\<^sub>v\ append_g.simps \.distinct \.inject wfG.simps wfG_elims by metis + qed + qed +qed + +section \Validity\ + +lemma subst_self_valid: + fixes v::v + assumes "\ ; \ ; G \ v \ \ z : b | c \" and "atom z \ v" + shows " \ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" +proof - + have "c= (CE_val (V_var z) == CE_val v )" using infer_v_form2 assms by presburger + hence "c[z::=v]\<^sub>c\<^sub>v = (CE_val (V_var z) == CE_val v )[z::=v]\<^sub>c\<^sub>v" by auto + also have "... = (((CE_val (V_var z))[z::=v]\<^sub>c\<^sub>e\<^sub>v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" by fastforce + also have "... = ((CE_val v) == ((CE_val v)[z::=v]\<^sub>c\<^sub>e\<^sub>v))" using subst_cev.simps subst_vv.simps by presburger + also have "... = (CE_val v == CE_val v )" using infer_v_form subst_cev.simps assms forget_subst_vv by presburger + finally have *:"c[z::=v]\<^sub>c\<^sub>v = (CE_val v == CE_val v )" by auto + + have **:"\ ; \ ; G\\<^sub>w\<^sub>f CE_val v : b" using wfCE_valI assms infer_v_v_wf b_of.simps by metis + + show ?thesis proof(rule validI) + show "\ ; \ ; G\\<^sub>w\<^sub>f c[z::=v]\<^sub>c\<^sub>v" proof - + have "\ ; \ ; G\\<^sub>w\<^sub>f v : b" using infer_v_v_wf assms b_of.simps by metis + moreover have "\ \\<^sub>w\<^sub>f ([]::\) \ \ ; \ ; G\\<^sub>w\<^sub>f []\<^sub>\" using wfD_emptyI wfPhi_emptyI infer_v_wf assms by auto + ultimately show ?thesis using * wfCE_valI wfC_eqI by metis + qed + show "\i. wfI \ G i \ is_satis_g i G \ is_satis i c[z::=v]\<^sub>c\<^sub>v" proof(rule,rule) + fix i + assume \wfI \ G i \ is_satis_g i G\ + thus \is_satis i c[z::=v]\<^sub>c\<^sub>v\ using * ** is_satis_eq by auto + qed + qed +qed + +lemma subst_valid_simple: + fixes v::v + assumes "\ ; \ ; G \ v \ \ z0 : b | c0 \" and + "atom z0 \ c" and "atom z0 \ v" + "\; \ ; (z0,b,c0)#\<^sub>\G \ c[z::=V_var z0]\<^sub>c\<^sub>v" + shows " \ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" +proof - + have " \ ; \ ; G \ c0[z0::=v]\<^sub>c\<^sub>v" using subst_self_valid assms by metis + moreover have "atom z0 \ G" using assms valid_wf_all by meson + moreover have "wfV \ \ G v b" using infer_v_v_wf assms b_of.simps by metis + moreover have "(c[z::=V_var z0]\<^sub>c\<^sub>v)[z0::=v]\<^sub>c\<^sub>v = c[z::=v]\<^sub>c\<^sub>v" using subst_v_simple_commute assms subst_v_c_def by metis + ultimately show ?thesis using valid_trans assms subst_defs by metis +qed + +lemma wfI_subst1: + assumes " wfI \ (G'[x::=v]\<^sub>\\<^sub>v @ G) i" and "wfG \ \ (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G)" and "eval_v i v sv" and "wfRCV \ sv b" + shows "wfI \ (G' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G) ( i( x \ sv))" +proof - + { + fix xa::x and ba::b and ca::c + assume as: "(xa,ba,ca) \ toSet ((G' @ ((x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ G)))" + then have " \s. Some s = (i(x \ sv)) xa \ wfRCV \ s ba" + proof(cases "x=xa") + case True + have "Some sv = (i(x \ sv)) x \ wfRCV \ sv b" using as assms wfI_def by auto + moreover have "b=ba" using assms as True wfG_member_unique by metis + ultimately show ?thesis using True by auto + next + case False + + then obtain ca' where "(xa, ba, ca') \ toSet (G'[x::=v]\<^sub>\\<^sub>v @ G)" using wfG_member_subst2 assms as by metis + then obtain s where " Some s = i xa \ wfRCV \ s ba" using wfI_def assms False by blast + thus ?thesis using False by auto + qed + } + from this show ?thesis using wfI_def allI by blast +qed + +lemma subst_valid: + fixes v::v and c'::c and \ ::\ + assumes "\ ; \ ; \ \ c[z::=v]\<^sub>c\<^sub>v" and "\ ; \ ; \\\<^sub>w\<^sub>f v : b" and + "\ ; \\\<^sub>w\<^sub>f \" and "atom x \ c" and "atom x \ \" and + "\ ; \\\<^sub>w\<^sub>f (\'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \)" and + "\ ; \ ; \'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \ \ c'" (is " \ ; \; ?G \ c'") + shows "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ c'[x::=v]\<^sub>c\<^sub>v" +proof - + have *:"wfC \ \ (\'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ \) c'" using valid.simps assms by metis + hence "wfC \ \ (\'[x::=v]\<^sub>\\<^sub>v @ \) (c'[x::=v]\<^sub>c\<^sub>v)" using wf_subst(2)[OF *] b_of.simps assms subst_g_inside wfC_wf by metis + moreover have "\i. wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ \) i \ is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ \) \ is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" + + proof(rule,rule) + fix i + assume as: " wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ \) i \ is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ \)" + + hence wfig: "wfI \ \ i" using wfI_suffix infer_v_wf assms by metis + then obtain s where s:"eval_v i v s" and b:"wfRCV \ s b" using eval_v_exist infer_v_v_wf b_of.simps assms by metis + + have is1: "is_satis_g ( i( x \ s)) (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" proof(rule is_satis_g_i_upd2) + show "is_satis (i(x \ s)) (c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)" proof - + have "is_satis i (c[z::=v]\<^sub>c\<^sub>v)" + using subst_valid_simple assms as valid.simps infer_v_wf assms + is_satis_g_suffix wfI_suffix by metis + hence "is_satis i ((c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v)" using assms subst_v_simple_commute[of x c z v] subst_v_c_def by metis + moreover have "\ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" using wfC_refl wfG_suffix assms by metis + moreover have "\ ; \ ; \\\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf b_of.simps by metis + ultimately show ?thesis using subst_c_satis[OF s , of \ \ x b "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v" \ "c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v"] wfig by auto + qed + show "atom x \ \" using assms by metis + show "wfG \ \ (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" using valid_wf_all assms by metis + show "\ ; \ ; \\\<^sub>w\<^sub>f v : b" using assms infer_v_v_wf by force + show "i \ v \ ~ s " using s by auto + show "\ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ i" using as by auto + show "i \ \'[x::=v]\<^sub>\\<^sub>v @ \ " using as by auto + qed + hence "is_satis ( i( x \ s)) c'" proof - + have "wfI \ (\' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) ( i( x \ s))" + using wfI_subst1[of \ \' x v \ i \ b c z s] as b s assms by metis + thus ?thesis using is1 valid.simps assms by presburger + qed + + thus "is_satis i (c'[x::=v]\<^sub>c\<^sub>v)" using subst_c_satis_full[OF s] valid.simps as infer_v_v_wf b_of.simps assms by metis + + qed + ultimately show ?thesis using valid.simps by auto +qed + +lemma subst_valid_infer_v: + fixes v::v and c'::c + assumes "\ ; \ ; G \ v \ \ z0 : b | c0 \" and "atom x \ c" and "atom x \ G" and "wfG \ \ (G'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ G)" and "atom z0 \ v" + " \;\;(z0,b,c0)#\<^sub>\G \ c[z::=V_var z0]\<^sub>c\<^sub>v" and "atom z0 \ c" and + " \;\;G'@(x,b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v ) #\<^sub>\ G \ c'" (is " \ ; \; ?G \ c'") + shows " \;\;G'[x::=v]\<^sub>\\<^sub>v@G \ c'[x::=v]\<^sub>c\<^sub>v" +proof - + have "\ ; \ ; G \ c[z::=v]\<^sub>c\<^sub>v" + using infer_v_wf subst_valid_simple valid.simps assms using subst_valid_simple assms valid.simps infer_v_wf assms + is_satis_g_suffix wfI_suffix by metis + moreover have "wfV \ \ G v b" and "wfG \ \ G" + using assms infer_v_wf b_of.simps apply metis using assms infer_v_wf by metis + ultimately show ?thesis using assms subst_valid by metis +qed + +section \Subtyping\ + +lemma subst_subtype: + fixes v::v + assumes "\ ; \ ; \ \ v \ (\z0:b|c0\)" and + " \;\;\ \ (\z0:b|c0\) \ (\ z : b | c \)" and + " \;\;\'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ (\ z1 : b1 | c1 \) \ (\ z2 : b1 | c2 \)" (is " \ ; \; ?G1 \ ?t1 \ ?t2" ) and + "atom z \ (x,v) \ atom z0 \ (c,x,v,z,\) \ atom z1 \ (x,v) \ atom z2 \ (x,v)" and "wsV \ \ \ v" + shows " \;\;\'[x::=v]\<^sub>\\<^sub>v@\ \ \ z1 : b1 | c1 \[x::=v]\<^sub>\\<^sub>v \ \ z2 : b1 | c2 \[x::=v]\<^sub>\\<^sub>v" +proof - + have z2: "atom z2 \ (x,v) " using assms by auto + hence "x \ z2" by auto + + obtain xx::x where xxf: "atom xx \ (x,z1, c1, z2, c2, \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \, c1[x::=v]\<^sub>c\<^sub>v, c2[x::=v]\<^sub>c\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v @ \, + (\ , \ , \'[x::=v]\<^sub>\\<^sub>v@\, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 , c2[x::=v]\<^sub>c\<^sub>v ))" (is "atom xx \ ?tup") + using obtain_fresh by blast + hence xxf2: "atom xx \ (z1, c1, z2, c2, \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)" using fresh_prod9 fresh_prod5 by fast + + have vd1: " \;\;((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v @ \ \ (c2[z2::=V_var xx]\<^sub>c\<^sub>v)[x::=v]\<^sub>c\<^sub>v" + proof(rule subst_valid_infer_v[of \ _ _ _ z0 b c0 _ c, where z=z]) + show "\ ; \ ; \ \ v \ \ z0 : b | c0 \" using assms by auto + + show xf: "atom x \ \" using subtype_g_wf wfG_inside_fresh_suffix assms by metis + + show "atom x \ c" proof - + have "wfT \ \ \ (\ z : b | c \)" using subtype_wf[OF assms(2)] by auto + moreover have "x \ z" using assms(4) + using fresh_Pair not_self_fresh by blast + ultimately show ?thesis using xf wfT_fresh_c assms by presburger + qed + + show " \ ; \\\<^sub>w\<^sub>f ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ " + proof(subst append_g.simps,rule wfG_consI) + show *: \ \ ; \\\<^sub>w\<^sub>f \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using subtype_g_wf assms by metis + show \atom xx \ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\ using xxf fresh_prod9 by metis + show \ \ ; \\\<^sub>w\<^sub>f b1 \ using subtype_elims[OF assms(3)] wfT_wfC wfC_wf wfG_cons by metis + show "\ ; \ ; (xx, b1, TRUE) #\<^sub>\ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c1[z1::=V_var xx]\<^sub>c\<^sub>v " proof(rule wfT_wfC) + have "\ z1 : b1 | c1 \ = \ xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \ " using xxf fresh_prod9 type_eq_subst xxf2 fresh_prodN by metis + thus "\ ; \ ; \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f \ xx : b1 | c1[z1::=V_var xx]\<^sub>c\<^sub>v \ " using subtype_wfT[OF assms(3)] by metis + show "atom xx \ \' @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" using xxf fresh_prod9 by metis + qed + qed + + show "atom z0 \ v" using assms fresh_prod5 by auto + have "\ ; \ ; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>v " + apply(rule obtain_fresh[of "(z0,c0, \, c, z)"],rule subtype_valid[OF assms(2), THEN valid_flip], + (fastforce simp add: assms fresh_prodN)+) done + thus "\ ; \ ; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>c\<^sub>v " using subst_defs by auto + + show "atom z0 \ c" using assms fresh_prod5 by auto + show "\ ; \ ; ((xx, b1, c1[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ \') @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ c2[z2::=V_var xx]\<^sub>c\<^sub>v " + using subtype_valid assms(3) xxf xxf2 fresh_prodN append_g.simps subst_defs by metis + qed + + have xfw1: "atom z1 \ v \ atom x \ [ xx ]\<^sup>v \ x \ z1" + apply(intro conjI) + apply(simp add: assms xxf fresh_at_base fresh_prodN freshers fresh_x_neq)+ + using fresh_x_neq fresh_prodN xxf apply blast + using fresh_x_neq fresh_prodN assms by blast + + have xfw2: "atom z2 \ v \ atom x \ [ xx ]\<^sup>v \ x \ z2" + apply(auto simp add: assms xxf fresh_at_base fresh_prodN freshers) + by(insert xxf fresh_at_base fresh_prodN assms, fast+) + + have wf1: "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \)" proof - + have "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z1 : b1 | c1 \)[x::=v]\<^sub>\\<^sub>v" + using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis + moreover have "atom z1 \ (x,v)" using assms by auto + ultimately show ?thesis using subst_tv.simps by auto + qed + moreover have wf2: "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \)" proof - + have "wfT \ \ (\'[x::=v]\<^sub>\\<^sub>v@\) (\ z2 : b1 | c2 \)[x::=v]\<^sub>\\<^sub>v" using wf_subst(4) assms b_of.simps infer_v_v_wf subtype_wf subst_tv.simps subst_g_inside wfT_wf by metis + moreover have "atom z2 \ (x,v)" using assms by auto + ultimately show ?thesis using subst_tv.simps by auto + qed + moreover have "\ ; \ ; (xx, b1, c1[x::=v]\<^sub>c\<^sub>v[z1::=V_var xx]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v @ \ ) \ (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" proof - + have "xx \ x" using xxf fresh_Pair fresh_at_base by fast + hence "((xx, b1, subst_cv c1 z1 (V_var xx) ) #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v = (xx, b1, (subst_cv c1 z1 (V_var xx) )[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v)" + using subst_gv.simps by auto + moreover have "(c1[z1::=V_var xx]\<^sub>c\<^sub>v )[x::=v]\<^sub>c\<^sub>v = (c1[x::=v]\<^sub>c\<^sub>v) [z1::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw1 by metis + moreover have "c2[z2::=[ xx ]\<^sup>v]\<^sub>c\<^sub>v[x::=v]\<^sub>c\<^sub>v = (c2[x::=v]\<^sub>c\<^sub>v)[z2::=V_var xx]\<^sub>c\<^sub>v" using subst_cv_commute_full xfw2 by metis + ultimately show ?thesis using vd1 append_g.simps by metis + qed + moreover have "atom xx \ (\ , \ , \'[x::=v]\<^sub>\\<^sub>v@\, z1 , c1[x::=v]\<^sub>c\<^sub>v , z2 ,c2[x::=v]\<^sub>c\<^sub>v )" + using xxf fresh_prodN by metis + ultimately have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ \ z1 : b1 | c1[x::=v]\<^sub>c\<^sub>v \ \ \ z2 : b1 | c2[x::=v]\<^sub>c\<^sub>v \" + using subtype_baseI subst_defs by metis + thus ?thesis using subst_tv.simps assms by presburger +qed + +lemma subst_subtype_tau: + fixes v::v + assumes "\ ; \ ; \ \ v \ \" and + "\ ; \ ; \ \ \ \ (\ z : b | c \)" + "\ ; \ ; \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ \1 \ \2" and + "atom z \ (x,v)" + shows "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v@\ \ \1[x::=v]\<^sub>\\<^sub>v \ \2[x::=v]\<^sub>\\<^sub>v" +proof - + obtain z0 and b0 and c0 where zbc0: "\=(\ z0 : b0 | c0 \) \ atom z0 \ (c,x,v,z,\)" + using obtain_fresh_z by metis + obtain z1 and b1 and c1 where zbc1: "\1=(\ z1 : b1 | c1 \) \ atom z1 \ (x,v)" + using obtain_fresh_z by metis + obtain z2 and b2 and c2 where zbc2: "\2=(\ z2 : b2 | c2 \) \ atom z2 \ (x,v)" + using obtain_fresh_z by metis + + have "b0=b" using subtype_eq_base zbc0 assms by blast + + hence vinf: "\ ; \ ; \ \ v \ \ z0 : b | c0 \" using assms zbc0 by blast + have vsub: "\ ; \ ; \ \\ z0 : b | c0 \ \ \ z : b | c \" using assms zbc0 \b0=b\ by blast + have beq:"b1=b2" using subtype_eq_base + using zbc1 zbc2 assms by blast + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ \ z1 : b1 | c1 \[x::=v]\<^sub>\\<^sub>v \ \ z2 : b1 | c2 \[x::=v]\<^sub>\\<^sub>v" + proof(rule subst_subtype[OF vinf vsub] ) + show "\ ; \ ; \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ \ z1 : b1 | c1 \ \ \ z2 : b1 | c2 \" + using beq assms zbc1 zbc2 by auto + show "atom z \ (x, v) \ atom z0 \ (c, x, v, z, \) \ atom z1 \ (x, v) \ atom z2 \ (x, v)" + using zbc0 zbc1 zbc2 assms by blast + show "wfV \ \ \ v (b_of \)" using infer_v_wf assms by simp + qed + + thus ?thesis using zbc1 zbc2 \b1=b2\ assms by blast +qed + +lemma subtype_if1: + fixes v::v + assumes "P ; \ ; \ \ t1 \ t2" and "wfV P \ \ v (base_for_lit l)" and + "atom z1 \ v" and "atom z2 \ v" and "atom z1 \ t1" and "atom z2 \ t2" and "atom z1 \ \" and "atom z2 \ \" + shows "P ; \ ; \ \ \ z1 : b_of t1 | CE_val v == CE_val (V_lit l) IMP (c_of t1 z1) \ \ \ z2 : b_of t2 | CE_val v == CE_val (V_lit l) IMP (c_of t2 z2) \" +proof - + obtain z1' where t1: "t1 = \ z1' : b_of t1 | c_of t1 z1'\ \ atom z1' \ (z1,\,t1)" using obtain_fresh_z_c_of by metis + obtain z2' where t2: "t2 = \ z2' : b_of t2 | c_of t2 z2'\ \ atom z2' \ (z2,t2) " using obtain_fresh_z_c_of by metis + have beq:"b_of t1 = b_of t2" using subtype_eq_base2 assms by auto + + have c1: "(c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t1 z1" using c_of_switch t1 assms by simp + have c2: "(c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>c\<^sub>v = c_of t2 z2" using c_of_switch t2 assms by simp + + have "P ; \ ; \ \ \ z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[z1]\<^sup>v]\<^sub>v \ \ \ z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[z2]\<^sup>v]\<^sub>v \" + proof(rule subtype_if) + show \P ; \ ; \ \ \ z1' : b_of t1 | c_of t1 z1' \ \ \ z2' : b_of t1 | c_of t2 z2' \\ using t1 t2 assms beq by auto + show \ P ; \ ; \ \\<^sub>w\<^sub>f \ z1 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t1 z1')[z1'::=[ z1 ]\<^sup>v]\<^sub>v \ \ using wfT_wfT_if_rev assms subtype_wfT c1 subst_defs by metis + show \ P ; \ ; \ \\<^sub>w\<^sub>f \ z2 : b_of t1 | [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP (c_of t2 z2')[z2'::=[ z2 ]\<^sup>v]\<^sub>v \ \ using wfT_wfT_if_rev assms subtype_wfT c2 subst_defs beq by metis + show \atom z1 \ v\ using assms by auto + show \atom z1' \ \\ using t1 by auto + show \atom z1 \ c_of t1 z1'\ using t1 assms c_of_fresh by force + show \atom z2 \ c_of t2 z2'\ using t2 assms c_of_fresh by force + show \atom z2 \ v\ using assms by auto + qed + then show ?thesis using t1 t2 assms c1 c2 beq subst_defs by metis +qed + +section \Values\ + +lemma subst_infer_aux: + fixes \\<^sub>1::\ and v'::v + assumes "\ ; \ ; \ \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>1" and "\ ; \ ; \' \ v' \ \\<^sub>2" and "b_of \\<^sub>1 = b_of \\<^sub>2" + shows "\\<^sub>1 = (\\<^sub>2[x::=v]\<^sub>\\<^sub>v)" +proof - + obtain z1 and b1 where zb1: "\\<^sub>1 = (\ z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \) \ atom z1 \ ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v),v'[x::=v]\<^sub>v\<^sub>v)" + using infer_v_form_fresh[OF assms(1)] by fastforce + obtain z2 and b2 where zb2: "\\<^sub>2 = (\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \) \ atom z2 \ ((CE_val (v'[x::=v]\<^sub>v\<^sub>v), CE_val v,x,v),v')" + using infer_v_form_fresh [OF assms(2)] by fastforce + have beq: "b1 = b2" using assms zb1 zb2 by simp + + hence "(\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val v') \)[x::=v]\<^sub>\\<^sub>v = (\ z2 : b2 | C_eq (CE_val (V_var z2)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \)" + using subst_tv.simps subst_cv.simps subst_ev.simps forget_subst_vv[of x "V_var z2"] zb2 by force + also have "... = (\ z1 : b1 | C_eq (CE_val (V_var z1)) (CE_val (v'[x::=v]\<^sub>v\<^sub>v)) \)" + using type_e_eq[of z2 "CE_val (v'[x::=v]\<^sub>v\<^sub>v)"z1 b1 ] zb1 zb2 fresh_PairD(1) assms beq by metis + finally show ?thesis using zb1 zb2 by argo +qed + +lemma subst_t_b_eq: + fixes x::x and v::v + shows "b_of (\[x::=v]\<^sub>\\<^sub>v) = b_of \" +proof - + obtain z and b and c where "\ = \ z : b | c \ \ atom z \ (x,v)" + using has_fresh_z by blast + thus ?thesis using subst_tv.simps by simp +qed + +lemma fresh_g_fresh_v: + fixes x::x + assumes "atom x \ \" and "wfV \ \ \ v b" + shows "atom x \ v" + using assms wfV_supp wfX_wfY wfG_atoms_supp_eq fresh_def + by (metis wfV_x_fresh) + +lemma infer_v_fresh_g_fresh_v: + fixes x::x and \::\ and v::v + assumes "atom x \ \'@\" and "\ ; \ ; \ \ v \ \" + shows "atom x \ v" +proof - + have "atom x \ \" using fresh_suffix assms by auto + moreover have "wfV \ \ \ v (b_of \)" using infer_v_wf assms by auto + ultimately show ?thesis using fresh_g_fresh_v by metis +qed + +lemma infer_v_fresh_g_fresh_xv: + fixes xa::x and v::v and \::\ + assumes "atom xa \ \'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)" and "\ ; \ ; \ \ v \ \" + shows "atom xa \ (x,v)" +proof - + have "atom xa \ x" using assms fresh_in_g fresh_def by blast + moreover have "\'@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) = ((\'@(x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\GNil)@\)" using append_g.simps append_g_assoc by simp + moreover hence "atom xa \ v" using infer_v_fresh_g_fresh_v assms by metis + ultimately show ?thesis by auto +qed + +lemma wfG_subst_infer_v: + fixes v::v + assumes "\ ; \ \\<^sub>w\<^sub>f \' @ (x, b, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" and "\ ; \ ; \ \ v \ \" and "b_of \ = b" + shows "\ ; \\\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \ " + using wfG_subst_wfV infer_v_v_wf assms by auto + +lemma fresh_subst_gv_inside: + fixes \::\ + assumes "atom z \ \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \" and "atom z \ v" + shows "atom z \ \'[x::=v]\<^sub>\\<^sub>v@\" + unfolding fresh_append_g using fresh_append_g assms fresh_subst_gv fresh_GCons by metis + +lemma subst_t: + fixes x::x and b::b and xa::x + assumes "atom z \ x" and "atom z \ v" + shows "(\ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v'[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \) = (\ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v)" + using assms subst_vv.simps subst_tv.simps subst_cv.simps subst_cev.simps by auto + +lemma infer_l_fresh: + assumes "\ l \ \" + shows "atom x \ \" +proof - + have "[] ; {||} \\<^sub>w\<^sub>f GNil" using wfG_nilI wfTh_emptyI by auto + hence "[] ; {||} ; GNil \\<^sub>w\<^sub>f \" using assms infer_l_wf by auto + thus ?thesis using fresh_def wfT_supp by force +qed + +lemma subst_infer_v: + fixes v::v and v'::v + assumes "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \\<^sub>2" and + "\ ; \ ; \ \ v \ \\<^sub>1" and + "\ ; \ ; \ \ \\<^sub>1 \ (\ z0 : b\<^sub>1 | c0 \)" and "atom z0 \ (x,v)" + shows "\ ; \ ; (\'[x::=v]\<^sub>\\<^sub>v)@\ \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)" v' \\<^sub>2 avoiding: x v rule: infer_v.strong_induct) + case (infer_v_varI \ \ b c xa z) + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof(cases "x=xa") + case True + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \" + proof(rule infer_v_g_weakening) + show *:\ \ ; \ ; \ \ v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \\ + using infer_v_form infer_v_varI + by (metis True lookup_inside_unique_b lookup_inside_wf ms_fresh_all(32) subtype_eq_base type_e_eq) + show \toSet \ \ toSet (\'[x::=v]\<^sub>\\<^sub>v @ \)\ by simp + have "\ ; \ ; \ \\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis + thus \ \ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \ \ + using wfG_subst[OF infer_v_varI(3), of \' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \ v] subst_g_inside infer_v_varI by metis + qed + + thus ?thesis using subst_vv.simps True by simp + next + case False + then obtain c' where c: "Some (b, c') = lookup (\'[x::=v]\<^sub>\\<^sub>v @ \) xa" using lookup_subst2 infer_v_varI by metis + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ [ xa ]\<^sup>v \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ xa ]\<^sup>v ]\<^sup>c\<^sup>e \" + proof + have "\ ; \ ; \ \\<^sub>w\<^sub>f v : b\<^sub>1" using infer_v_wf subtype_eq_base2 b_of.simps infer_v_varI by metis + thus "\ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \" using infer_v_varI + using wfG_subst[OF infer_v_varI(3), of \' x b\<^sub>1 "c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v" \ v] subst_g_inside infer_v_varI by metis + show "atom z \ xa" using infer_v_varI by auto + show "Some (b, c') = lookup (\'[x::=v]\<^sub>\\<^sub>v @ \) xa" using c by auto + show "atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)" by (fresh_mth add: infer_v_varI fresh_subst_gv_inside) + qed + then show ?thesis using subst_vv.simps False by simp + qed + thus ?case using subst_t fresh_prodN infer_v_varI by metis +next + case (infer_v_litI \ \ l \) + show ?case unfolding subst_vv.simps proof + show "\ ; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \" using wfG_subst_infer_v infer_v_litI subtype_eq_base2 b_of.simps by metis + have "atom x \ \" using infer_v_litI infer_l_fresh by metis + thus "\ l \ \[x::=v]\<^sub>\\<^sub>v" using infer_v_litI type_v_subst_fresh by simp + qed +next + case (infer_v_pairI z v1 v2 \ \ t1 t2) + have " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ + \ \ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v \ \ z : [ b_of t1[x::=v]\<^sub>\\<^sub>v , b_of + t2[x::=v]\<^sub>\\<^sub>v ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1[x::=v]\<^sub>v\<^sub>v , v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>v ]\<^sup>c\<^sup>e \" + proof + show \atom z \ (v1[x::=v]\<^sub>v\<^sub>v, v2[x::=v]\<^sub>v\<^sub>v)\ by (fresh_mth add: infer_v_pairI) + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)\ by (fresh_mth add: infer_v_pairI fresh_subst_gv_inside) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ t1[x::=v]\<^sub>\\<^sub>v\ using infer_v_pairI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ t2[x::=v]\<^sub>\\<^sub>v\ using infer_v_pairI by metis + qed + then show ?case using subst_vv.simps subst_tv.simps infer_v_pairI b_of_subst by simp +next + case (infer_v_consI s dclist \ dc tc \ va tv z) + + have " \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ (V_cons s dc va[x::=v]\<^sub>v\<^sub>v) \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show td:\AF_typedef s dclist \ set \\ using infer_v_consI by auto + show dc:\(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ va[x::=v]\<^sub>v\<^sub>v \ tv[x::=v]\<^sub>\\<^sub>v\ using infer_v_consI by auto + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[x::=v]\<^sub>\\<^sub>v\ + using subst_subtype_tau infer_v_consI by metis + moreover have "atom x \ tc" using wfTh_lookup_supp_empty[OF td dc] infer_v_wf infer_v_consI fresh_def by fast + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc\ by simp + show \atom z \ va[x::=v]\<^sub>v\<^sub>v\ using infer_v_consI by auto + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \)\ by (fresh_mth add: infer_v_consI fresh_subst_gv_inside) + qed + thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_consI by metis + +next + case (infer_v_conspI s bv dclist \ dc tc \ va tv b z) + have "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ (V_consp s dc b va[x::=v]\<^sub>v\<^sub>v) \ \ z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show td:\AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show dc:\(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ va[x::=v]\<^sub>v\<^sub>v \ tv[x::=v]\<^sub>\\<^sub>v\ using infer_v_conspI by metis + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v\ + using subst_subtype_tau infer_v_conspI by metis + moreover have "atom x \ tc[bv::=b]\<^sub>\\<^sub>b" proof - + have "supp tc \ { atom bv }" using wfTh_poly_lookup_supp infer_v_conspI wfX_wfY by metis + hence "atom x \ tc" using x_not_in_b_set + using fresh_def by fastforce + moreover have "atom x \ b" using x_fresh_b by auto + ultimately show ?thesis using fresh_subst_if subst_b_\_def by metis + qed + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ tv[x::=v]\<^sub>\\<^sub>v \ tc[bv::=b]\<^sub>\\<^sub>b\ by simp + show \atom z \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \, va[x::=v]\<^sub>v\<^sub>v, b)\ proof - + have "atom z \ va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis + moreover have "atom z \ \'[x::=v]\<^sub>\\<^sub>v @ \" using fresh_subst_gv_inside infer_v_conspI by metis + ultimately show ?thesis using fresh_prodN infer_v_conspI by metis + qed + show \atom bv \ (\, \, \'[x::=v]\<^sub>\\<^sub>v @ \, va[x::=v]\<^sub>v\<^sub>v, b)\ proof - + have "atom bv \ va[x::=v]\<^sub>v\<^sub>v" using fresh_subst_v_if infer_v_conspI subst_v_v_def by metis + moreover have "atom bv \ \'[x::=v]\<^sub>\\<^sub>v @ \" using fresh_subst_gv_inside infer_v_conspI by metis + ultimately show ?thesis using fresh_prodN infer_v_conspI by metis + qed + show "\ ; \ \\<^sub>w\<^sub>f b" using infer_v_conspI by auto + qed + thus ?case using subst_vv.simps subst_t[of z x v ] infer_v_conspI by metis + +qed + +lemma subst_infer_check_v: + fixes v::v and v'::v + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and + "check_v \ \ (\'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)) v' \\<^sub>2" and + "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "check_v \ \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) (v'[x::=v]\<^sub>v\<^sub>v) (\\<^sub>2[x::=v]\<^sub>\\<^sub>v)" +proof - + obtain \\<^sub>2' where t2: "infer_v \ \ (\' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) v' \\<^sub>2' \ \ ; \ ; (\' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \) \ \\<^sub>2' \ \\<^sub>2" + using check_v_elims assms by blast + hence "infer_v \ \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) (v'[x::=v]\<^sub>v\<^sub>v) (\\<^sub>2'[x::=v]\<^sub>\\<^sub>v)" + using subst_infer_v[OF _ assms(1) assms(3) assms(4)] by blast + moreover hence "\; \ ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\\<^sub>2'[x::=v]\<^sub>\\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using subst_subtype assms t2 by (meson subst_subtype_tau subtype_trans) + ultimately show ?thesis using check_v.intros by blast +qed + +lemma type_veq_subst[simp]: + assumes "atom z \ (x,v)" + shows "\ z : b | CE_val (V_var z) == CE_val v' \[x::=v]\<^sub>\\<^sub>v = \ z : b | CE_val (V_var z) == CE_val v'[x::=v]\<^sub>v\<^sub>v \" + using assms by auto + +lemma subst_infer_v_form: + fixes v::v and v'::v and \::\ + assumes "\ ; \ ; \ \ v \ \\<^sub>1" and + "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \\<^sub>2" and "b= b_of \\<^sub>2" + "\ ; \ ; \ \ \\<^sub>1 \ (\ z0 : b\<^sub>1 | c0 \)" and "atom z0 \ (x,v)" and "atom z3' \ (x,v,v', \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) )" + shows \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : b | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \\ +proof - + have "\ ; \ ; \'@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\) \ v' \ \ z3' : b_of \\<^sub>2 | C_eq (CE_val (V_var z3')) (CE_val v') \" + proof(rule infer_v_form4) + show \ \ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \\<^sub>2\ using assms by metis + show \atom z3' \ (v', \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \)\ using assms fresh_prodN by metis + show \b_of \\<^sub>2 = b_of \\<^sub>2\ by auto + qed + hence \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : b_of \\<^sub>2 | CE_val (V_var z3') == CE_val v' \[x::=v]\<^sub>\\<^sub>v\ + using subst_infer_v assms by metis + thus ?thesis using type_veq_subst fresh_prodN assms by metis +qed + +section \Expressions\ + +text \ + For operator, fst and snd cases, we use elimination to get one or more values, apply the substitution lemma for values. The types always have +the same form and are equal under substitution. + For function application, the subst value is a subtype of the value which is a subtype of the argument. The return of the function is the same + under substitution. +\ + +text \ Observe a similar pattern for each case \ + +lemma subst_infer_e: + fixes v::v and e::e and \'::\ + assumes + "\ ; \ ; \ ; G ; \ \ e \ \\<^sub>2" and "G = (\'@((x,b\<^sub>1,subst_cv c0 z0 (V_var x))#\<^sub>\\))" + "\ ; \ ; \ \ v \ \\<^sub>1" and + "\; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "\ ; \ ; \ ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) ; (\[x::=v]\<^sub>\\<^sub>v) \ (subst_ev e x v ) \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct avoiding: x v rule: infer_e.strong_induct) + case (infer_e_valI \ \ \'' \ \ v' \) + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_val (v'[x::=v]\<^sub>v\<^sub>v)) \ \[x::=v]\<^sub>\\<^sub>v" + proof + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v" using wfD_subst infer_e_valI subtype_eq_base2 + by (metis b_of.simps infer_v_v_wf subst_g_inside_simple wfD_wf wf_subst(11)) + show "\\\<^sub>w\<^sub>f \" using infer_e_valI by auto + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" using subst_infer_v infer_e_valI using wfD_subst infer_e_valI subtype_eq_base2 + by metis + qed + thus ?case using subst_ev.simps by simp +next + case (infer_e_plusI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op Plus v1 v2, CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v , CE_op Plus [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_plusI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : B_int | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : B_int | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = B_int \ b2' = B_int" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op Plus (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ + using infer_e_plusI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_plusI by blast + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_plusI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_plusI z2 bb by metis + show \atom z3' \ AE_op Plus v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_prod6 * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : B_int | CE_val (V_var z3') == CE_op Plus ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_int | CE_val (V_var z3') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + by(subst subst_tv.simps,auto simp add: * ) + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_leqI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op LEq v1 v2, CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op LEq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_leqI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : B_int | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : B_int | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = B_int \ b2' = B_int" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op LEq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_leqI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_leqI(2) by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_int | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_leqI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : B_int | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_leqI z2 bb by metis + show \atom z3' \ AE_op LEq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : B_bool | CE_val (V_var z3') == CE_op LEq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bool | CE_val (V_var z3') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_eqI \ \ \'' \ \ v1 z1 bb c1 v2 z2 c2 z3) + + hence z3f: "atom z3 \ CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using e.fresh ce.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_op Eq v1 v2, CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , CE_op Eq [v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e [v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e , AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v,\'[x::=v]\<^sub>\\<^sub>v @ \ )" + using obtain_fresh by metis + hence **:"(\ z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_eqI fresh_Pair z3f by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z1 : bb | c1 \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + obtain z2' b2' c2' where z2:"atom z2' \ (x,v) \ \ z2 : bb | c2 \ = \ z2' : b2' | c2' \" using obtain_fresh_z by metis + + have bb:"b1' = bb \ b2' = bb" using z1 z2 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_op Eq (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_eqI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_eqI(2) by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1' : bb | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_eqI z1 bb by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2' : bb | c2'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_eqI z2 bb by metis + show \atom z3' \ AE_op Eq v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed + moreover have "\ z3' : B_bool | CE_val (V_var z3') == CE_op Eq ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bool | CE_val (V_var z3') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_appI \ \ \'' \ \ f x' b c \' s' v' \) + + hence "x \ x'" using \atom x' \ \''\ using wfG_inside_x_neq wfX_wfY by metis + + show ?case proof(subst subst_ev.simps,rule) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using infer_e_appI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_appI by metis + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x' b c \' s'))) = lookup_fun \ f\ using infer_e_appI by metis + + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b | c \[x::=v]\<^sub>\\<^sub>v\ proof(rule subst_infer_check_v ) + show "\ ; \ ; \ \ v \ \\<^sub>1" using infer_e_appI by metis + show "\ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ x' : b | c \" using infer_e_appI by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_appI by metis + show "atom z0 \ (x, v)" using infer_e_appI by metis + qed + moreover have "atom x \ c" using wfPhi_f_simple_supp_c infer_e_appI fresh_def \x\x'\ + atom_eq_iff empty_iff infer_e_appI.hyps insert_iff subset_singletonD by metis + + moreover hence "atom x \ \ x' : b | c \" using \.fresh supp_b_empty fresh_def by blast + ultimately show \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b | c \\ using forget_subst_tv by metis + + have *: "atom x' \ (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appI check_v_wf by blast + + show \atom x' \ (\, \, \, \'[x::=v]\<^sub>\\<^sub>v @ \, \[x::=v]\<^sub>\\<^sub>v, v'[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)\ + apply(unfold fresh_prodN, intro conjI) + apply (fresh_subst_mth_aux add: infer_e_appI fresh_subst_gv wfD_wf subst_g_inside) + using infer_e_appI fresh_subst_gv wfD_wf subst_g_inside apply metis + using infer_e_appI fresh_subst_dv_if apply metis + done + + have "supp \' \ { atom x' } \ supp \" using infer_e_appI wfT_supp wfPhi_f_simple_wfT + by (meson infer_e_appI.hyps(2) le_supI1 wfPhi_f_simple_supp_t) + hence "atom x \ \'" using \x\x'\ fresh_def supp_at_base x_not_in_b_set by fastforce + thus \\'[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \[x::=v]\<^sub>\\<^sub>v\ using subst_tv_commute infer_e_appI subst_defs by metis + qed +next + case (infer_e_appPI \ \ \'' \ \ b' f bv x' b c \' s' v' \) + + hence "x \ x'" using \atom x' \ \''\ using wfG_inside_x_neq wfX_wfY by metis + + show ?case proof(subst subst_ev.simps,rule) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using infer_e_appPI wfD_subst subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_appPI(4) by auto + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI(5) by auto + show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b c \' s'))) = lookup_fun \ f" using infer_e_appPI(6) by auto + + show "\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \" proof - + have \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \[x::=v]\<^sub>\\<^sub>v\ proof(rule subst_infer_check_v ) + show "\ ; \ ; \ \ v \ \\<^sub>1" using infer_e_appPI by metis + show "\ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" using infer_e_appPI subst_defs by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_appPI by metis + show "atom z0 \ (x, v)" using infer_e_appPI by metis + qed + moreover have "atom x \ c" proof - + have "supp c \ {atom x', atom bv}" using wfPhi_f_poly_supp_c[OF infer_e_appPI(6)] infer_e_appPI by metis + thus ?thesis unfolding fresh_def using \x\x'\ atom_eq_iff by auto + qed + moreover hence "atom x \ \ x' : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" using \.fresh supp_b_empty fresh_def subst_b_fresh_x + by (metis subst_b_c_def) + ultimately show ?thesis using forget_subst_tv subst_defs by metis + qed + have "supp \' \ { atom x', atom bv }" using wfPhi_f_poly_supp_t infer_e_appPI by metis + hence "atom x \ \'" using fresh_def \x\x'\ by force + hence *:"atom x \ \'[bv::=b']\<^sub>\\<^sub>b" using subst_b_fresh_x subst_b_\_def by metis + have "atom x' \ (x,v)" using infer_v_fresh_g_fresh_xv infer_e_appPI check_v_wf by blast + thus "atom x' \ \'[x::=v]\<^sub>\\<^sub>v @ \" using infer_e_appPI fresh_subst_gv wfD_wf subst_g_inside fresh_Pair by metis + show "\'[bv::=b']\<^sub>b[x'::=v'[x::=v]\<^sub>v\<^sub>v]\<^sub>v = \[x::=v]\<^sub>\\<^sub>v" using infer_e_appPI subst_tv_commute[OF * ] subst_defs by metis + show "atom bv \ (\, \, \, \'[x::=v]\<^sub>\\<^sub>v @ \, \[x::=v]\<^sub>\\<^sub>v, b', v'[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)" + by (fresh_mth add: infer_e_appPI fresh_subst_gv_inside) + qed +next + case (infer_e_fstI \ \ \'' \ \ v' z' b1 b2 c z) + + hence zf: "atom z \ CE_fst [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_fst v', CE_fst [v']\<^sup>c\<^sup>e , AE_fst v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ )" using obtain_fresh by auto + hence **:"(\ z : b1 | CE_val (V_var z) == CE_fst [v']\<^sup>c\<^sup>e \) = \ z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_fstI(4) fresh_Pair zf by metis + + obtain z1' b1' c1' where z1:"atom z1' \ (x,v) \ \ z' : B_pair b1 b2 | c \ = \ z1' : b1' | c1' \" using obtain_fresh_z by metis + + have bb:"b1' = B_pair b1 b2" using z1 \.eq_iff by metis + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_fst v'[x::=v]\<^sub>v\<^sub>v) \ \ z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_fstI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_fstI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_fstI z1 bb by metis + + show \atom z3' \ AE_fst v'[x::=v]\<^sub>v\<^sub>v \ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : b1 | CE_val (V_var z3') == CE_fst [v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e \ = \ z3' : b1 | CE_val (V_var z3') == CE_fst [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_sndI \ \ \'' \ \ v' z' b1 b2 c z) + hence zf: "atom z \ CE_snd [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,AE_snd v', CE_snd [v']\<^sup>c\<^sup>e , AE_snd v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ ,v', \'')" using obtain_fresh by auto + hence **:"(\ z : b2 | CE_val (V_var z) == CE_snd [v']\<^sup>c\<^sup>e \) = \ z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_sndI(4) fresh_Pair zf by metis + + obtain z1' b2' c1' where z1:"atom z1' \ (x,v) \ \ z' : B_pair b1 b2 | c \ = \ z1' : b2' | c1' \" using obtain_fresh_z by metis + + have bb:"b2' = B_pair b1 b2" using z1 \.eq_iff by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_snd (v'[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_sndI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_sndI by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z1' : B_pair b1 b2 | c1'[x::=v]\<^sub>c\<^sub>v \\ using subst_tv.simps subst_infer_v infer_e_sndI z1 bb by metis + + show \atom z3' \ AE_snd v'[x::=v]\<^sub>v\<^sub>v \ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using * by auto + qed + moreover have "\ z3' : b2 | CE_val (V_var z3') == CE_snd ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : b2 | CE_val (V_var z3') == CE_snd [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + by(subst subst_tv.simps, auto simp add: fresh_prodN *) + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_lenI \ \ \'' \ \ v' z' c z) + hence zf: "atom z \ CE_len [v']\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + obtain z3'::x where *:"atom z3' \ (x,v,AE_len v', CE_len [v']\<^sup>c\<^sup>e , AE_len v'[x::=v]\<^sub>v\<^sub>v ,\'[x::=v]\<^sub>\\<^sub>v @ \ , \'',v')" using obtain_fresh by auto + hence **:"(\ z : B_int | CE_val (V_var z) == CE_len [v']\<^sup>c\<^sup>e \) = \ z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \" + using type_e_eq infer_e_lenI fresh_Pair zf by metis + + have ***: "\ ; \ ; \'' \ v' \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \" + using infer_e_lenI infer_v_form3[OF infer_e_lenI(3), of z3'] b_of.simps * fresh_Pair by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_len (v'[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_lenI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_lenI by metis + + have \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v' \[x::=v]\<^sub>\\<^sub>v\ + proof(rule subst_infer_v) + show \ \ ; \ ; \ \ v \ \\<^sub>1\ using infer_e_lenI by metis + show \ \ ; \ ; \' @ (x, b\<^sub>1, c0[z0::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ v' \ \ z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v' ]\<^sup>c\<^sup>e \\ using *** infer_e_lenI by metis + show "\ ; \ ; \ \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" using infer_e_lenI by metis + show "atom z0 \ (x, v)" using infer_e_lenI by metis + qed + + thus \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v'[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v'[x::=v]\<^sub>v\<^sub>v \\ + using subst_tv.simps subst_ev.simps fresh_Pair * fresh_prodN subst_vv.simps by auto + + show \atom z3' \ AE_len v'[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_Pair * by metis + qed + + moreover have "\ z3' : B_int | CE_val (V_var z3') == CE_len ([v'[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_int | CE_val (V_var z3') == CE_len [v']\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + + ultimately show ?case using subst_ev.simps * ** by metis +next + case (infer_e_mvarI \ \ \'' \ \ u \) + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_mvar u) \ \[x::=v]\<^sub>\\<^sub>v" + proof + show \ \ ; \\\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \\ proof - + have "wfV \ \ \ v (b_of \\<^sub>1)" using infer_v_wf infer_e_mvarI by auto + moreover have "b_of \\<^sub>1 = b\<^sub>1" using subtype_eq_base2 infer_e_mvarI b_of.simps by simp + ultimately show ?thesis using wf_subst(3)[OF infer_e_mvarI(1), of \' x b\<^sub>1 "c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v" \ v] infer_e_mvarI subst_g_inside by metis + qed + show \ \\\<^sub>w\<^sub>f \ \ using infer_e_mvarI by auto + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_mvarI subtype_eq_base2 b_of.simps by metis + show \(u, \[x::=v]\<^sub>\\<^sub>v) \ setD \[x::=v]\<^sub>\\<^sub>v\ using infer_e_mvarI subst_dv_member by metis + qed + moreover have " (AE_mvar u) = (AE_mvar u)[x::=v]\<^sub>e\<^sub>v" using subst_ev.simps by auto + ultimately show ?case by auto + +next + case (infer_e_concatI \ \ \'' \ \ v1 z1 c1 v2 z2 c2 z3) + + hence zf: "atom z3 \ CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e" using ce.fresh e.fresh opp.fresh by metis + + obtain z3'::x where *:"atom z3' \ (x,v,v1,v2,AE_concat v1 v2, CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e , AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v) ,\'[x::=v]\<^sub>\\<^sub>v @ \ , \'',v1 , v2)" using obtain_fresh by auto + + hence **:"(\ z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \) = \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_concatI fresh_Pair zf by metis + have zfx: "atom x \ z3'" using fresh_at_base fresh_prodN * by auto + + have v1: "\ ; \ ; \'' \ v1 \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v1 \" + using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis + have v2: "\ ; \ ; \'' \ v2 \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val v2 \" + using infer_e_concatI infer_v_form3 b_of.simps * fresh_Pair by metis + + have "\ ; \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ ; \[x::=v]\<^sub>\\<^sub>v \ (AE_concat (v1[x::=v]\<^sub>v\<^sub>v) (v2[x::=v]\<^sub>v\<^sub>v)) \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_concatI subtype_eq_base2 b_of.simps by metis + show \ \\\<^sub>w\<^sub>f \ \ by(simp add: infer_e_concatI) + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val (v1[x::=v]\<^sub>v\<^sub>v) \\ + using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z3' : B_bitvec | CE_val (V_var z3') == CE_val (v2[x::=v]\<^sub>v\<^sub>v) \\ + using subst_infer_v_form infer_e_concatI fresh_prodN * b_of.simps by metis + show \atom z3' \ AE_concat v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using fresh_Pair * by metis + show \atom z3' \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_Pair * by metis + qed + + moreover have "\ z3' : B_bitvec | CE_val (V_var z3') == CE_concat ([v1[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) ([v2[x::=v]\<^sub>v\<^sub>v]\<^sup>c\<^sup>e) \ = \ z3' : B_bitvec | CE_val (V_var z3') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \[x::=v]\<^sub>\\<^sub>v" + using subst_tv.simps subst_ev.simps * by auto + + ultimately show ?case using subst_ev.simps ** * by metis +next + case (infer_e_splitI \ \ \'' \ \ v1 z1 c1 v2 z2 z3) + hence *:"atom z3 \ (x,v)" using fresh_Pair by auto + have \x \ z3 \ using infer_e_splitI by force + have "\ ; \ ; \ ; (\'[x::=v]\<^sub>\\<^sub>v @ \) ; \[x::=v]\<^sub>\\<^sub>v \ (AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v) \ + \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND + [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e \" + proof + show \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wfD_subst infer_e_splitI subtype_eq_base2 b_of.simps by metis + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_splitI by auto + have \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1 : B_bitvec | c1 \[x::=v]\<^sub>\\<^sub>v\ + using subst_infer_v infer_e_splitI by metis + thus \ \ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ \ \ v1[x::=v]\<^sub>v\<^sub>v \ \ z1 : B_bitvec | c1[x::=v]\<^sub>c\<^sub>v \\ + using infer_e_splitI subst_tv.simps fresh_Pair by metis + have \x \ z2 \ using infer_e_splitI by force + have "(\ z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \) = + (\ z2 : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) + AND ([ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \[x::=v]\<^sub>\\<^sub>v)" + unfolding subst_cv.simps subst_cev.simps subst_vv.simps using \x \ z2\ infer_e_splitI fresh_Pair by simp + thus \\ ; \ ; \'[x::=v]\<^sub>\\<^sub>v @ + \ \ v2[x::=v]\<^sub>v\<^sub>v \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e + AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \\ + using infer_e_splitI subst_infer_check_v fresh_Pair by metis + + show \atom z1 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + show \atom z2 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + show \atom z3 \ AE_split v1[x::=v]\<^sub>v\<^sub>v v2[x::=v]\<^sub>v\<^sub>v\ using infer_e_splitI fresh_subst_vv_if by auto + + show \atom z3 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + show \atom z2 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + show \atom z1 \ \'[x::=v]\<^sub>\\<^sub>v @ \\ using fresh_subst_gv_inside infer_e_splitI by metis + qed + thus ?case apply (subst subst_tv.simps) + using infer_e_splitI fresh_Pair apply metis + unfolding subst_tv.simps subst_ev.simps subst_cv.simps subst_cev.simps subst_vv.simps * + using \x \ z3\ by simp +qed + +lemma infer_e_uniqueness: + assumes "\ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>2" + shows "\\<^sub>1 = \\<^sub>2" + using assms proof(nominal_induct rule:e.strong_induct) + case (AE_val x) + then show ?case using infer_e_elims(7)[OF AE_val(1)] infer_e_elims(7)[OF AE_val(2)] infer_v_uniqueness by metis +next + case (AE_app f v) + + obtain x1 b1 c1 s1' \1' where t1: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f \ \\<^sub>1 = \1'[x1::=v]\<^sub>\\<^sub>v" using infer_e_app2E[OF AE_app(1)] by metis + moreover obtain x2 b2 c2 s2' \2' where t2: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f \ \\<^sub>2 = \2'[x2::=v]\<^sub>\\<^sub>v" using infer_e_app2E[OF AE_app(2)] by metis + + have "\1'[x1::=v]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>v" using t1 and t2 fun_ret_unique by metis + thus ?thesis using t1 t2 by auto +next + case (AE_appP f b v) + obtain bv1 x1 b1 c1 s1' \1' where t1: "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f \ \\<^sub>1 = \1'[bv1::=b]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v" using infer_e_appP2E[OF AE_appP(1)] by metis + moreover obtain bv2 x2 b2 c2 s2' \2' where t2: "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f \ \\<^sub>2 = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>v" using infer_e_appP2E[OF AE_appP(2)] by metis + + have "\1'[bv1::=b]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>v" using t1 and t2 fun_poly_ret_unique by metis + thus ?thesis using t1 t2 by auto +next + case (AE_op opp v1 v2) + show ?case proof(rule opp.exhaust) + assume "opp = plus" + hence "\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(11)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>1\ ] infer_e_elims(11)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Plus v1 v2 \ \\<^sub>2\ ] + by force + next + assume "opp = leq" + hence "opp = LEq" using opp.strong_exhaust by auto + hence "\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(12)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>1\ ] infer_e_elims(12)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op LEq v1 v2 \ \\<^sub>2\ ] + by force + next + assume "opp = eq" + hence "opp = Eq" using opp.strong_exhaust by auto + hence "\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>1" and "\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>2" using AE_op by auto + thus ?thesis using infer_e_elims(25)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>1\ ] infer_e_elims(25)[OF \\ ; \ ; \ ; GNil ; \ \ AE_op Eq v1 v2 \ \\<^sub>2\ ] + by force + qed +next + case (AE_concat v1 v2) + + obtain z3::x where t1:"\\<^sub>1 = \ z3 : B_bitvec | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ \ atom z3 \ v1 \ atom z3 \ v2 " using infer_e_elims(18)[OF AE_concat(1)] by metis + obtain z3'::x where t2:"\\<^sub>2 = \ z3' : B_bitvec | [ [ z3' ]\<^sup>v ]\<^sup>c\<^sup>e == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ \ atom z3' \ v1 \ atom z3' \ v2" using infer_e_elims(18)[OF AE_concat(2)] by metis + + thus ?case using t1 t2 type_e_eq ce.fresh by metis + +next + case (AE_fst v) + + obtain z1 and b1 where "\\<^sub>1 = \ z1 : b1 | CE_val (V_var z1) == (CE_fst [v]\<^sup>c\<^sup>e) \" using infer_v_form AE_fst by auto + + obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where + f1: "\\<^sub>2 = \ xx : bb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxa : B_pair bb bba | cc \ \ atom xx \ v" + using infer_e_elims(8)[OF AE_fst(2)] by metis + obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where + f2: "\\<^sub>1 = \ xxb : bbb | CE_val (V_var xxb) == CE_fst [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxc : B_pair bbb bbc | cca \ \ atom xxb \ v" + using infer_e_elims(8)[OF AE_fst(1)] by metis + then have "B_pair bb bba = B_pair bbb bbc" + using f1 by (metis (no_types) b_of.simps infer_v_uniqueness) + then have "\ xx : bbb | CE_val (V_var xx) == CE_fst [v]\<^sup>c\<^sup>e \ = \\<^sub>2" + using f1 by auto + then show ?thesis + using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple) +next + case (AE_snd v) + obtain xx :: x and bb :: b and xxa :: x and bba :: b and cc :: c where + f1: "\\<^sub>2 = \ xx : bba | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxa : B_pair bb bba | cc \ \ atom xx \ v" + using infer_e_elims(9)[OF AE_snd(2)] by metis + obtain xxb :: x and bbb :: b and xxc :: x and bbc :: b and cca :: c where + f2: "\\<^sub>1 = \ xxb : bbc | CE_val (V_var xxb) == CE_snd [v]\<^sup>c\<^sup>e \ \ \ ; \ ; GNil\\<^sub>w\<^sub>f \ \ \ ; \ ; GNil \ v \ \ xxc : B_pair bbb bbc | cca \ \ atom xxb \ v" + using infer_e_elims(9)[OF AE_snd(1)] by metis + then have "B_pair bb bba = B_pair bbb bbc" + using f1 by (metis (no_types) b_of.simps infer_v_uniqueness) + then have "\ xx : bbc | CE_val (V_var xx) == CE_snd [v]\<^sup>c\<^sup>e \ = \\<^sub>2" + using f1 by auto + then show ?thesis + using f2 by (meson ce.fresh fresh_GNil type_e_eq wfG_x_fresh_in_v_simple) +next + case (AE_mvar x) + then show ?case using infer_e_elims(10)[OF AE_mvar(1)] infer_e_elims(10)[OF AE_mvar(2)] wfD_unique by metis +next + case (AE_len x) + then show ?case using infer_e_elims(16)[OF AE_len(1)] infer_e_elims(16)[OF AE_len(2)] by force +next + case (AE_split x1a x2) + then show ?case using infer_e_elims(22)[OF AE_split(1)] infer_e_elims(22)[OF AE_split(2)] by force +qed + +section \Statements\ + +lemma subst_infer_check_v1: + fixes v::v and v'::v and \::\ + assumes "\ = \\<^sub>1@((x,b\<^sub>1,c0[z0::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>2)" and + "\ ; \ ; \\<^sub>2 \ v \ \\<^sub>1" and + "\ ; \ ; \ \ v' \ \\<^sub>2" and + "\ ; \ ; \\<^sub>2 \ \\<^sub>1 \ \ z0 : b\<^sub>1 | c0 \" and "atom z0 \ (x,v)" + shows "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \\<^sub>2[x::=v]\<^sub>\\<^sub>v" + using subst_g_inside check_v_wf assms subst_infer_check_v by metis + +lemma infer_v_c_valid: + assumes " \ ; \ ; \ \ v \ \" and "\ ; \ ; \ \ \ \ \ z : b | c \" + shows \\ ; \ ; \ \ c[z::=v]\<^sub>c\<^sub>v \ +proof - + obtain z1 and b1 and c1 where *:"\ = \ z1 : b1 | c1 \ \ atom z1 \ (c,v,\)" using obtain_fresh_z by metis + then have "b1 = b" using assms subtype_eq_base by metis + moreover then have "\ ; \ ; \ \ v \ \ z1 : b | c1 \" using assms * by auto + + moreover have "\ ; \ ; (z1, b, c1) #\<^sub>\ \ \ c[z::=[ z1 ]\<^sup>v]\<^sub>c\<^sub>v " proof - + have "\ ; \ ; (z1, b, c1[z1::=[ z1 ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c[z::=[ z1 ]\<^sup>v]\<^sub>v " + using subtype_valid[OF assms(2), of z1 z1 b c1 z c ] * fresh_prodN \b1 = b\ by metis + moreover have "c1[z1::=[ z1 ]\<^sup>v]\<^sub>v = c1" using subst_v_v_def by simp + ultimately show ?thesis using subst_v_c_def by metis + qed + ultimately show ?thesis using * fresh_prodN subst_valid_simple by metis +qed + +text \ Substitution Lemma for Statements \ + +lemma subst_infer_check_s: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and + \\<^sub>1::\ and \\<^sub>2::\ and css::branch_list + assumes "\ ; \ ; \\<^sub>1 \ v \ \" and "\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" and + "atom z \ (x, v)" + shows "\ ; \ ; \ ; \; \ \ s \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + and + "\ ; \ ; \ ; \; \; tid ; cons ; const ; v' \ cs \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v; + tid ; cons ; const ; v'[x::=v]\<^sub>v\<^sub>v \ cs[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + and + "\ ; \ ; \ ; \; \; tid ; dclist ; v' \ css \ \' \ + \ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1)) \ + \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v; tid ; dclist ; v'[x::=v]\<^sub>v\<^sub>v \ + subst_branchlv css x v \ \'[x::=v]\<^sub>\\<^sub>v" + + using assms proof(nominal_induct \' and \' and \' avoiding: x v arbitrary: \\<^sub>2 and \\<^sub>2 and \\<^sub>2 + rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \ \ v' \' \'') + + have sg: " \[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using check_valI by subst_mth + have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ (AS_val (v'[x::=v]\<^sub>v\<^sub>v)) \ \''[x::=v]\<^sub>\\<^sub>v" proof + have "\ ; \ ; \\<^sub>1\\<^sub>w\<^sub>f v : b " using infer_v_v_wf subtype_eq_base2 b_of.simps check_valI by metis + thus \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v\\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v\ using wf_subst(15) check_valI by auto + show \ \\\<^sub>w\<^sub>f \ \ using check_valI by auto + show \ \ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v\ proof(subst sg, rule subst_infer_v) + show "\ ; \ ; \\<^sub>1 \ v \ \" using check_valI by auto + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ v' \ \'" using check_valI by metis + show "\ ; \ ; \\<^sub>1 \ \ \ \ z: b | c \" using check_valI by auto + show "atom z \ (x, v)" using check_valI by auto + qed + show \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v \ \''[x::=v]\<^sub>\\<^sub>v\ using subst_subtype_tau check_valI sg by metis + qed + + thus ?case using Typing.check_valI subst_sv.simps sg by auto +next + case (check_letI xa \ \ \ \ \ ea \a za sa ba ca) + have *:"(AS_let xa ea sa)[x::=v]\<^sub>s\<^sub>v=(AS_let xa (ea[x::=v]\<^sub>e\<^sub>v) sa[x::=v]\<^sub>s\<^sub>v)" + using subst_sv.simps \ atom xa \ x\ \ atom xa \ v\ by auto + show ?case unfolding * proof + + show "atom xa \ (\,\,\,\[x::=v]\<^sub>\\<^sub>v,\[x::=v]\<^sub>\\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v,\a[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_letI) + + show "atom za \ (xa,\,\,\,\[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v,ea[x::=v]\<^sub>e\<^sub>v, + \a[x::=v]\<^sub>\\<^sub>v,sa[x::=v]\<^sub>s\<^sub>v)" + by(subst_tuple_mth add: check_letI) + + show "\; \; \; \[x::=v]\<^sub>\\<^sub>v; \[x::=v]\<^sub>\\<^sub>v \ + ea[x::=v]\<^sub>e\<^sub>v \ \ za : ba | ca[x::=v]\<^sub>c\<^sub>v \" + proof - + have "\; \; \; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1; \[x::=v]\<^sub>\\<^sub>v \ + ea[x::=v]\<^sub>e\<^sub>v \ \ za : ba | ca \[x::=v]\<^sub>\\<^sub>v" + using check_letI subst_infer_e by metis + thus ?thesis using check_letI subst_tv.simps + by (metis fresh_prod2I infer_e_wf subst_g_inside_simple) + qed + + show "\; \; \; (xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v; + \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + proof - + have "\; \; \; ((xa, ba, ca[za::=V_var xa]\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v ; + \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + apply(rule check_letI(23)[of "(xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>2"]) + by(metis check_letI append_g.simps subst_defs)+ + + moreover have "(xa, ba, ca[x::=v]\<^sub>c\<^sub>v[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v = + ((xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v" + using subst_cv_commute subst_gv.simps check_letI + by (metis ms_fresh_all(39) ms_fresh_all(49) subst_cv_commute_full) + ultimately show ?thesis + using subst_defs by auto + qed + qed +next + case (check_assertI xa \ \ \ \ \ ca \ s) + show ?case unfolding subst_sv.simps proof + show \atom xa \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, ca[x::=v]\<^sub>c\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, s[x::=v]\<^sub>s\<^sub>v)\ + by(subst_tuple_mth add: check_assertI) + have "xa \ x" using check_assertI by fastforce + thus \ \ ; \ ; \ ; (xa, B_bool, ca[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v\ + using check_assertI(12)[of "(xa, B_bool, c) #\<^sub>\ \\<^sub>2" x v] check_assertI subst_gv.simps append_g.simps by metis + have \\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ ca[x::=v]\<^sub>c\<^sub>v \ proof(rule subst_valid ) + show \\ ; \ ; \\<^sub>1 \ c[z::=v]\<^sub>c\<^sub>v \ using infer_v_c_valid check_assertI by metis + show \ \ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f v : b \ using check_assertI infer_v_wf b_of.simps subtype_eq_base + by (metis subtype_eq_base2) + show \ \ ; \ \\<^sub>w\<^sub>f \\<^sub>1 \ using check_assertI infer_v_wf by metis + have " \ ; \ \\<^sub>w\<^sub>f \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1" using check_assertI wfX_wfY by metis + thus \atom x \ \\<^sub>1\ using check_assertI wfG_suffix wfG_elims by metis + + moreover have "\ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f \ z : b | c \" using subtype_wfT check_assertI by metis + moreover have "x \ z" using fresh_Pair check_assertI fresh_x_neq by metis + ultimately show \atom x \ c\ using check_assertI wfT_fresh_c by metis + + show \ \ ; \ \\<^sub>w\<^sub>f \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ using check_assertI wfX_wfY by metis + show \\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ ca \ using check_assertI by auto + qed + thus \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ ca[x::=v]\<^sub>c\<^sub>v \ using check_assertI + proof - + show ?thesis + by (metis (no_types) \\ = \\<^sub>2 @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\ \\ ; \ ; \ \ ca\ \\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ ca[x::=v]\<^sub>c\<^sub>v\ subst_g_inside valid_g_wf) (* 93 ms *) + qed + + have "\ ; \ ; \\<^sub>1 \\<^sub>w\<^sub>f v : b" using infer_v_wf b_of.simps check_assertI + by (metis subtype_eq_base2) + thus \ \ ; \ ; \[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v \ using wf_subst2(6) check_assertI by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist vv cs \ css) + show ?case unfolding * using subst_sv.simps check_branch_list_consI by simp +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + show ?case unfolding * using subst_sv.simps check_branch_list_finalI by simp +next + case (check_branch_s_branchI \ \ \ \ \ const xa \ tid cons va sa) + hence *:"(AS_branch cons xa sa)[x::=v]\<^sub>s\<^sub>v = (AS_branch cons xa sa[x::=v]\<^sub>s\<^sub>v)" using subst_branchv.simps fresh_Pair by metis + show ?case unfolding * proof + + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v\\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v " + using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis + + show "\\<^sub>w\<^sub>f \ " using check_branch_s_branchI by metis + + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v]\<^sub>\\<^sub>v " + using wf_subst check_branch_s_branchI subtype_eq_base2 b_of.simps infer_v_wf by metis + + show wft:"\ ; {||} ; GNil\\<^sub>w\<^sub>f const " using check_branch_s_branchI by metis + + show "atom xa \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, tid, cons, const,va[x::=v]\<^sub>v\<^sub>v, \[x::=v]\<^sub>\\<^sub>v)" + apply(unfold fresh_prodN, (simp add: check_branch_s_branchI )+) + apply(rule,metis fresh_z_subst_g check_branch_s_branchI fresh_Pair ) + by(metis fresh_subst_dv check_branch_s_branchI fresh_Pair ) + + have "\ ; \ ; \ ; ((xa, b_of const, CE_val va == CE_val(V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + using check_branch_s_branchI by (metis append_g.simps(2)) + + moreover have "(xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of (const) xa) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v = + ((xa, b_of const , CE_val va == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v" + proof - + have *:"xa \ x" using check_branch_s_branchI fresh_at_base by metis + have "atom x \ const" using wfT_nil_supp[OF wft] fresh_def by auto + hence "atom x \ (const,xa)" using fresh_at_base wfT_nil_supp[OF wft] fresh_Pair fresh_def * by auto + moreover hence "(c_of (const) xa)[x::=v]\<^sub>c\<^sub>v = c_of (const) xa" + using c_of_fresh[of x const xa] forget_subst_cv wfT_nil_supp wft by metis + moreover hence "(V_cons tid cons (V_var xa))[x::=v]\<^sub>v\<^sub>v = (V_cons tid cons (V_var xa))" using check_branch_s_branchI subst_vv.simps * by metis + ultimately show ?thesis using subst_gv.simps check_branch_s_branchI subst_cv.simps subst_cev.simps * by presburger + qed + + ultimately show "\ ; \ ; \ ; (xa, b_of const, CE_val va[x::=v]\<^sub>v\<^sub>v == CE_val (V_cons tid cons (V_var xa)) AND c_of const xa) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ sa[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + by metis + qed + +next + case (check_let2I xa \ \ \ G \ t s1 \a s2 ) + hence *:"(AS_let2 xa t s1 s2)[x::=v]\<^sub>s\<^sub>v = (AS_let2 xa t[x::=v]\<^sub>\\<^sub>v (s1[x::=v]\<^sub>s\<^sub>v) s2[x::=v]\<^sub>s\<^sub>v)" using subst_sv.simps fresh_Pair by metis + have "xa \ x" using check_let2I fresh_at_base by metis + show ?case unfolding * proof + show "atom xa \ (\, \, \, G[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, t[x::=v]\<^sub>\\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, \a[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_let2I) + show "\ ; \ ; \ ; G[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" using check_let2I by metis + + have "\ ; \ ; \ ; ((xa, b_of t, c_of t xa) #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + proof(rule check_let2I(14)) + show \(xa, b_of t, c_of t xa) #\<^sub>\ G = (((xa, b_of t, c_of t xa)#\<^sub>\ \\<^sub>2)) @ (x, b, c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\ + using check_let2I append_g.simps by metis + show \ \ ; \ ; \\<^sub>1 \ v \ \\ using check_let2I by metis + show \\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \\ using check_let2I by metis + show \atom z \ (x, v)\ using check_let2I by metis + qed + moreover have "c_of t[x::=v]\<^sub>\\<^sub>v xa = (c_of t xa)[x::=v]\<^sub>c\<^sub>v" using subst_v_c_of fresh_Pair check_let2I by metis + moreover have "b_of t[x::=v]\<^sub>\\<^sub>v = b_of t" using b_of.simps subst_tv.simps b_of_subst by metis + ultimately show " \ ; \ ; \ ; (xa, b_of t[x::=v]\<^sub>\\<^sub>v, c_of t[x::=v]\<^sub>\\<^sub>v xa) #\<^sub>\ G[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \a[x::=v]\<^sub>\\<^sub>v" + using check_let2I(14) subst_gv.simps \xa \ x\ b_of.simps by metis + qed + +next + + case (check_varI u \ \ \ \ \ \' va \'' s) + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside check_s_wf check_varI by meson + + have "\ ; \ ;\ ; subst_gv \ x v ; \[x::=v]\<^sub>\\<^sub>v \ AS_var u \'[x::=v]\<^sub>\\<^sub>v (va[x::=v]\<^sub>v\<^sub>v) (subst_sv s x v) \ \''[x::=v]\<^sub>\\<^sub>v" + proof(rule Typing.check_varI) + show "atom u \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, \''[x::=v]\<^sub>\\<^sub>v)" + by(subst_tuple_mth add: check_varI) + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ va[x::=v]\<^sub>v\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" using check_varI subst_infer_check_v ** by metis + show "\ ; \ ; \ ; subst_gv \ x v ; (u, \'[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \''[x::=v]\<^sub>\\<^sub>v" proof - + have "wfD \ \ (\\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1) ((u,\')#\<^sub>\ \)" using check_varI check_s_wf by meson + moreover have "wfV \ \ \\<^sub>1 v (b_of \)" using infer_v_wf check_varI(6) check_varI by metis + have "wfD \ \ (\[x::=v]\<^sub>\\<^sub>v) ((u, \'[x::=v]\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v]\<^sub>\\<^sub>v)" proof(subst subst_dv.simps(2)[symmetric], subst **, rule wfD_subst) + show "\ ; \ ; \\<^sub>1 \ v \ \" using check_varI by auto + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1\\<^sub>w\<^sub>f (u, \') #\<^sub>\ \" using check_varI check_s_wf by simp + show "b_of \ = b" using check_varI subtype_eq_base2 b_of.simps by auto + qed + thus ?thesis using check_varI by auto + qed + qed + moreover have "atom u \ (x,v)" using u_fresh_xv by auto + ultimately show ?case using subst_sv.simps(7) by auto + +next + case (check_assignI P \ \ \ \ u \1 v' z1 \') (* may need to revisit subst in \ as well *) + + have "wfG P \ \" using check_v_wf check_assignI by simp + hence gs: "\\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 = \[x::=v]\<^sub>\\<^sub>v" using subst_g_inside check_assignI by simp + + have "P ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_assign u (v'[x::=v]\<^sub>v\<^sub>v) \ \'[x::=v]\<^sub>\\<^sub>v" + proof(rule Typing.check_assignI) + show "P \\<^sub>w\<^sub>f \ " using check_assignI by auto + show "wfD P \ (\[x::=v]\<^sub>\\<^sub>v) \[x::=v]\<^sub>\\<^sub>v" using wf_subst(15)[OF check_assignI(2)] gs infer_v_v_wf check_assignI b_of.simps subtype_eq_base2 by metis + thus "(u, \1[x::=v]\<^sub>\\<^sub>v) \ setD \[x::=v]\<^sub>\\<^sub>v" using check_assignI subst_dv_member by metis + thus "P ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \1[x::=v]\<^sub>\\<^sub>v" using subst_infer_check_v check_assignI gs by metis + + have "P ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ \ z : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" proof(rule subst_subtype_tau) + show "P ; \ ; \\<^sub>1 \ v \ \" using check_assignI by auto + show "P ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" using check_assignI by meson + show "P ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ \ z : B_unit | TRUE \ \ \'" using check_assignI + by (metis Abs1_eq_iff(3) \.eq_iff c.fresh(1) c.perm_simps(1)) + show "atom z \ (x, v)" using check_assignI by auto + qed + moreover have "\ z : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z : B_unit | TRUE \" using subst_tv.simps subst_cv.simps check_assignI by presburger + ultimately show "P ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \ z : B_unit | TRUE \ \ \'[x::=v]\<^sub>\\<^sub>v" using gs by auto + qed + thus ?case using subst_sv.simps(5) by auto + +next + case (check_whileI \ \ \ \ \ s1 z' s2 \') + have " wfG \ \ (\\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1)" using check_whileI check_s_wf by meson + hence **: " \[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_whileI by auto + have teq: "(\ z : B_unit | TRUE \)[x::=v]\<^sub>\\<^sub>v = (\ z : B_unit | TRUE \)" by(auto simp add: subst_sv.simps check_whileI) + moreover have "(\ z : B_unit | TRUE \) = (\ z' : B_unit | TRUE \)" using type_eq_flip c.fresh flip_fresh_fresh by metis + ultimately have teq2:"(\ z' : B_unit | TRUE \)[x::=v]\<^sub>\\<^sub>v = (\ z' : B_unit | TRUE \)" by metis + + hence "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ \ z' : B_bool | TRUE \" using check_whileI subst_sv.simps subst_top_eq by metis + moreover have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \ z' : B_unit | TRUE \" using check_whileI subst_top_eq by metis + moreover have "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ \ z' : B_unit | TRUE \ \ \'[x::=v]\<^sub>\\<^sub>v" proof - + have "\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ \ z' : B_unit | TRUE \[x::=v]\<^sub>\\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" proof(rule subst_subtype_tau) + show "\ ; \ ; \\<^sub>1 \ v \ \" by(auto simp add: check_whileI) + show "\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" by(auto simp add: check_whileI) + show "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ \ z' : B_unit | TRUE \ \ \'" using check_whileI by metis + show "atom z \ (x, v)" by(auto simp add: check_whileI) + qed + thus ?thesis using teq2 ** by auto + qed + + ultimately have " \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_while s1[x::=v]\<^sub>s\<^sub>v s2[x::=v]\<^sub>s\<^sub>v \ \'[x::=v]\<^sub>\\<^sub>v" + using Typing.check_whileI by metis + then show ?case using subst_sv.simps by metis +next + case (check_seqI P \ \ \ \ s1 z s2 \ ) + hence "P ; \; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_seq (s1[x::=v]\<^sub>s\<^sub>v) (s2[x::=v]\<^sub>s\<^sub>v) \ \[x::=v]\<^sub>\\<^sub>v" using Typing.check_seqI subst_top_eq check_seqI by metis + then show ?case using subst_sv.simps by metis +next + case (check_caseI \ \ \ \ \ tid dclist v' cs \ za) + + have wf: "wfG \ \ \" using check_caseI check_v_wf by simp + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_caseI by auto + + have "\ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ AS_match (v'[x::=v]\<^sub>v\<^sub>v) (subst_branchlv cs x v) \ \[x::=v]\<^sub>\\<^sub>v" proof(rule Typing.check_caseI ) + show "check_branch_list \ \ \ (\[x::=v]\<^sub>\\<^sub>v) \[x::=v]\<^sub>\\<^sub>v tid dclist v'[x::=v]\<^sub>v\<^sub>v (subst_branchlv cs x v ) (\[x::=v]\<^sub>\\<^sub>v)" using check_caseI by auto + show "AF_typedef tid dclist \ set \" using check_caseI by auto + show "\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ v'[x::=v]\<^sub>v\<^sub>v \ \ za : B_id tid | TRUE \" proof - + have "\ ; \ ; \\<^sub>2 @ (x, b, c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1 \ v' \ \ za : B_id tid | TRUE \" + using check_caseI by argo + hence "\ ; \ ; \\<^sub>2[x::=v]\<^sub>\\<^sub>v @ \\<^sub>1 \ v'[x::=v]\<^sub>v\<^sub>v \ (\ za : B_id tid | TRUE \)[x::=v]\<^sub>\\<^sub>v" + using check_caseI subst_infer_check_v[OF check_caseI(7) _ check_caseI(8) check_caseI(9)] by meson + moreover have "(\ za : B_id tid | TRUE \) = ((\ za : B_id tid | TRUE \)[x::=v]\<^sub>\\<^sub>v)" + using subst_cv.simps subst_tv.simps subst_cv_true by fast + ultimately show ?thesis using check_caseI ** by argo + qed + show "wfTh \" using check_caseI by auto + qed + thus ?case using subst_branchlv.simps subst_sv.simps(4) by metis +next + case (check_ifI z' \ \ \ \ \ va s1 s2 \') + show ?case unfolding subst_sv.simps proof + show \atom z' \ (\, \, \, \[x::=v]\<^sub>\\<^sub>v, \[x::=v]\<^sub>\\<^sub>v, va[x::=v]\<^sub>v\<^sub>v, s1[x::=v]\<^sub>s\<^sub>v, s2[x::=v]\<^sub>s\<^sub>v, \'[x::=v]\<^sub>\\<^sub>v)\ + by(subst_tuple_mth add: check_ifI) + have *:"\ z' : B_bool | TRUE \[x::=v]\<^sub>\\<^sub>v = \ z' : B_bool | TRUE \" using subst_tv.simps check_ifI + by (metis freshers(19) subst_cv.simps(1) type_eq_subst) + have **: "\[x::=v]\<^sub>\\<^sub>v = \\<^sub>2[x::=v]\<^sub>\\<^sub>v@\\<^sub>1" using subst_g_inside wf check_ifI check_v_wf by metis + show \\ ; \ ; \[x::=v]\<^sub>\\<^sub>v \ va[x::=v]\<^sub>v\<^sub>v \ \ z' : B_bool | TRUE \\ + proof(subst *[symmetric], rule subst_infer_check_v1[where \\<^sub>1=\\<^sub>2 and \\<^sub>2=\\<^sub>1]) + show "\ = \\<^sub>2 @ ((x, b ,c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\<^sub>1)" using check_ifI by metis + show \ \ ; \ ; \\<^sub>1 \ v \ \\ using check_ifI by metis + show \\ ; \ ; \ \ va \ \ z' : B_bool | TRUE \\ using check_ifI by metis + show \\ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \\ using check_ifI by metis + show \atom z \ (x, v)\ using check_ifI by metis + qed + + have " \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \ = \ z' : b_of \' | [ va ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' z' \[x::=v]\<^sub>\\<^sub>v" + by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of) + + thus \ \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s1[x::=v]\<^sub>s\<^sub>v \ \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \\ + using check_ifI by metis + have " \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \ = \ z' : b_of \' | [ va ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' z' \[x::=v]\<^sub>\\<^sub>v" + by(simp add: subst_tv.simps fresh_Pair check_ifI b_of_subst subst_v_c_of) + thus \ \ ; \ ; \ ; \[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s2[x::=v]\<^sub>s\<^sub>v \ \ z' : b_of \'[x::=v]\<^sub>\\<^sub>v | [ va[x::=v]\<^sub>v\<^sub>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \'[x::=v]\<^sub>\\<^sub>v z' \\ + using check_ifI by metis + qed +qed + +lemma subst_check_check_s: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \\<^sub>1::\ and \\<^sub>2::\ + assumes "\ ; \ ; \\<^sub>1 \ v \ \ z : b | c \" and "atom z \ (x, v)" + and "check_s \ \ \ \ \ s \'" and "\ = (\\<^sub>2@((x,b,c[z::=[x]\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\\<^sub>1))" + shows "check_s \ \ \ (subst_gv \ x v) \[x::=v]\<^sub>\\<^sub>v (s[x::=v]\<^sub>s\<^sub>v) (subst_tv \' x v )" +proof - + obtain \ where "\ ; \ ; \\<^sub>1 \ v \ \ \ \ ; \ ; \\<^sub>1 \ \ \ \ z : b | c \" using check_v_elims assms by auto + thus ?thesis using subst_infer_check_s assms by metis +qed + +text \ If a statement checks against a type @{text "\"} then it checks against a supertype of @{text "\"} \ +lemma check_s_supertype: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \::\ and \'::\ and css::branch_list + shows "check_s \ \ \ G \ s t1 \ \ ; \ ; G \ t1 \ t2 \ check_s \ \ \ G \ s t2" and + "check_branch_s \ \ \ G \ tid cons const v cs t1 \ \ ; \ ; G \ t1 \ t2 \ check_branch_s \ \ \ G \ tid cons const v cs t2" and + "check_branch_list \ \ \ G \ tid dclist v css t1 \ \ ; \ ; G \ t1 \ t2 \ check_branch_list \ \ \ G \ tid dclist v css t2" +proof(induct arbitrary: t2 and t2 and t2 rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \ ) + hence " \ ; \ ; \ \ \' \ t2" using subtype_trans by meson + then show ?case using subtype_trans Typing.check_valI check_valI by metis + +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof(rule Typing.check_letI) + show "atom x \(\, \, \, \, \, e, t2)" using check_letI subtype_fresh_tau fresh_prodN by metis + show "atom z \ (x, \, \, \, \, \, e, t2, s)" using check_letI(2) subtype_fresh_tau[of z \ \ _ _ t2] fresh_prodN check_letI(6) by auto + show "\ ; \ ; \ ; \ ; \ \ e \ \ z : b | c \" using check_letI by meson + + have "wfG \ \ ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \)" using check_letI check_s_wf subst_defs by metis + moreover have "toSet \ \ toSet ((x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \)" by auto + ultimately have " \ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ \ \ t2" using subtype_weakening[OF check_letI(6)] by auto + thus "\ ; \ ; \ ; (x, b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ ; \ \ s \ t2" using check_letI subst_defs by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using Typing.check_branch_list_consI by auto +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using Typing.check_branch_list_finalI by auto +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + have "atom x \ t2" using subtype_fresh_tau[of x \ ] check_branch_s_branchI(5,8) fresh_prodN by metis + thus "atom x \ (\, \, \, \, \, tid, cons, const, v, t2)" using check_branch_s_branchI fresh_prodN by metis + show "wfT \ \ \ t2" using subtype_wf check_branch_s_branchI by meson + show "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \ ; \ \ s \ t2" proof - + have "wfG \ \ ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" using check_s_wf check_branch_s_branchI by metis + moreover have "toSet \ \ toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \)" by auto + hence "\ ; \ ; ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) \ \ \ t2" + using check_branch_s_branchI subtype_weakening + using calculation by presburger + thus ?thesis using check_branch_s_branchI by presburger + qed + qed(auto simp add: check_branch_s_branchI) +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof(rule Typing.check_ifI) + have *:"atom z \ t2" using subtype_fresh_tau[of z \ \ ] check_ifI fresh_prodN by auto + thus \atom z \ (\, \, \, \, \, v, s1, s2, t2)\ using check_ifI fresh_prodN by auto + show \\ ; \ ; \ \ v \ \ z : B_bool | TRUE \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \ \ s1 \ \ z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \\ + using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis + + show \ \ ; \ ; \ ; \ ; \ \ s2 \ \ z : b_of t2 | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t2 z \\ + using check_ifI subtype_if1 fresh_prodN base_for_lit.simps b_of.simps * check_v_wf by metis + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + + show ?case proof + have "atom x \ t2" using subtype_fresh_tau[OF _ _ \\ ; \ ; \ \ \ \ t2\] check_assertI fresh_prodN by simp + thus "atom x \ (\, \, \, \, \, c, t2, s)" using subtype_fresh_tau check_assertI fresh_prodN by simp + have "\ ; \ ; (x, B_bool, c) #\<^sub>\ \ \ \ \ t2" apply(rule subtype_weakening) + using check_assertI apply simp + using toSet.simps apply blast + using check_assertI check_s_wf by simp + thus "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ t2" using check_assertI by simp + show "\ ; \ ; \ \ c " using check_assertI by auto + show "\ ; \ ; \ \\<^sub>w\<^sub>f \ " using check_assertI by auto + qed +next + case (check_let2I x P \ \ G \ t s1 \ s2 ) + have "wfG P \ ((x, b_of t, c_of t x) #\<^sub>\ G)" + using check_let2I check_s_wf by metis + moreover have "toSet G \ toSet ((x, b_of t, c_of t x) #\<^sub>\ G)" by auto + ultimately have *:"P ; \ ; (x, b_of t, c_of t x) #\<^sub>\ G \ \ \ t2" using check_let2I subtype_weakening by metis + show ?case proof(rule Typing.check_let2I) + have "atom x \ t2" using subtype_fresh_tau[of x \ ] check_let2I fresh_prodN by metis + thus "atom x \ (P, \, \, G, \, t, s1, t2)" using check_let2I fresh_prodN by metis + show "P ; \ ; \ ; G ; \ \ s1 \ t" using check_let2I by blast + show "P ; \ ; \ ;(x, b_of t, c_of t x ) #\<^sub>\ G ; \ \ s2 \ t2" using check_let2I * by blast + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof(rule Typing.check_varI) + have "atom u \ t2" using u_fresh_t by auto + thus \atom u \ (\, \, \, \, \, \', v, t2)\ using check_varI fresh_prodN by auto + show \\ ; \ ; \ \ v \ \'\ using check_varI by auto + show \ \ ; \ ; \ ; \ ; (u, \') #\<^sub>\ \ \ s \ t2\ using check_varI by auto + qed +next + case (check_assignI \ u \ P G v z \') + then show ?case using Typing.check_assignI by (meson subtype_trans) +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using Typing.check_whileI by (meson subtype_trans) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case using Typing.check_seqI by blast +next + case (check_caseI \ \ \ tid cs \ v z) + then show ?case using Typing.check_caseI subtype_trans by meson + +qed + +lemma subtype_let: + fixes s'::s and cs::branch_s and css::branch_list and v::v + shows "\ ; \ ; \ ; GNil ; \ \ AS_let x e\<^sub>1 s \ \ \ \ ; \ ; \ ; GNil ; \ \ e\<^sub>1 \ \\<^sub>1 \ + \ ; \ ; \ ; GNil ; \ \ e\<^sub>2 \ \\<^sub>2 \ \ ; \ ; GNil \ \\<^sub>2 \ \\<^sub>1 \ \ ; \ ; \ ; GNil ; \ \ AS_let x e\<^sub>2 s \ \" and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" +proof(nominal_induct GNil \ "AS_let x e\<^sub>1 s" \ and \ and \ avoiding: e\<^sub>2 \\<^sub>1 \\<^sub>2 + rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_letI x1 \ \ \ \ \1 z1 s1 b1 c1) + obtain z2 and b2 and c2 where t2:"\\<^sub>2 = \ z2 : b2 | c2 \ \ atom z2 \ (x1, \, \, \, GNil, \, e\<^sub>2, \1, s1) " + using obtain_fresh_z by metis + + obtain z1a and b1a and c1a where t1:"\\<^sub>1 = \ z1a : b1a | c1a \ \ atom z1a \ x1" using infer_e_uniqueness check_letI by metis + hence t3: " \ z1a : b1a | c1a \ = \ z1 : b1 | c1 \ " using infer_e_uniqueness check_letI by metis + + have beq: "b1a = b2 \ b2 = b1" using check_letI subtype_eq_base t1 t2 t3 by metis + + have " \ ; \ ; \ ; GNil ; \ \ AS_let x1 e\<^sub>2 s1 \ \1" proof + show \atom x1 \ (\, \, \, GNil, \, e\<^sub>2, \1)\ using check_letI t2 fresh_prodN by metis + show \atom z2 \ (x1, \, \, \, GNil, \, e\<^sub>2, \1, s1)\ using check_letI t2 using check_letI t2 fresh_prodN by metis + show \\ ; \ ; \ ; GNil ; \ \ e\<^sub>2 \ \ z2 : b2 | c2 \\ using check_letI t2 by metis + + have \ \ ; \ ; \ ; GNil@(x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ + proof(rule ctx_subtype_s) + have "c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v = c1[z1::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v" using subst_v_flip_eq_two subst_v_c_def t3 \.eq_iff by metis + thus \ \ ; \ ; \ ; GNil @ (x1, b2, c1a[z1a::=[ x1 ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ using check_letI beq append_g.simps subst_defs by metis + show \\ ; \ ; GNil \ \ z2 : b2 | c2 \ \ \ z1a : b2 | c1a \\ using check_letI beq t1 t2 by metis + have "atom x1 \ c2" using t2 check_letI \_fresh_c fresh_prodN by blast + moreover have "atom x1 \ c1a" using t1 check_letI \_fresh_c fresh_prodN by blast + ultimately show \atom x1 \ (z1a, z2, c1a, c2)\ using t1 t2 fresh_prodN fresh_x_neq by metis + qed + + thus \ \ ; \ ; \ ; (x1, b2, c2[z2::=[ x1 ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil ; \ \ s1 \ \1\ using append_g.simps subst_defs by metis + qed + + moreover have "AS_let x1 e\<^sub>2 s1 = AS_let x e\<^sub>2 s" using check_letI s_branch_s_branch_list.eq_iff by metis + + ultimately show ?case by metis + +qed(auto+) + +end \ No newline at end of file diff --git a/thys/MiniSail/MiniSail.thy b/thys/MiniSail/MiniSail.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/MiniSail.thy @@ -0,0 +1,10 @@ +(*<*) +theory MiniSail + imports + "Safety" +begin +(*>*) + +(*<*) +end +(*>*) \ No newline at end of file diff --git a/thys/MiniSail/Nominal-Utils.thy b/thys/MiniSail/Nominal-Utils.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Nominal-Utils.thy @@ -0,0 +1,586 @@ +(*<*) +theory "Nominal-Utils" +imports Nominal2.Nominal2 "HOL-Library.AList" +begin +(*>*) + +chapter \Prelude\ +text \Some useful Nominal lemmas. Many of these are from Launchbury.Nominal-Utils.\ + +section \Lemmas helping with equivariance proofs\ + +lemma perm_rel_lemma: + assumes "\ \ x y. r (\ \ x) (\ \ y) \ r x y" + shows "r (\ \ x) (\ \ y) \ r x y" (is "?l \ ?r") +by (metis (full_types) assms permute_minus_cancel(2)) + +lemma perm_rel_lemma2: + assumes "\ \ x y. r x y \ r (\ \ x) (\ \ y)" + shows "r x y \ r (\ \ x) (\ \ y)" (is "?l \ ?r") +by (metis (full_types) assms permute_minus_cancel(2)) + +lemma fun_eqvtI: + assumes f_eqvt[eqvt]: "(\ p x. p \ (f x) = f (p \ x))" + shows "p \ f = f" by perm_simp rule + +lemma eqvt_at_apply: + assumes "eqvt_at f x" + shows "(p \ f) x = f x" +by (metis (hide_lams, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) + +lemma eqvt_at_apply': + assumes "eqvt_at f x" + shows "p \ f x = f (p \ x)" +by (metis (hide_lams, no_types) assms eqvt_at_def) + +lemma eqvt_at_apply'': + assumes "eqvt_at f x" + shows "(p \ f) (p \ x) = f (p \ x)" +by (metis (hide_lams, no_types) assms eqvt_at_def permute_fun_def permute_minus_cancel(1)) + +lemma size_list_eqvt[eqvt]: "p \ size_list f x = size_list (p \ f) (p \ x)" +proof (induction x) + case (Cons x xs) + have "f x = p \ (f x)" by (simp add: permute_pure) + also have "... = (p \ f) (p \ x)" by simp + with Cons + show ?case by (auto simp add: permute_pure) +qed simp + +section \ Freshness via equivariance \ + +lemma eqvt_fresh_cong1: "(\p x. p \ (f x) = f (p \ x)) \ a \ x \ a \ f x " + apply (rule fresh_fun_eqvt_app[of f]) + apply (rule eqvtI) + apply (rule eq_reflection) + apply (rule ext) + apply (metis permute_fun_def permute_minus_cancel(1)) + apply assumption + done + +lemma eqvt_fresh_cong2: + assumes eqvt: "(\p x y. p \ (f x y) = f (p \ x) (p \ y))" + and fresh1: "a \ x" and fresh2: "a \ y" + shows "a \ f x y" +proof- + have "eqvt (\ (x,y). f x y)" + using eqvt + apply (- , auto simp add: eqvt_def) + by (rule ext, auto, metis permute_minus_cancel(1)) + moreover + have "a \ (x, y)" using fresh1 fresh2 by auto + ultimately + have "a \ (\ (x,y). f x y) (x, y)" by (rule fresh_fun_eqvt_app) + thus ?thesis by simp +qed + +lemma eqvt_fresh_star_cong1: + assumes eqvt: "(\p x. p \ (f x) = f (p \ x))" + and fresh1: "a \* x" + shows "a \* f x" + by (metis fresh_star_def eqvt_fresh_cong1 assms) + +lemma eqvt_fresh_star_cong2: + assumes eqvt: "(\p x y. p \ (f x y) = f (p \ x) (p \ y))" + and fresh1: "a \* x" and fresh2: "a \* y" + shows "a \* f x y" + by (metis fresh_star_def eqvt_fresh_cong2 assms) + +lemma eqvt_fresh_cong3: + assumes eqvt: "(\p x y z. p \ (f x y z) = f (p \ x) (p \ y) (p \ z))" + and fresh1: "a \ x" and fresh2: "a \ y" and fresh3: "a \ z" + shows "a \ f x y z" +proof- + have "eqvt (\ (x,y,z). f x y z)" + using eqvt + apply (- , auto simp add: eqvt_def) + by(rule ext, auto, metis permute_minus_cancel(1)) + moreover + have "a \ (x, y, z)" using fresh1 fresh2 fresh3 by auto + ultimately + have "a \ (\ (x,y,z). f x y z) (x, y, z)" by (rule fresh_fun_eqvt_app) + thus ?thesis by simp +qed + +lemma eqvt_fresh_star_cong3: + assumes eqvt: "(\p x y z. p \ (f x y z) = f (p \ x) (p \ y) (p \ z))" + and fresh1: "a \* x" and fresh2: "a \* y" and fresh3: "a \* z" + shows "a \* f x y z" + by (metis fresh_star_def eqvt_fresh_cong3 assms) + +section \ Additional simplification rules \ + +lemma not_self_fresh[simp]: "atom x \ x \ False" + by (metis fresh_at_base(2)) + +lemma fresh_star_singleton: "{ x } \* e \ x \ e" + by (simp add: fresh_star_def) + +section \ Additional equivariance lemmas \ + +lemma eqvt_cases: + fixes f x \ + assumes eqvt: "\x. \ \ f x = f (\ \ x)" + obtains "f x" "f (\ \ x)" | "\ f x " " \ f (\ \ x)" + using assms[symmetric] + by (cases "f x") auto + +lemma range_eqvt: "\ \ range Y = range (\ \ Y)" + unfolding image_eqvt UNIV_eqvt .. + +lemma case_option_eqvt[eqvt]: + "\ \ case_option d f x = case_option (\ \ d) (\ \ f) (\ \ x)" + by(cases x)(simp_all) + +lemma supp_option_eqvt: + "supp (case_option d f x) \ supp d \ supp f \ supp x" + apply (cases x) + apply (auto simp add: supp_Some ) + apply (metis (mono_tags) Un_iff subsetCE supp_fun_app) + done + +lemma funpow_eqvt[simp,eqvt]: + "\ \ ((f :: 'a \ 'a::pt) ^^ n) = (\ \ f) ^^ (\ \ n)" + by (induct n,simp, rule ext, simp, perm_simp,simp) + +lemma delete_eqvt[eqvt]: + "\ \ AList.delete x \ = AList.delete (\ \ x) (\ \ \)" +by (induct \, auto) + +lemma restrict_eqvt[eqvt]: + "\ \ AList.restrict S \ = AList.restrict (\ \ S) (\ \ \)" +unfolding AList.restrict_eq by perm_simp rule + +lemma supp_restrict: + "supp (AList.restrict S \) \ supp \" + by (induction \) (auto simp add: supp_Pair supp_Cons) + +lemma clearjunk_eqvt[eqvt]: + "\ \ AList.clearjunk \ = AList.clearjunk (\ \ \)" + by (induction \ rule: clearjunk.induct) auto + +lemma map_ran_eqvt[eqvt]: + "\ \ map_ran f \ = map_ran (\ \ f) (\ \ \)" +by (induct \, auto) + +lemma dom_perm: + "dom (\ \ f) = \ \ (dom f)" + unfolding dom_def by (perm_simp) (simp) + +lemmas dom_perm_rev[simp,eqvt] = dom_perm[symmetric] + +lemma ran_perm[simp]: + "\ \ (ran f) = ran (\ \ f)" + unfolding ran_def by (perm_simp) (simp) + +lemma map_add_eqvt[eqvt]: + "\ \ (m1 ++ m2) = (\ \ m1) ++ (\ \ m2)" + unfolding map_add_def + by (perm_simp, rule) + +lemma map_of_eqvt[eqvt]: + "\ \ map_of l = map_of (\ \ l)" + by (induct l, simp add: permute_fun_def,simp,perm_simp,auto) + +lemma concat_eqvt[eqvt]: "\ \ concat l = concat (\ \ l)" + by (induction l)(auto simp add: append_eqvt) + +lemma tranclp_eqvt[eqvt]: "\ \ tranclp P v\<^sub>1 v\<^sub>2 = tranclp (\ \ P) (\ \ v\<^sub>1) (\ \ v\<^sub>2)" + unfolding tranclp_def by perm_simp rule + +lemma rtranclp_eqvt[eqvt]: "\ \ rtranclp P v\<^sub>1 v\<^sub>2 = rtranclp (\ \ P) (\ \ v\<^sub>1) (\ \ v\<^sub>2)" + unfolding rtranclp_def by perm_simp rule + +lemma Set_filter_eqvt[eqvt]: "\ \ Set.filter P S = Set.filter (\ \ P) (\ \ S)" + unfolding Set.filter_def + by perm_simp rule + +lemma Sigma_eqvt'[eqvt]: "\ \ Sigma = Sigma" + apply (rule ext) + apply (rule ext) + apply (subst permute_fun_def) + apply (subst permute_fun_def) + unfolding Sigma_def + apply perm_simp + apply (simp add: permute_self) + done + +lemma override_on_eqvt[eqvt]: + "\ \ (override_on m1 m2 S) = override_on (\ \ m1) (\ \ m2) (\ \ S)" + by (auto simp add: override_on_def ) + +lemma card_eqvt[eqvt]: + "\ \ (card S) = card (\ \ S)" +by (cases "finite S", induct rule: finite_induct) (auto simp add: card_insert_if mem_permute_iff permute_pure) + +(* Helper lemmas provided by Christian Urban *) + +lemma Projl_permute: + assumes a: "\y. f = Inl y" + shows "(p \ (Sum_Type.projl f)) = Sum_Type.projl (p \ f)" +using a by auto + +lemma Projr_permute: + assumes a: "\y. f = Inr y" + shows "(p \ (Sum_Type.projr f)) = Sum_Type.projr (p \ f)" +using a by auto + +section \ Freshness lemmas \ + +lemma fresh_list_elem: + assumes "a \ \" + and "e \ set \" + shows "a \ e" +using assms +by(induct \)(auto simp add: fresh_Cons) + +lemma set_not_fresh: + "x \ set L \ \(atom x \ L)" + by (metis fresh_list_elem not_self_fresh) + +lemma pure_fresh_star[simp]: "a \* (x :: 'a :: pure)" + by (simp add: fresh_star_def pure_fresh) + +lemma supp_set_mem: "x \ set L \ supp x \ supp L" + by (induct L) (auto simp add: supp_Cons) + +lemma set_supp_mono: "set L \ set L2 \ supp L \ supp L2" + by (induct L)(auto simp add: supp_Cons supp_Nil dest:supp_set_mem) + +lemma fresh_star_at_base: + fixes x :: "'a :: at_base" + shows "S \* x \ atom x \ S" + by (metis fresh_at_base(2) fresh_star_def) + +section \ Freshness and support for subsets of variables \ + +lemma supp_mono: "finite (B::'a::fs set) \ A \ B \ supp A \ supp B" + by (metis infinite_super subset_Un_eq supp_of_finite_union) + +lemma fresh_subset: + "finite B \ x \ (B :: 'a::at_base set) \ A \ B \ x \ A" + by (auto dest:supp_mono simp add: fresh_def) + +lemma fresh_star_subset: + "finite B \ x \* (B :: 'a::at_base set) \ A \ B \ x \* A" + by (metis fresh_star_def fresh_subset) + +lemma fresh_star_set_subset: + "x \* (B :: 'a::at_base list) \ set A \ set B \ x \* A" + by (metis fresh_star_set fresh_star_subset[OF finite_set]) + +section \ The set of free variables of an expression \ + +definition fv :: "'a::pt \ 'b::at_base set" + where "fv e = {v. atom v \ supp e}" + +lemma fv_eqvt[simp,eqvt]: "\ \ (fv e) = fv (\ \ e)" + unfolding fv_def by simp + +lemma fv_Nil[simp]: "fv [] = {}" + by (auto simp add: fv_def supp_Nil) +lemma fv_Cons[simp]: "fv (x # xs) = fv x \ fv xs" + by (auto simp add: fv_def supp_Cons) +lemma fv_Pair[simp]: "fv (x, y) = fv x \ fv y" + by (auto simp add: fv_def supp_Pair) +lemma fv_append[simp]: "fv (x @ y) = fv x \ fv y" + by (auto simp add: fv_def supp_append) +lemma fv_at_base[simp]: "fv a = {a::'a::at_base}" + by (auto simp add: fv_def supp_at_base) +lemma fv_pure[simp]: "fv (a::'a::pure) = {}" + by (auto simp add: fv_def pure_supp) + +lemma fv_set_at_base[simp]: "fv (l :: ('a :: at_base) list) = set l" + by (induction l) auto + +lemma flip_not_fv: "a \ fv x \ b \ fv x \ (a \ b) \ x = x" + by (metis flip_def fresh_def fv_def mem_Collect_eq swap_fresh_fresh) + +lemma fv_not_fresh: "atom x \ e \ x \ fv e" + unfolding fv_def fresh_def by blast + +lemma fresh_fv: "finite (fv e :: 'a set) \ atom (x :: ('a::at_base)) \ (fv e :: 'a set) \ atom x \ e" + unfolding fv_def fresh_def + by (auto simp add: supp_finite_set_at_base) + +lemma finite_fv[simp]: "finite (fv (e::'a::fs) :: ('b::at_base) set)" +proof- + have "finite (supp e)" by (metis finite_supp) + hence "finite (atom -` supp e :: 'b set)" + apply (rule finite_vimageI) + apply (rule inj_onI) + apply (simp) + done + moreover + have "(atom -` supp e :: 'b set) = fv e" unfolding fv_def by auto + ultimately + show ?thesis by simp +qed + +definition fv_list :: "'a::fs \ 'b::at_base list" + where "fv_list e = (SOME l. set l = fv e)" + +lemma set_fv_list[simp]: "set (fv_list e) = (fv e :: ('b::at_base) set)" +proof- + have "finite (fv e :: 'b set)" by (rule finite_fv) + from finite_list[OF finite_fv] + obtain l where "set l = (fv e :: 'b set)".. + thus ?thesis + unfolding fv_list_def by (rule someI) +qed + +lemma fresh_fv_list[simp]: + "a \ (fv_list e :: 'b::at_base list) \ a \ (fv e :: 'b::at_base set)" +proof- + have "a \ (fv_list e :: 'b::at_base list) \ a \ set (fv_list e :: 'b::at_base list)" + by (rule fresh_set[symmetric]) + also have "\ \ a \ (fv e :: 'b::at_base set)" by simp + finally show ?thesis. +qed + +section \ Other useful lemmas \ + +lemma pure_permute_id: "permute p = (\ x. (x::'a::pure))" + by rule (simp add: permute_pure) + +lemma supp_set_elem_finite: + assumes "finite S" + and "(m::'a::fs) \ S" + and "y \ supp m" + shows "y \ supp S" + using assms supp_of_finite_sets + by auto + +lemmas fresh_star_Cons = fresh_star_list(2) + +lemma mem_permute_set: + shows "x \ p \ S \ (- p \ x) \ S" + by (metis mem_permute_iff permute_minus_cancel(2)) + +lemma flip_set_both_not_in: + assumes "x \ S" and "x' \ S" + shows "((x' \ x) \ S) = S" + unfolding permute_set_def + by (auto) (metis assms flip_at_base_simps(3))+ + +lemma inj_atom: "inj atom" by (metis atom_eq_iff injI) + +lemmas image_Int[OF inj_atom, simp] + +lemma eqvt_uncurry: "eqvt f \ eqvt (case_prod f)" + unfolding eqvt_def + by perm_simp simp + +lemma supp_fun_app_eqvt2: + assumes a: "eqvt f" + shows "supp (f x y) \ supp x \ supp y" +proof- + from supp_fun_app_eqvt[OF eqvt_uncurry [OF a]] + have "supp (case_prod f (x,y)) \ supp (x,y)". + thus ?thesis by (simp add: supp_Pair) +qed + +lemma supp_fun_app_eqvt3: + assumes a: "eqvt f" + shows "supp (f x y z) \ supp x \ supp y \ supp z" +proof- + from supp_fun_app_eqvt2[OF eqvt_uncurry [OF a]] + have "supp (case_prod f (x,y) z) \ supp (x,y) \ supp z". + thus ?thesis by (simp add: supp_Pair) +qed + +(* Fighting eta-contraction *) +lemma permute_0[simp]: "permute 0 = (\ x. x)" + by auto +lemma permute_comp[simp]: "permute x \ permute y = permute (x + y)" by auto + +lemma map_permute: "map (permute p) = permute p" + apply rule + apply (induct_tac x) + apply auto + done + +lemma fresh_star_restrictA[intro]: "a \* \ \ a \* AList.restrict V \" + by (induction \) (auto simp add: fresh_star_Cons) + +lemma Abs_lst_Nil_eq[simp]: "[[]]lst. (x::'a::fs) = [xs]lst. x' \ (([],x) = (xs, x'))" + apply rule + apply (frule Abs_lst_fcb2[where f = "\ x y _ . (x,y)" and as = "[]" and bs = "xs" and c = "()"]) + apply (auto simp add: fresh_star_def) + done + +lemma Abs_lst_Nil_eq2[simp]: "[xs]lst. (x::'a::fs) = [[]]lst. x' \ ((xs,x) = ([], x'))" + by (subst eq_commute) auto + +lemma prod_cases8 [cases type]: + obtains (fields) a b c d e f g h where "y = (a, b, c, d, e, f, g,h)" + by (cases y, case_tac g) blast + +lemma prod_induct8 [case_names fields, induct type]: + "(\a b c d e f g h. P (a, b, c, d, e, f, g, h)) \ P x" + by (cases x) blast + +lemma prod_cases9 [cases type]: + obtains (fields) a b c d e f g h i where "y = (a, b, c, d, e, f, g,h,i)" + by (cases y, case_tac h) blast + +lemma prod_induct9 [case_names fields, induct type]: + "(\a b c d e f g h i. P (a, b, c, d, e, f, g, h, i)) \ P x" + by (cases x) blast + +named_theorems nominal_prod_simps + +named_theorems ms_fresh "Facts for helping with freshness proofs" + +lemma fresh_prod2[nominal_prod_simps,ms_fresh]: "x \ (a,b) = (x \ a \ x \ b )" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod3[nominal_prod_simps,ms_fresh]: "x \ (a,b,c) = (x \ a \ x \ b \ x \ c)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod4[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d) = (x \ a \ x \ b \ x \ c \ x \ d)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod5[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod6[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod7[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod8[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h )" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod9[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod10[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i,j) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i \ x \ j)" + using fresh_def supp_Pair by fastforce + +lemma fresh_prod12[nominal_prod_simps,ms_fresh]: "x \ (a,b,c,d,e,f,g,h,i,j,k,l) = (x \ a \ x \ b \ x \ c \ x \ d \ x \ e \ x \ f \ x \ g \ x \ h \ x \ i \ x \ j \ x \ k \ x \ l)" + using fresh_def supp_Pair by fastforce + +lemmas fresh_prodN = fresh_Pair fresh_prod3 fresh_prod4 fresh_prod5 fresh_prod6 fresh_prod7 fresh_prod8 fresh_prod9 fresh_prod10 fresh_prod12 + +lemma fresh_prod2I: + fixes x and x1 and x2 + assumes "x \ x1" and "x \ x2" + shows "x \ (x1,x2)" using fresh_prod2 assms by auto + +lemma fresh_prod3I: + fixes x and x1 and x2 and x3 + assumes "x \ x1" and "x \ x2" and "x \ x3" + shows "x \ (x1,x2,x3)" using fresh_prod3 assms by auto + +lemma fresh_prod4I: + fixes x and x1 and x2 and x3 and x4 + assumes "x \ x1" and "x \ x2" and "x \ x3" and "x \ x4" + shows "x \ (x1,x2,x3,x4)" using fresh_prod4 assms by auto + +lemma fresh_prod5I: + fixes x and x1 and x2 and x3 and x4 and x5 + assumes "x \ x1" and "x \ x2" and "x \ x3" and "x \ x4" and "x \ x5" + shows "x \ (x1,x2,x3,x4,x5)" using fresh_prod5 assms by auto + +lemma flip_collapse[simp]: + fixes b1::"'a::pt" and bv1::"'b::at" and bv2::"'b::at" + assumes "atom bv2 \ b1" and "atom c \ (bv1,bv2,b1)" and "bv1 \ bv2" + shows "(bv2 \ c) \ (bv1 \ bv2) \ b1 = (bv1 \ c) \ b1" +proof - + have "c \ bv1" and "bv2 \ bv1" using assms by auto+ + hence "(bv2 \ c) + (bv1 \ bv2) + (bv2 \ c) = (bv1 \ c)" using flip_triple[of c bv1 bv2] flip_commute by metis + hence "(bv2 \ c) \ (bv1 \ bv2) \ (bv2 \ c) \ b1 = (bv1 \ c) \ b1" using permute_plus by metis + thus ?thesis using assms flip_fresh_fresh by force +qed + +lemma triple_eqvt[simp]: + "p \ (x, b,c) = (p \ x, p \ b , p \ c)" +proof - + have "(x,b,c) = (x,(b,c))" by simp + thus ?thesis using Pair_eqvt by simp +qed + +lemma lst_fst: + fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs" + assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))" + shows " ([[atom x]]lst. t1 = [[atom x']]lst. t1')" +proof - + have "(\c. atom c \ (t2,t2') \ atom c \ (x, x', t1, t1') \ (x \ c) \ t1 = (x' \ c) \ t1')" + proof(rule,rule,rule) + fix c::'a + assume "atom c \ (t2,t2')" and "atom c \ (x, x', t1, t1')" + hence "atom c \ (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp + thus " (x \ c) \ t1 = (x' \ c) \ t1'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp + qed + thus ?thesis using Abs1_eq_iff_all(3)[of x t1 x' t1' "(t2,t2')"] by simp +qed + +lemma lst_snd: + fixes x::"'a::at" and t1::"'b::fs" and x'::"'a::at" and t2::"'c::fs" + assumes " ([[atom x]]lst. (t1,t2) = [[atom x']]lst. (t1',t2'))" + shows " ([[atom x]]lst. t2 = [[atom x']]lst. t2')" +proof - + have "(\c. atom c \ (t1,t1') \ atom c \ (x, x', t2, t2') \ (x \ c) \ t2 = (x' \ c) \ t2')" + proof(rule,rule,rule) + fix c::'a + assume "atom c \ (t1,t1')" and "atom c \ (x, x', t2, t2')" + hence "atom c \ (x, x', (t1,t2), (t1',t2'))" using fresh_prod2 by simp + thus " (x \ c) \ t2 = (x' \ c) \ t2'" using assms Abs1_eq_iff_all(3) Pair_eqvt by simp + qed + thus ?thesis using Abs1_eq_iff_all(3)[of x t2 x' t2' "(t1,t1')"] by simp +qed + +lemma lst_head_cons_pair: + fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" + assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)" + shows "[[atom y1]]lst. (x1,xs1) = [[atom y2]]lst. (x2,xs2)" +proof(subst Abs1_eq_iff_all(3)[of y1 "(x1,xs1)" y2 "(x2,xs2)"],rule,rule,rule) + fix c::'a + assume "atom c \ (x1#xs1,x2#xs2)" and "atom c \ (y1, y2, (x1, xs1), x2, xs2)" + thus "(y1 \ c) \ (x1, xs1) = (y2 \ c) \ (x2, xs2)" using assms Abs1_eq_iff_all(3) by auto +qed + +lemma lst_head_cons_neq_nil: + fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" + assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (xs2)" + shows "xs2 \ []" +proof + assume as:"xs2 = []" + thus False using Abs1_eq_iff(3)[of y1 "x1#xs1" y2 Nil] assms as by auto +qed + +lemma lst_head_cons: + fixes y1::"'a ::at" and y2::"'a::at" and x1::"'b::fs" and x2::"'b::fs" and xs1::"('b::fs) list" and xs2::"('b::fs) list" + assumes "[[atom y1]]lst. (x1 # xs1) = [[atom y2]]lst. (x2 # xs2)" + shows "[[atom y1]]lst. x1 = [[atom y2]]lst. x2" and "[[atom y1]]lst. xs1 = [[atom y2]]lst. xs2" + using lst_head_cons_pair lst_fst lst_snd assms by metis+ + +lemma lst_pure: + fixes x1::"'a ::at" and t1::"'b::pure" and x2::"'a ::at" and t2::"'b::pure" + assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" + shows "t1=t2" + using assms Abs1_eq_iff_all(3) pure_fresh flip_fresh_fresh + by (metis Abs1_eq(3) permute_pure) + +lemma lst_supp: + assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" + shows "supp t1 - {atom x1} = supp t2 - {atom x2}" +proof - + have "supp ([[atom x1]]lst.t1) = supp ([[atom x2]]lst.t2)" using assms by auto + thus ?thesis using Abs_finite_supp + by (metis assms empty_set list.simps(15) supp_lst.simps) +qed + +lemma lst_supp_subset: + assumes "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "supp t1 \ {atom x1} \ B" + shows "supp t2 \ {atom x2} \ B" + using assms lst_supp by fast + +lemma projl_inl_eqvt: + fixes \ :: perm + shows "\ \ (projl (Inl x)) = projl (Inl (\ \ x))" +unfolding projl_def Inl_eqvt by simp + +end diff --git a/thys/MiniSail/Operational.thy b/thys/MiniSail/Operational.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Operational.thy @@ -0,0 +1,171 @@ +(*<*) +theory Operational + imports Typing +begin + (*>*) + +chapter \Operational Semantics\ + +text \ Here we define the operational semantics in terms of a small-step reduction relation. \ + +section \Reduction Rules\ + +text \ The store for mutable variables \ +type_synonym \ = "(u*v) list" + +nominal_function update_d :: "\ \ u \ v \ \" where + "update_d [] _ _ = []" +| "update_d ((u',v')#\) u v = (if u = u' then ((u,v)#\) else ((u',v')# (update_d \ u v)))" + by(auto,simp add: eqvt_def update_d_graph_aux_def ,metis neq_Nil_conv old.prod.exhaust) +nominal_termination (eqvt) by lexicographic_order + +text \ Relates constructor to the branch in the case and binding variable and statement \ +inductive find_branch :: "dc \ branch_list \ branch_s \ bool" where + find_branch_finalI: "dc' = dc \ find_branch dc' (AS_final (AS_branch dc x s )) (AS_branch dc x s)" +| find_branch_branch_eqI: "dc' = dc \ find_branch dc' (AS_cons (AS_branch dc x s) css) (AS_branch dc x s)" +| find_branch_branch_neqI: "\ dc \ dc'; find_branch dc' css cs \ \ find_branch dc' (AS_cons (AS_branch dc x s) css) cs" +equivariance find_branch +nominal_inductive find_branch . + +inductive_cases find_branch_elims[elim!]: + "find_branch dc (AS_final cs') cs" + "find_branch dc (AS_cons cs' css) cs" + +nominal_function lookup_branch :: "dc \ branch_list \ branch_s option" where + "lookup_branch dc (AS_final (AS_branch dc' x s)) = (if dc = dc' then (Some (AS_branch dc' x s)) else None)" +| "lookup_branch dc (AS_cons (AS_branch dc' x s) css) = (if dc = dc' then (Some (AS_branch dc' x s)) else lookup_branch dc css)" + apply(auto,simp add: eqvt_def lookup_branch_graph_aux_def) + by(metis neq_Nil_conv old.prod.exhaust s_branch_s_branch_list.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +text \ Reduction rules \ +inductive reduce_stmt :: "\ \ \ \ s \ \ \ s \ bool" (" _ \ \ _ , _\ \ \ _ , _\" [50, 50, 50] 50) where + reduce_if_trueI: " \ \ \\, AS_if [L_true]\<^sup>v s1 s2\ \ \\, s1\ " +| reduce_if_falseI: " \ \ \\, AS_if [L_false]\<^sup>v s1 s2\ \ \\, s2\ " +| reduce_let_valI: " \ \ \\, AS_let x (AE_val v) s\ \ \\, s[x::=v]\<^sub>s\<^sub>v\" +| reduce_let_plusI: " \ \ \\, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \ " +| reduce_let_leqI: "b = (if (n1 \ n2) then L_true else L_false) \ + \ \ \\, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit b)) s\" +| reduce_let_eqI: "b = (if (n1 = n2) then L_true else L_false) \ + \ \ \\, AS_let x ((AE_op Eq (V_lit n1) (V_lit n2))) s\ \ + \\, AS_let x (AE_val (V_lit b)) s\" +| reduce_let_appI: "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ z b c \ s'))) = lookup_fun \ f \ + \ \ \\, AS_let x ((AE_app f v)) s\ \ \\, AS_let2 x \[z::=v]\<^sub>\\<^sub>v s'[z::=v]\<^sub>s\<^sub>v s\ " +| reduce_let_appPI: "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \ s'))) = lookup_fun \ f \ + \ \ \\, AS_let x ((AE_appP f b' v)) s\ \ \\, AS_let2 x \[bv::=b']\<^sub>\\<^sub>b[z::=v]\<^sub>\\<^sub>v s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s\ " +| reduce_let_fstI: "\ \ \\, AS_let x (AE_fst (V_pair v1 v2)) s\ \ \\, AS_let x (AE_val v1) s\" +| reduce_let_sndI: "\ \ \\, AS_let x (AE_snd (V_pair v1 v2)) s\ \ \\, AS_let x (AE_val v2) s\" +| reduce_let_concatI: "\ \ \\, AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s\ \ + \\, AS_let x (AE_val (V_lit (L_bitvec (v1@v2)))) s\" +| reduce_let_splitI: " split n v (v1 , v2 ) \ \ \ \\, AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s\ \ + \\, AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s\" +| reduce_let_lenI: "\ \ \\, AS_let x (AE_len (V_lit (L_bitvec v))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num (int (List.length v))))) s\" +| reduce_let_mvar: "(u,v) \ set \ \ \ \ \\, AS_let x (AE_mvar u) s\ \ \\, AS_let x (AE_val v) s \" +| reduce_assert1I: "\ \ \\, AS_assert c (AS_val v)\ \ \\, AS_val v\" +| reduce_assert2I: " \ \ \\, s \ \ \ \', s'\ \ \ \ \\, AS_assert c s\ \ \ \', AS_assert c s'\" +| reduce_varI: "atom u \ \ \ \ \ \\, AS_var u \ v s \ \ \ ((u,v)#\) , s\" +| reduce_assignI: " \ \ \\, AS_assign u v \ \ \ update_d \ u v , AS_val (V_lit L_unit)\" +| reduce_seq1I: "\ \ \\, AS_seq (AS_val (V_lit L_unit )) s \ \ \\, s\" +| reduce_seq2I: "\s1 \ AS_val v ; \ \ \\, s1\ \ \ \', s1'\ \ \ + \ \ \\, AS_seq s1 s2\ \ \ \', AS_seq s1' s2\" +| reduce_let2_valI: "\ \ \\, AS_let2 x t (AS_val v) s\ \ \\, s[x::=v]\<^sub>s\<^sub>v\" +| reduce_let2I: " \ \ \\, s1 \ \ \ \', s1'\ \ \ \ \\, AS_let2 x t s1 s2\ \ \ \', AS_let2 x t s1' s2\" + +| reduce_caseI: "\ Some (AS_branch dc x' s') = lookup_branch dc cs \ \ \ \ \\, AS_match (V_cons tyid dc v) cs\ \ \\, s'[x'::=v]\<^sub>s\<^sub>v\ " +| reduce_whileI: "\ atom x \ (s1,s2); atom z \ x \ \ + \ \ \\, AS_while s1 s2 \ \ + \\, AS_let2 x (\ z : B_bool | TRUE \) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit))) \" + +equivariance reduce_stmt +nominal_inductive reduce_stmt . + +inductive_cases reduce_stmt_elims[elim!]: + "\ \ \\, AS_if (V_lit L_true) s1 s2\ \ \\ , s1\" + "\ \ \\, AS_if (V_lit L_false) s1 s2\ \ \\,s2\" + "\ \ \\, AS_let x (AE_val v) s\ \ \\,s'\" + "\ \ \\, AS_let x (AE_op Plus ((V_lit (L_num n1))) ((V_lit (L_num n2)))) s\ \ + \\, AS_let x (AE_val (V_lit (L_num ( (( n1)+(n2)))))) s \" + "\ \ \\, AS_let x ((AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)))) s\ \ \\, AS_let x (AE_val (V_lit b)) s\" + "\ \ \\, AS_let x ((AE_app f v)) s\ \ \\, AS_let2 x \ (subst_sv s' x v ) s\ " + "\ \ \\, AS_let x ((AE_len v)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_let x ((AE_concat v1 v2)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_seq s1 s2\ \ \ \' ,s'\" + "\ \ \\, AS_let x ((AE_appP f b v)) s\ \ \\, AS_let2 x \ (subst_sv s' z v) s\ " + "\ \ \\, AS_let x ((AE_split v1 v2)) s\ \ \\, AS_let x v' s\ " + "\ \ \\, AS_assert c s \ \ \\, s'\ " + "\ \ \\, AS_let x ((AE_op Eq (V_lit (n1)) (V_lit (n2)))) s\ \ \\, AS_let x (AE_val (V_lit b)) s\" + +inductive reduce_stmt_many :: "\ \ \ \ s \ \ \ s \ bool" ("_ \ \ _ , _\ \\<^sup>* \ _ , _\" [50, 50, 50] 50) where + reduce_stmt_many_oneI: "\ \ \\, s\ \ \\', s'\ \ \ \ \\ , s\ \\<^sup>* \\', s'\ " +| reduce_stmt_many_manyI: "\ \ \ \\, s\ \ \\', s'\ ; \ \ \\', s'\ \\<^sup>* \\'', s''\ \ \ \ \ \\, s\ \\<^sup>* \\'', s''\" + +nominal_function convert_fds :: "fun_def list \ (f*fun_def) list" where + "convert_fds [] = []" +| "convert_fds ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))#fs) = ((f,AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))#convert_fds fs)" +| "convert_fds ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)))#fs) = ((f,AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)))#convert_fds fs)" + apply(auto) + apply (simp add: eqvt_def convert_fds_graph_aux_def ) + using fun_def.exhaust fun_typ.exhaust fun_typ_q.exhaust neq_Nil_conv + by metis +nominal_termination (eqvt) by lexicographic_order + +nominal_function convert_tds :: "type_def list \ (f*type_def) list" where + "convert_tds [] = []" +| "convert_tds ((AF_typedef s dclist)#fs) = ((s,AF_typedef s dclist)#convert_tds fs)" +| "convert_tds ((AF_typedef_poly s bv dclist)#fs) = ((s,AF_typedef_poly s bv dclist)#convert_tds fs)" + apply(auto) + apply (simp add: eqvt_def convert_tds_graph_aux_def ) + by (metis type_def.exhaust neq_Nil_conv) +nominal_termination (eqvt) by lexicographic_order + +inductive reduce_prog :: "p \ v \ bool" where + "\ reduce_stmt_many \ [] s \ (AS_val v) \ \ reduce_prog (AP_prog \ \ [] s) v" + +section \Reduction Typing\ + +text \ Checks that the store is consistent with @{typ \} \ +inductive delta_sim :: "\ \ \ \ \ \ bool" ( "_ \ _ \ _ " [50,50] 50 ) where + delta_sim_nilI: "\ \ [] \ []\<^sub>\ " +| delta_sim_consI: "\ \ \ \ \ \ ; \ ; {||} ; GNil \ v \ \ ; u \ fst ` set \ \ \ \ \ ((u,v)#\) \ ((u,\)#\<^sub>\\)" + +equivariance delta_sim +nominal_inductive delta_sim . + +inductive_cases delta_sim_elims[elim!]: + "\ \ [] \ []\<^sub>\" + "\ \ ((u,v)#ds) \ (u,\) #\<^sub>\ D" + "\ \ ((u,v)#ds) \ D" + +text \A typing judgement that combines typing of the statement, the store and the condition that definitions are well-typed\ +inductive config_type :: "\ \ \ \ \ \ \ \ s \ \ \ bool" ("_ ; _ ; _ \ \ _ , _\ \ _ " [50, 50, 50] 50) where + config_typeI: "\ \ ; \ ; {||} ; GNil ; \ \ s \ \; + (\ fd \ set \. \ ; \ \ fd) ; + \ \ \ \ \ \ + \ \ ; \ ; \ \ \ \ , s\ \ \" +equivariance config_type +nominal_inductive config_type . + +inductive_cases config_type_elims [elim!]: + " \ ; \ ; \ \ \ \ , s\ \ \" + +nominal_function \_of :: "var_def list \ \" where + "\_of [] = []" +| "\_of ((AV_def u t v)#vs) = (u,v) # (\_of vs)" + apply auto + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust + by (metis var_def.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive config_type_prog :: "p \ \ \ bool" (" \ \ _\ \ _") where + "\ + \ ; \ ; \_of \ \ \ \_of \ , s\ \ \ +\ \ \ \ AP_prog \ \ \ s\ \ \" + +inductive_cases config_type_prog_elims [elim!]: + "\ \ AP_prog \ \ \ s\ \ \" + +end diff --git a/thys/MiniSail/RCLogic.thy b/thys/MiniSail/RCLogic.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/RCLogic.thy @@ -0,0 +1,222 @@ +(*<*) +theory RCLogic + imports Wellformed +begin + (*>*) + +hide_const Syntax.dom + +chapter \Refinement Constraint Logic\ + +text \ Semantics for the logic we use in the refinement constraints. It is a multi-sorted, quantifier free +logic with polymorphic datatypes and linear arithmetic. We could have modelled by using one of the +encodings to FOL however we wanted to explore using a more direct model. \ + +section \Evaluation and Satisfiability\ + +subsection \Valuation\ + +text \ Refinement constraint logic values. SUt is a value for the uninterpreted + sort that corresponds to basic type variables. For now we only need one of these universes. + We wrap an smt\_val inside it during a process we call 'boxing' + which is introduced in the RCLogicL theory \ +nominal_datatype rcl_val = SBitvec "bit list" | SNum int | SBool bool | SPair rcl_val rcl_val | + SCons tyid string rcl_val | SConsp tyid string b rcl_val | + SUnit | SUt rcl_val + +text \ RCL sorts. Represent our domains. The universe is the union of all of the these. + S\_Ut is the single uninterpreted sort. These map almost directly to basic type + but we have them to distinguish syntax (basic types) and semantics (RCL sorts) \ +nominal_datatype rcl_sort = S_bool | S_int | S_unit | S_pair rcl_sort rcl_sort | S_id tyid | S_app tyid rcl_sort | S_bitvec | S_ut + +type_synonym valuation = "(x,rcl_val) map" + +type_synonym type_valuation = "(bv,rcl_sort) map" + +text \Well-sortedness for RCL values\ +inductive wfRCV:: "\ \ rcl_val \ b \ bool" ( " _ \ _ : _" [50,50] 50) where + wfRCV_BBitvecI: "P \ (SBitvec bv) : B_bitvec" +| wfRCV_BIntI: "P \ (SNum n) : B_int" +| wfRCV_BBoolI: "P \ (SBool b) : B_bool" +| wfRCV_BPairI: "\ P \ s1 : b1 ; P \ s2 : b2 \ \ P \ (SPair s1 s2) : (B_pair b1 b2)" +| wfRCV_BConsI: "\ AF_typedef s dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + \ \ s1 : b \ \ \ \(SCons s dc s1) : (B_id s)" +| wfRCV_BConsPI:"\ AF_typedef_poly s bv dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + atom bv \ (\, SConsp s dc b' s1, B_app s b'); + \ \ s1 : b[bv::=b']\<^sub>b\<^sub>b \ \ \ \(SConsp s dc b' s1) : (B_app s b')" +| wfRCV_BUnitI: "P \ SUnit : B_unit" +| wfRCV_BVarI: "P \ (SUt n) : (B_var bv)" +equivariance wfRCV +nominal_inductive wfRCV + avoids wfRCV_BConsPI: bv +proof(goal_cases) + case (1 s bv dclist \ dc x b c b' s1) + then show ?case using fresh_star_def by auto +next + case (2 s bv dclist \ dc x b c s1 b') + then show ?case by auto +qed + +inductive_cases wfRCV_elims : + "wfRCV P s B_bitvec" + "wfRCV P s (B_pair b1 b2)" + "wfRCV P s (B_int)" + "wfRCV P s (B_bool)" + "wfRCV P s (B_id ss)" + "wfRCV P s (B_var bv)" + "wfRCV P s (B_unit)" + "wfRCV P s (B_app tyid b)" + "wfRCV P (SBitvec bv) b" + "wfRCV P (SNum n) b" + "wfRCV P (SBool n) b" + "wfRCV P (SPair s1 s2) b" + "wfRCV P (SCons s dc s1) b" + "wfRCV P (SConsp s dc b' s1) b" + "wfRCV P SUnit b" + "wfRCV P (SUt s1) b" + +text \ Sometimes we want to assert @{text "P \ s ~ b[bv=b']"} and we want to know what b is +however substitution is not injective so we can't write this in terms of @{text "wfRCV"}. +So we define a relation that makes the components of the substitution explicit. \ + +inductive wfRCV_subst:: "\ \ rcl_val \ b \ (bv*b) option \ bool" where + wfRCV_subst_BBitvecI: "wfRCV_subst P (SBitvec bv) B_bitvec sub " +| wfRCV_subst_BIntI: "wfRCV_subst P (SNum n) B_int sub " +| wfRCV_subst_BBoolI: "wfRCV_subst P (SBool b) B_bool sub " +| wfRCV_subst_BPairI: "\ wfRCV_subst P s1 b1 sub ; wfRCV_subst P s2 b2 sub \ \ wfRCV_subst P (SPair s1 s2) (B_pair b1 b2) sub" +| wfRCV_subst_BConsI: "\ AF_typedef s dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + wfRCV_subst \ s1 b None \ \ wfRCV_subst \ (SCons s dc s1) (B_id s) sub" +| wfRCV_subst_BConspI: "\ AF_typedef_poly s bv dclist \ set \; + (dc, \ x : b | c \) \ set dclist ; + wfRCV_subst \ s1 (b[bv::=b']\<^sub>b\<^sub>b) sub \ \ wfRCV_subst \ (SConsp s dc b' s1) (B_app s b') sub" +| wfRCV_subst_BUnitI: "wfRCV_subst P SUnit B_unit sub " +| wfRCV_subst_BVar1I: "bvar \ bv \ wfRCV_subst P (SUt n) (B_var bv) (Some (bvar,bin))" +| wfRCV_subst_BVar2I: "\ bvar = bv; wfRCV_subst P s bin None \ \ wfRCV_subst P s (B_var bv) (Some (bvar,bin))" +| wfRCV_subst_BVar3I: "wfRCV_subst P (SUt n) (B_var bv) None" +equivariance wfRCV_subst +nominal_inductive wfRCV_subst . + +subsection \Evaluation base-types\ + +inductive eval_b :: "type_valuation \ b \ rcl_sort \ bool" ( "_ \ _ \ ~ _ " ) where + "v \ B_bool \ ~ S_bool" +| "v \ B_int \ ~ S_int" +| "Some s = v bv \ v \ B_var bv \ ~ s" +equivariance eval_b +nominal_inductive eval_b . + +subsection \Wellformed vvaluations\ + +definition wfI :: "\ \ \ \ valuation \ bool" ( " _ ; _ \ _" ) where + "\ ; \ \ i = (\ (x,b,c) \ toSet \. \s. Some s = i x \ \ \ s : b)" + +subsection \Evaluating Terms\ + +nominal_function eval_l :: "l \ rcl_val" ( "\ _ \ " ) where + "\ L_true \ = SBool True" +| "\ L_false \ = SBool False" +| "\ L_num n \ = SNum n" +| "\ L_unit \ = SUnit" +| "\ L_bitvec n \ = SBitvec n" + apply(auto simp: eqvt_def eval_l_graph_aux_def) + by (metis l.exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive eval_v :: "valuation \ v \ rcl_val \ bool" ( "_ \ _ \ ~ _ " ) where + eval_v_litI: "i \ V_lit l \ ~ \ l \ " +| eval_v_varI: "Some sv = i x \ i \ V_var x \ ~ sv" +| eval_v_pairI: "\ i \ v1 \ ~ s1 ; i \ v2 \ ~ s2 \ \ i \ V_pair v1 v2 \ ~ SPair s1 s2" +| eval_v_consI: "i \ v \ ~ s \ i \ V_cons tyid dc v \ ~ SCons tyid dc s" +| eval_v_conspI: "i \ v \ ~ s \ i \ V_consp tyid dc b v \ ~ SConsp tyid dc b s" +equivariance eval_v +nominal_inductive eval_v . + +inductive_cases eval_v_elims: + "i \ V_lit l \ ~ s" + "i \ V_var x \ ~ s" + "i \ V_pair v1 v2 \ ~ s" + "i \ V_cons tyid dc v \ ~ s" + "i \ V_consp tyid dc b v \ ~ s" + +inductive eval_e :: "valuation \ ce \ rcl_val \ bool" ( "_ \ _ \ ~ _ " ) where + eval_e_valI: "i \ v \ ~ sv \ i \ CE_val v \ ~ sv" +| eval_e_plusI: "\ i \ v1 \ ~ SNum n1; i \ v2 \ ~ SNum n2 \ \ i \ (CE_op Plus v1 v2) \ ~ (SNum (n1+n2))" +| eval_e_leqI: "\ i \ v1 \ ~ (SNum n1); i \ v2 \ ~ (SNum n2) \ \ i \ (CE_op LEq v1 v2) \ ~ (SBool (n1 \ n2))" +| eval_e_eqI: "\ i \ v1 \ ~ s1; i \ v2 \ ~ s2 \ \ i \ (CE_op Eq v1 v2) \ ~ (SBool (s1 = s2))" +| eval_e_fstI: "\ i \ v \ ~ SPair v1 v2 \ \ i \ (CE_fst v) \ ~ v1" +| eval_e_sndI: "\ i \ v \ ~ SPair v1 v2 \ \ i \ (CE_snd v) \ ~ v2" +| eval_e_concatI:"\ i \ v1 \ ~ (SBitvec bv1); i \ v2 \ ~ (SBitvec bv2) \ \ i \ (CE_concat v1 v2) \ ~ (SBitvec (bv1@bv2))" +| eval_e_lenI:"\ i \ v \ ~ (SBitvec bv) \ \ i \ (CE_len v) \ ~ (SNum (int (List.length bv)))" +equivariance eval_e +nominal_inductive eval_e . + +inductive_cases eval_e_elims: + "i \ (CE_val v) \ ~ s" + "i \ (CE_op Plus v1 v2) \ ~ s" + "i \ (CE_op LEq v1 v2) \ ~ s" + "i \ (CE_op Eq v1 v2) \ ~ s" + "i \ (CE_fst v) \ ~ s" + "i \ (CE_snd v) \ ~ s" + "i \ (CE_concat v1 v2) \ ~ s" + "i \ (CE_len v) \ ~ s" + +inductive eval_c :: "valuation \ c \ bool \ bool" ( " _ \ _ \ ~ _ ") where + eval_c_trueI: "i \ C_true \ ~ True" +| eval_c_falseI:"i \ C_false \ ~ False" +| eval_c_conjI: "\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_conj c1 c2) \ ~ (b1 \ b2)" +| eval_c_disjI: "\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_disj c1 c2) \ ~ (b1 \ b2)" +| eval_c_impI:"\ i \ c1 \ ~ b1 ; i \ c2 \ ~ b2 \ \ i \ (C_imp c1 c2) \ ~ (b1 \ b2)" +| eval_c_notI:"\ i \ c \ ~ b \ \ i \ (C_not c) \ ~ (\ b)" +| eval_c_eqI:"\ i \ e1 \ ~ sv1; i \ e2 \ ~ sv2 \ \ i \ (C_eq e1 e2) \ ~ (sv1=sv2)" +equivariance eval_c +nominal_inductive eval_c . + +inductive_cases eval_c_elims: + "i \ C_true \ ~ True" + "i \ C_false \ ~ False" + "i \ (C_conj c1 c2)\ ~ s" + "i \ (C_disj c1 c2)\ ~ s" + "i \ (C_imp c1 c2)\ ~ s" + "i \ (C_not c) \ ~ s" + "i \ (C_eq e1 e2)\ ~ s" + "i \ C_true \ ~ s" + "i \ C_false \ ~ s" + +subsection \Satisfiability\ + +inductive is_satis :: "valuation \ c \ bool" ( " _ \ _ " ) where + "i \ c \ ~ True \ i \ c" +equivariance is_satis +nominal_inductive is_satis . + +nominal_function is_satis_g :: "valuation \ \ \ bool" ( " _ \ _ " ) where + "i \ GNil = True" +| "i \ ((x,b,c) #\<^sub>\ G) = ( i \ c \ i \ G)" + apply(auto simp: eqvt_def is_satis_g_graph_aux_def) + by (metis \.exhaust old.prod.exhaust) +nominal_termination (eqvt) by lexicographic_order + +section \Validity\ + +nominal_function valid :: "\ \ \ \ \ \ c \ bool" ("_ ; _ ; _ \ _ " [50, 50] 50) where + "P ; B ; G \ c = ( (P ; B ; G \\<^sub>w\<^sub>f c) \ (\i. (P ; G \ i) \ i \ G \ i \ c))" + by (auto simp: eqvt_def wfI_def valid_graph_aux_def) +nominal_termination (eqvt) by lexicographic_order + +section \Lemmas\ +text \Lemmas needed for Examples\ + +lemma valid_trueI [intro]: + fixes G::\ + assumes "P ; B \\<^sub>w\<^sub>f G" + shows "P ; B ; G \ C_true" +proof - + have "\i. i \ C_true" using is_satis.simps eval_c_trueI by simp + moreover have "P ; B ; G \\<^sub>w\<^sub>f C_true" using wfC_trueI assms by simp + ultimately show ?thesis using valid.simps by simp +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/RCLogicL.thy b/thys/MiniSail/RCLogicL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/RCLogicL.thy @@ -0,0 +1,2729 @@ +(*<*) +theory RCLogicL + imports RCLogic WellformedL +begin + (*>*) + +hide_const Syntax.dom + +chapter \Refinement Constraint Logic Lemmas\ + +section \Lemmas\ + +lemma wfI_domi: + assumes "\ ; \ \ i" + shows "fst ` toSet \ \ dom i" + using wfI_def toSet.simps assms by fastforce + +lemma wfI_lookup: + fixes G::\ and b::b + assumes "Some (b,c) = lookup G x" and "P ; G \ i" and "Some s = i x" and "P ; B \\<^sub>w\<^sub>f G" + shows "P \ s : b" +proof - + have "(x,b,c) \ toSet G" using lookup.simps assms + using lookup_in_g by blast + then obtain s' where *:"Some s' = i x \ wfRCV P s' b" using wfI_def assms by auto + hence "s' = s" using assms by (metis option.inject) + thus ?thesis using * by auto +qed + +lemma wfI_restrict_weakening: + assumes "wfI \ \' i'" and "i = restrict_map i' (fst `toSet \)" and "toSet \ \ toSet \'" + shows "\ ; \ \ i" +proof - + { fix x + assume "x \ toSet \" + have "case x of (x, b, c) \ \s. Some s = i x \ \ \ s : b" using assms wfI_def + proof - + have "case x of (x, b, c) \ \s. Some s = i' x \ \ \ s : b" + using \x \ toSet \\ assms wfI_def by auto + then have "\s. Some s = i (fst x) \ wfRCV \ s (fst (snd x))" + by (simp add: \x \ toSet \\ assms(2) case_prod_unfold) + then show ?thesis + by (simp add: case_prod_unfold) + qed + } + thus ?thesis using wfI_def assms by auto +qed + +lemma wfI_suffix: + fixes G::\ + assumes "wfI P (G'@G) i" and "P ; B \\<^sub>w\<^sub>f G" + shows "P ; G \ i" + using wfI_def append_g.simps assms toSet.simps by simp + +lemma wfI_replace_inside: + assumes "wfI \ (\' @ (x, b, c) #\<^sub>\ \) i" + shows "wfI \ (\' @ (x, b, c') #\<^sub>\ \) i" + using wfI_def toSet_splitP assms by simp + +section \Existence of evaluation\ + +lemma eval_l_base: + "\ \ \ l \ : (base_for_lit l)" + apply(nominal_induct l rule:l.strong_induct) + using wfRCV.intros eval_l.simps base_for_lit.simps by auto+ + +lemma obtain_fresh_bv_dclist: + fixes tm::"'a::fs" + assumes "(dc, \ x : b | c \) \ set dclist" + obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ tm" +proof - + obtain bv1 dclist1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \ atom bv1 \ tm" + using obtain_fresh_bv by metis + moreover hence "[[atom bv]]lst. dclist = [[atom bv1]]lst. dclist1" using type_def.eq_iff by metis + moreover then obtain x1 b1 c1 where "(dc, \ x1 : b1 | c1 \) \ set dclist1" using td_lookup_eq_iff assms by metis + ultimately show ?thesis using that by blast +qed + +lemma obtain_fresh_bv_dclist_b_iff: + fixes tm::"'a::fs" + assumes "(dc, \ x : b | c \) \ set dclist" and "AF_typedef_poly tyid bv dclist \ set P" and "\\<^sub>w\<^sub>f P" + obtains bv1::bv and dclist1 x1 b1 c1 where "AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ tm \ b[bv::=b']\<^sub>b\<^sub>b=b1[bv1::=b']\<^sub>b\<^sub>b" +proof - + obtain bv1 dclist1 x1 b1 c1 where *:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 \ atom bv1 \ tm + \ (dc, \ x1 : b1 | c1 \) \ set dclist1" using obtain_fresh_bv_dclist assms by metis + + hence "AF_typedef_poly tyid bv1 dclist1 \ set P" using assms by metis + + hence "b[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b" + using wfTh_typedef_poly_b_eq_iff[OF assms(2) assms(1) _ _ assms(3),of bv1 dclist1 x1 b1 c1 b'] * by metis + + from this that show ?thesis using * by metis +qed + +lemma eval_v_exist: + fixes \::\ and v::v and b::b + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "\s. i \ v \ ~ s \ P \ s : b " + using assms proof(nominal_induct v arbitrary: b rule:v.strong_induct) + case (V_lit x) + then show ?case using eval_l_base eval_v.intros eval_l.simps wfV_elims rcl_val.supp pure_supp by metis +next + case (V_var x) + then obtain c where *:"Some (b,c) = lookup \ x" using wfV_elims by metis + hence "x \ fst ` toSet \" using lookup_x by blast + hence "x \ dom i" using wfI_domi using assms by blast + then obtain s where "i x = Some s" by auto + moreover hence "P \ s : b" using wfRCV.intros wfI_lookup * V_var + by (metis wfV_wf) + + ultimately show ?case using eval_v.intros rcl_val.supp b.supp by metis +next + case (V_pair v1 v2) + then obtain b1 and b2 where *:"P ; B ; \ \\<^sub>w\<^sub>f v1 : b1 \ P ; B ; \ \\<^sub>w\<^sub>f v2 : b2 \ b = B_pair b1 b2" using wfV_elims by metis + then obtain s1 and s2 where "eval_v i v1 s1 \ wfRCV P s1 b1 \ eval_v i v2 s2 \ wfRCV P s2 b2" using V_pair by metis + thus ?case using eval_v.intros wfRCV.intros * by metis +next + case (V_cons tyid dc v) + then obtain s and b'::b and dclist and x::x and c::c where "(wfV P B \ v b') \ i \ v \ ~ s \ P \ s : b' \ b = B_id tyid \ + AF_typedef tyid dclist \ set P \ (dc, \ x : b' | c \) \ set dclist" using wfV_elims(4) by metis + then show ?case using eval_v.intros(4) wfRCV.intros(5) V_cons by metis +next + case (V_consp tyid dc b' v) + + obtain b'a::b and bv and dclist and x::x and c::c where *:"(wfV P B \ v b'a[bv::=b']\<^sub>b\<^sub>b) \ b = B_app tyid b' \ + AF_typedef_poly tyid bv dclist \ set P \ (dc, \ x : b'a | c \) \ set dclist \ + atom bv \ (P, B_app tyid b',B) " using wf_strong_elim(1)[OF V_consp(3)] by metis + + then obtain s where **:"i \ v \ ~ s \ P \ s : b'a[bv::=b']\<^sub>b\<^sub>b" using V_consp by auto + + have " \\<^sub>w\<^sub>f P" using wfX_wfY V_consp by metis + then obtain bv1::bv and dclist1 x1 b1 c1 where 3:"AF_typedef_poly tyid bv dclist = AF_typedef_poly tyid bv1 dclist1 + \ (dc, \ x1 : b1 | c1 \) \ set dclist1 \ atom bv1 \ (P, SConsp tyid dc b' s, B_app tyid b') + \ b'a[bv::=b']\<^sub>b\<^sub>b = b1[bv1::=b']\<^sub>b\<^sub>b" + using * obtain_fresh_bv_dclist_b_iff by blast + + have " i \ V_consp tyid dc b' v \ ~ SConsp tyid dc b' s" using eval_v.intros by (simp add: "**") + + moreover have " P \ SConsp tyid dc b' s : B_app tyid b'" proof + show \AF_typedef_poly tyid bv1 dclist1 \ set P\ using 3 * by metis + show \(dc, \ x1 : b1 | c1 \) \ set dclist1\ using 3 by auto + thus \atom bv1 \ (P, SConsp tyid dc b' s, B_app tyid b')\ using * 3 fresh_prodN by metis + show \ P \ s : b1[bv1::=b']\<^sub>b\<^sub>b\ using 3 ** by auto + qed + + ultimately show ?case using eval_v.intros wfRCV.intros V_consp * by auto +qed + +lemma eval_v_uniqueness: + fixes v::v + assumes "i \ v \ ~ s" and "i \ v \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct v arbitrary: s s' rule:v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_l.simps by metis +next + case (V_var x) + then show ?case using eval_v_elims by (metis option.inject) +next + case (V_pair v1 v2) + obtain s1 and s2 where s:"i \ v1 \ ~ s1 \ i \ v2 \ ~ s2 \ s = SPair s1 s2" using eval_v_elims V_pair by metis + obtain s1' and s2' where s':"i \ v1 \ ~ s1' \ i \ v2 \ ~ s2' \ s' = SPair s1' s2'" using eval_v_elims V_pair by metis + then show ?case using eval_v_elims using V_pair s s' by auto +next + case (V_cons tyid dc v1) + obtain sv1 where 1:"i \ v1 \ ~ sv1 \ s = SCons tyid dc sv1" using eval_v_elims V_cons by metis + moreover obtain sv2 where 2:"i \ v1 \ ~ sv2 \ s' = SCons tyid dc sv2" using eval_v_elims V_cons by metis + ultimately have "sv1 = sv2" using V_cons by auto + then show ?case using 1 2 by auto +next + case (V_consp tyid dc b v1) + then show ?case using eval_v_elims by metis + +qed + +lemma eval_v_base: + fixes \::\ and v::v and b::b + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f v : b" and "i \ v \ ~ s" + shows "P \ s : b" + using eval_v_exist eval_v_uniqueness assms by metis + +lemma eval_e_uniqueness: + fixes e::ce + assumes "i \ e \ ~ s" and "i \ e \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct e arbitrary: s s' rule:ce.strong_induct) + case (CE_val x) + then show ?case using eval_v_uniqueness eval_e_elims by metis +next + case (CE_op opp x1 x2) + consider "opp = Plus" | "opp = LEq" | "opp = Eq" using opp.exhaust by metis + thus ?case proof(cases) + case 1 + hence a1:"eval_e i (CE_op Plus x1 x2) s" and a2:"eval_e i (CE_op Plus x1 x2) s'" using CE_op by auto + then show ?thesis using eval_e_elims(2)[OF a1] eval_e_elims(2)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + next + case 2 + hence a1:"eval_e i (CE_op LEq x1 x2) s" and a2:"eval_e i (CE_op LEq x1 x2) s'" using CE_op by auto + then show ?thesis using eval_v_uniqueness eval_e_elims(3)[OF a1] eval_e_elims(3)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + next + case 3 + hence a1:"eval_e i (CE_op Eq x1 x2) s" and a2:"eval_e i (CE_op Eq x1 x2) s'" using CE_op by auto + then show ?thesis using eval_v_uniqueness eval_e_elims(4)[OF a1] eval_e_elims(4)[OF a2] + CE_op eval_e_plusI + by (metis rcl_val.eq_iff(2)) + qed +next + case (CE_concat x1 x2) + hence a1:"eval_e i (CE_concat x1 x2) s" and a2:"eval_e i (CE_concat x1 x2) s'" using CE_concat by auto + show ?case using eval_e_elims(7)[OF a1] eval_e_elims(7)[OF a2] CE_concat eval_e_concatI rcl_val.eq_iff + proof - + assume "\P. (\bv1 bv2. \s' = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P" + obtain bbs :: "bit list" and bbsa :: "bit list" where + "i \ x2 \ ~ SBitvec bbs \ i \ x1 \ ~ SBitvec bbsa \ SBitvec (bbsa @ bbs) = s'" + by (metis \\P. (\bv1 bv2. \s' = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P\) (* 93 ms *) + then have "s' = s" + by (metis (no_types) \\P. (\bv1 bv2. \s = SBitvec (bv1 @ bv2); i \ x1 \ ~ SBitvec bv1 ; i \ x2 \ ~ SBitvec bv2 \ \ P) \ P\ \\s' s. \i \ x1 \ ~ s ; i \ x1 \ ~ s' \ \ s = s'\ \\s' s. \i \ x2 \ ~ s ; i \ x2 \ ~ s' \ \ s = s'\ rcl_val.eq_iff(1)) (* 125 ms *) + then show ?thesis + by metis (* 0.0 ms *) + qed +next + case (CE_fst x) + then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff) +next + case (CE_snd x) + then show ?case using eval_v_uniqueness by (meson eval_e_elims rcl_val.eq_iff) +next + case (CE_len x) + then show ?case using eval_e_elims rcl_val.eq_iff + proof - + obtain bbs :: "rcl_val \ ce \ (x \ rcl_val option) \ bit list" where + "\x0 x1 x2. (\v3. x0 = SNum (int (length v3)) \ x2 \ x1 \ ~ SBitvec v3 ) = (x0 = SNum (int (length (bbs x0 x1 x2))) \ x2 \ x1 \ ~ SBitvec (bbs x0 x1 x2) )" + by moura (* 0.0 ms *) + then have "\f c r. \ f \ [| c |]\<^sup>c\<^sup>e \ ~ r \ r = SNum (int (length (bbs r c f))) \ f \ c \ ~ SBitvec (bbs r c f)" + by (meson eval_e_elims(8)) (* 46 ms *) + then show ?thesis + by (metis (no_types) CE_len.hyps CE_len.prems(1) CE_len.prems(2) rcl_val.eq_iff(1)) (* 31 ms *) + qed + +qed + +lemma wfV_eval_bitvec: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_bitvec" and "P ; \ \ i" + shows "\bv. eval_v i v (SBitvec bv)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s B_bitvec" using eval_v_exist assms by metis + moreover then obtain bv where "s = SBitvec bv" using wfRCV_elims(1)[of P s] by metis + ultimately show ?thesis by metis +qed + +lemma wfV_eval_pair: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_pair b1 b2" and "P ; \ \ i" + shows "\s1 s2. eval_v i v (SPair s1 s2)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s (B_pair b1 b2)" using eval_v_exist assms by metis + moreover then obtain s1 and s2 where "s = SPair s1 s2" using wfRCV_elims(2)[of P s] by metis + ultimately show ?thesis by metis +qed + +lemma wfV_eval_int: + fixes v::v + assumes "P ; B ; \ \\<^sub>w\<^sub>f v : B_int" and "P ; \ \ i" + shows "\n. eval_v i v (SNum n)" +proof - + obtain s where "i \ v \ ~ s \ wfRCV P s (B_int)" using eval_v_exist assms by metis + moreover then obtain n where "s = SNum n" using wfRCV_elims(3)[of P s] by metis + ultimately show ?thesis by metis +qed + +text \Well sorted value with a well sorted valuation evaluates\ +lemma wfI_wfV_eval_v: + fixes v::v and b::b + assumes "\ ; B ; \ \\<^sub>w\<^sub>f v : b" and "wfI \ \ i" + shows "\s. i \ v \ ~ s \ \ \ s : b" + using eval_v_exist assms by auto + +lemma wfI_wfCE_eval_e: + fixes e::ce and b::b + assumes "wfCE P B G e b" and "P ; G \ i" + shows "\s. i \ e \ ~ s \ P \ s : b" + using assms proof(nominal_induct e arbitrary: b rule: ce.strong_induct) + case (CE_val v) + obtain s where "i \ v \ ~ s \ P \ s : b" using wfI_wfV_eval_v[of P B G v b i] assms wfCE_elims(1) [of P B G v b] CE_val by auto + then show ?case using CE_val eval_e.intros(1)[of i v s ] by auto +next + case (CE_op opp v1 v2) + + consider "opp =Plus" | "opp=LEq" | "opp=Eq" using opp.exhaust by auto + + thus ?case proof(cases) + case 1 + hence "wfCE P B G v1 B_int \ wfCE P B G v2 B_int" using wfCE_elims(2) CE_op + + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 B_int \ eval_e i v2 s2 \ wfRCV P s2 B_int" + using wfI_wfV_eval_v CE_op by metis + then obtain n1 and n2 where **:"s2=SNum n2 \ s1 = SNum n1" using wfRCV_elims by meson + hence "eval_e i (CE_op Plus v1 v2) (SNum (n1+n2))" using eval_e_plusI * ** by simp + moreover have "wfRCV P (SNum (n1+n2)) B_int" using wfRCV.intros by auto + ultimately show ?thesis using 1 + using CE_op.prems(1) wfCE_elims(2) by blast + next + case 2 + hence "wfCE P B G v1 B_int \ wfCE P B G v2 B_int" using wfCE_elims(3) CE_op + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 B_int \ eval_e i v2 s2 \ wfRCV P s2 B_int" + using wfI_wfV_eval_v CE_op by metis + then obtain n1 and n2 where **:"s2=SNum n2 \ s1 = SNum n1" using wfRCV_elims by meson + hence "eval_e i (CE_op LEq v1 v2) (SBool (n1 \ n2))" using eval_e_leqI * ** by simp + moreover have "wfRCV P (SBool (n1\n2)) B_bool" using wfRCV.intros by auto + ultimately show ?thesis using 2 + using CE_op.prems wfCE_elims by metis + next + case 3 + then obtain b2 where "wfCE P B G v1 b2 \ wfCE P B G v2 b2" using wfCE_elims(9) CE_op + by blast + then obtain s1 and s2 where *: "eval_e i v1 s1 \ wfRCV P s1 b2 \ eval_e i v2 s2 \ wfRCV P s2 b2" + using wfI_wfV_eval_v CE_op by metis + hence "eval_e i (CE_op Eq v1 v2) (SBool (s1 = s2))" using eval_e_leqI * + by (simp add: eval_e_eqI) + moreover have "wfRCV P (SBool (s1 = s2)) B_bool" using wfRCV.intros by auto + ultimately show ?thesis using 3 + using CE_op.prems wfCE_elims by metis + qed +next + case (CE_concat v1 v2) + then obtain s1 and s2 where *:"b = B_bitvec \ eval_e i v1 s1 \ eval_e i v2 s2 \ + wfRCV P s1 B_bitvec \ wfRCV P s2 B_bitvec" using + CE_concat + by (meson wfCE_elims(6)) + thus ?case using eval_e_concatI wfRCV.intros(1) wfRCV_elims + proof - + obtain bbs :: "type_def list \ rcl_val \ bit list" where + "\ts s. \ ts \ s : B_bitvec \ s = SBitvec (bbs ts s)" + using wfRCV_elims(1) by moura (* 156 ms *) + then show ?thesis + by (metis (no_types) "local.*" wfRCV_BBitvecI eval_e_concatI) (* 78 ms *) + qed +next + case (CE_fst v1) + thus ?case using eval_e_fstI wfRCV.intros wfCE_elims wfI_wfV_eval_v + by (metis wfRCV_elims(2) rcl_val.eq_iff(4)) +next + case (CE_snd v1) + thus ?case using eval_e_sndI wfRCV.intros wfCE_elims wfI_wfV_eval_v + by (metis wfRCV_elims(2) rcl_val.eq_iff(4)) +next + case (CE_len x) + thus ?case using eval_e_lenI wfRCV.intros wfCE_elims wfI_wfV_eval_v wfV_eval_bitvec + by (metis wfRCV_elims(1)) +qed + +lemma eval_e_exist: + fixes \::\ and e::ce + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f e : b" + shows "\s. i \ e \ ~ s" + using assms proof(nominal_induct e arbitrary: b rule:ce.strong_induct) + case (CE_val v) + then show ?case using eval_v_exist wfCE_elims eval_e.intros by metis +next + case (CE_op op v1 v2) + + show ?case proof(rule opp.exhaust) + assume \op = Plus\ + hence "P ; B ; \ \\<^sub>w\<^sub>f v1 : B_int \ P ; B ; \ \\<^sub>w\<^sub>f v2 : B_int \ b = B_int" using wfCE_elims CE_op by metis + then obtain n1 and n2 where "eval_e i v1 (SNum n1) \ eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_plusI[of i v1 _ v2] \op=Plus\ by auto + next + assume \op = LEq\ + hence "P ; B ; \ \\<^sub>w\<^sub>f v1 : B_int \ P ; B ; \ \\<^sub>w\<^sub>f v2 : B_int \ b = B_bool" using wfCE_elims CE_op by metis + then obtain n1 and n2 where "eval_e i v1 (SNum n1) \ eval_e i v2 (SNum n2)" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_leqI[of i v1 _ v2] eval_v_exist \op=LEq\ CE_op by auto + next + assume \op = Eq\ + then obtain b1 where "P ; B ; \ \\<^sub>w\<^sub>f v1 : b1 \ P ; B ; \ \\<^sub>w\<^sub>f v2 : b1 \ b = B_bool" using wfCE_elims CE_op by metis + then obtain s1 and s2 where "eval_e i v1 s1 \ eval_e i v2 s2" using CE_op eval_v_exist wfV_eval_int + by (metis wfI_wfCE_eval_e wfRCV_elims(3)) + then show \\a. eval_e i (CE_op op v1 v2) a\ using eval_e_eqI[of i v1 _ v2] eval_v_exist \op=Eq\ CE_op by auto + qed +next + case (CE_concat v1 v2) + then obtain bv1 and bv2 where "eval_e i v1 (SBitvec bv1) \ eval_e i v2 (SBitvec bv2)" + using wfV_eval_bitvec wfCE_elims(6) + by (meson eval_e_elims(7) wfI_wfCE_eval_e) + then show ?case using eval_e.intros by metis +next + case (CE_fst ce) + then obtain b2 where b:"P ; B ; \ \\<^sub>w\<^sub>f ce : B_pair b b2" using wfCE_elims by metis + then obtain s where s:"i \ ce \ ~ s" using CE_fst by auto + then obtain s1 and s2 where "s = (SPair s1 s2)" using eval_e_elims(4) CE_fst wfI_wfCE_eval_e[of P B \ ce "B_pair b b2" i,OF b] wfRCV_elims(2)[of P s b b2] + by (metis eval_e_uniqueness) + then show ?case using s eval_e.intros by metis +next + case (CE_snd ce) + then obtain b1 where b:"P ; B ; \ \\<^sub>w\<^sub>f ce : B_pair b1 b" using wfCE_elims by metis + then obtain s where s:"i \ ce \ ~ s" using CE_snd by auto + then obtain s1 and s2 where "s = (SPair s1 s2)" + using eval_e_elims(5) CE_snd wfI_wfCE_eval_e[of P B \ ce "B_pair b1 b" i,OF b] wfRCV_elims(2)[of P s b b1] + eval_e_uniqueness + by (metis wfRCV_elims(2)) + then show ?case using s eval_e.intros by metis +next + case (CE_len v1) + then obtain bv1 where "eval_e i v1 (SBitvec bv1)" + using wfV_eval_bitvec CE_len wfCE_elims eval_e_uniqueness + by (metis eval_e_elims(7) wfCE_concatI wfI_wfCE_eval_e) + then show ?case using eval_e.intros by metis +qed + +lemma eval_c_exist: + fixes \::\ and c::c + assumes "P ; \ \ i" and "P ; B ; \ \\<^sub>w\<^sub>f c" + shows "\s. i \ c \ ~ s" + using assms proof(nominal_induct c rule: c.strong_induct) + case C_true + then show ?case using eval_c.intros wfC_elims by metis +next + case C_false + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_conj c1 c2) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_disj x1 x2) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_not x) + then show ?case using eval_c.intros wfC_elims by metis +next + case (C_imp x1 x2) + then show ?case using eval_c.intros eval_e_exist wfC_elims by metis +next + case (C_eq x1 x2) + then show ?case using eval_c.intros eval_e_exist wfC_elims by metis +qed + +lemma eval_c_uniqueness: + fixes c::c + assumes "i \ c \ ~ s" and "i \ c \ ~ s'" + shows "s=s'" + using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct) + case C_true + then show ?case using eval_c_elims by metis +next + case C_false + then show ?case using eval_c_elims by metis +next + case (C_conj x1 x2) + then show ?case using eval_c_elims(3) by (metis (full_types)) +next + case (C_disj x1 x2) + then show ?case using eval_c_elims(4) by (metis (full_types)) +next + case (C_not x) + then show ?case using eval_c_elims(6) by (metis (full_types)) +next + case (C_imp x1 x2) + then show ?case using eval_c_elims(5) by (metis (full_types)) +next + case (C_eq x1 x2) + then show ?case using eval_e_uniqueness eval_c_elims(7) by metis +qed + +lemma wfI_wfC_eval_c: + fixes c::c + assumes "wfC P B G c" and "P ; G \ i" + shows "\s. i \ c \ ~ s" + using assms proof(nominal_induct c rule: c.strong_induct) +qed(metis wfC_elims wfI_wfCE_eval_e eval_c.intros)+ + +section \Satisfiability\ + +lemma satis_reflI: + fixes c::c + assumes "i \ ((x, b, c) #\<^sub>\ G)" + shows "i \ c" + using assms by auto + +lemma is_satis_mp: + fixes c1::c and c2::c + assumes "i \ (c1 IMP c2)" and "i \ c1" + shows "i \ c2" + using assms proof - + have "eval_c i (c1 IMP c2) True" using is_satis.simps using assms by blast + then obtain b1 and b2 where "True = (b1 \ b2) \ eval_c i c1 b1 \ eval_c i c2 b2" + using eval_c_elims(5) by metis + moreover have "eval_c i c1 True" using is_satis.simps using assms by blast + moreover have "b1 = True" using calculation eval_c_uniqueness by blast + ultimately have "eval_c i c2 True" by auto + thus ?thesis using is_satis.intros by auto +qed + +lemma is_satis_imp: + fixes c1::c and c2::c + assumes "i \ c1 \ i \ c2" and "i \ c1 \ ~ b1" and "i \ c2 \ ~ b2" + shows "i \ (c1 IMP c2)" +proof(cases b1) + case True + hence "i \ c2" using assms is_satis.simps by simp + hence "b2 = True" using is_satis.simps assms + using eval_c_uniqueness by blast + then show ?thesis using eval_c_impI is_satis.simps assms by force +next + case False + then show ?thesis using assms eval_c_impI is_satis.simps by metis +qed + +lemma is_satis_iff: + "i \ G = (\x b c. (x,b,c) \ toSet G \ i \ c)" + by(induct G,auto) + +lemma is_satis_g_append: + "i \ (G1@G2) = (i \ G1 \ i \ G2)" + using is_satis_g.simps is_satis_iff by auto + +section \Substitution for Evaluation\ + +lemma eval_v_i_upd: + fixes v::v + assumes "atom x \ v" and "i \ v \ ~ s'" + shows "eval_v ((i ( x \s))) v s' " + using assms proof(nominal_induct v arbitrary: s' rule:v.strong_induct) + case (V_lit x) + then show ?case by (metis eval_v_elims(1) eval_v_litI) +next + case (V_var y) + then obtain s where *: "Some s = i y \ s=s'" using eval_v_elims by metis + moreover have "x \ y" using \atom x \ V_var y\ v.supp by simp + ultimately have "(i ( x \s)) y = Some s" + by (simp add: \Some s = i y \ s = s'\) + then show ?case using eval_v_varI * \x \ y\ + by (simp add: eval_v.eval_v_varI) +next + case (V_pair v1 v2) + hence "atom x \ v1 \ atom x \ v2" using v.supp by simp + moreover obtain s1 and s2 where *: "eval_v i v1 s1 \ eval_v i v2 s2 \ s' = SPair s1 s2" using eval_v_elims V_pair by metis + ultimately have "eval_v ((i ( x \s))) v1 s1 \ eval_v ((i ( x \s))) v2 s2" using V_pair by blast + thus ?case using eval_v_pairI * by meson +next + case (V_cons tyid dc v1) + hence "atom x \ v1" using v.supp by simp + moreover obtain s1 where *: "eval_v i v1 s1 \ s' = SCons tyid dc s1" using eval_v_elims V_cons by metis + ultimately have "eval_v ((i ( x \s))) v1 s1" using V_cons by blast + thus ?case using eval_v_consI * by meson +next + case (V_consp tyid dc b1 v1) + + hence "atom x \ v1" using v.supp by simp + moreover obtain s1 where *: "eval_v i v1 s1 \ s' = SConsp tyid dc b1 s1" using eval_v_elims V_consp by metis + ultimately have "eval_v ((i ( x \s))) v1 s1" using V_consp by blast + thus ?case using eval_v_conspI * by meson +qed + +lemma eval_e_i_upd: + fixes e::ce + assumes "i \ e \ ~ s'" and "atom x \ e" + shows " (i ( x \s)) \ e \ ~ s'" + using assms apply(induct rule: eval_e.induct) using eval_v_i_upd eval_e_elims + by (meson ce.fresh eval_e.intros)+ + +lemma eval_c_i_upd: + fixes c::c + assumes "i \ c \ ~ s'" and "atom x \ c" + shows "((i ( x \s))) \ c \ ~ s' " + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using RCLogic.eval_c_eqI eval_e_i_upd c.fresh by metis +qed(simp add: eval_c.intros)+ + +lemma subst_v_eval_v[simp]: + fixes v::v and v'::v + assumes "i \ v \ ~ s" and "i \ (v'[x::=v]\<^sub>v\<^sub>v) \ ~ s'" + shows "(i ( x \ s )) \ v' \ ~ s'" + using assms proof(nominal_induct v' arbitrary: s' rule:v.strong_induct) + case (V_lit x) + then show ?case using subst_vv.simps + by (metis eval_v_elims(1) eval_v_litI) +next + case (V_var x') + then show ?case proof(cases "x=x'") + case True + hence "(V_var x')[x::=v]\<^sub>v\<^sub>v = v" using subst_vv.simps by auto + then show ?thesis using V_var eval_v_elims eval_v_varI eval_v_uniqueness True + by (simp add: eval_v.eval_v_varI) + next + case False + hence "atom x \ (V_var x')" by simp + then show ?thesis using eval_v_i_upd False V_var by fastforce + qed +next + case (V_pair v1 v2) + then obtain s1 and s2 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \ eval_v i (v2[x::=v]\<^sub>v\<^sub>v) s2 \ s' = SPair s1 s2" using V_pair eval_v_elims subst_vv.simps by metis + hence "(i ( x \ s )) \ v1 \ ~ s1 \ (i ( x \ s )) \ v2 \ ~ s2" using V_pair by metis + thus ?case using eval_v_pairI subst_vv.simps * V_pair by metis +next + case (V_cons tyid dc v1) + then obtain s1 where "eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1" using eval_v_elims subst_vv.simps by metis + thus ?case using eval_v_consI V_cons + by (metis eval_v_elims subst_vv.simps) +next + case (V_consp tyid dc b1 v1) + + then obtain s1 where *:"eval_v i (v1[x::=v]\<^sub>v\<^sub>v) s1 \ s' = SConsp tyid dc b1 s1" using eval_v_elims subst_vv.simps by metis + hence "i ( x \ s ) \ v1 \ ~ s1" using V_consp by metis + thus ?case using * eval_v_conspI by metis +qed + +lemma subst_e_eval_v[simp]: + fixes y::x and e::ce and v::v and e'::ce + assumes "i \ e' \ ~ s'" and "e'=(e[y::=v]\<^sub>c\<^sub>e\<^sub>v)" and "i \ v \ ~ s" + shows "(i ( y \ s )) \ e \ ~ s'" + using assms proof(induct arbitrary: e rule: eval_e.induct) + case (eval_e_valI i v1 sv) + then obtain v1' where *:"e = CE_val v1' \ v1 = v1'[y::=v]\<^sub>v\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_v i (v1'[y::=v]\<^sub>v\<^sub>v) sv" using eval_e_valI by simp + hence "eval_v (i ( y \ s )) v1' sv" using subst_v_eval_v eval_e_valI by simp + then show ?case using RCLogic.eval_e_valI * by meson +next + case (eval_e_plusI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op Plus v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_plusI by simp + hence "eval_e (i ( y \ s )) v1' (SNum n1) \ eval_e (i ( y \ s )) v2' (SNum n2)" using subst_v_eval_v eval_e_plusI + using "local.*" by blast + then show ?case using RCLogic.eval_e_plusI * by meson +next + case (eval_e_leqI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op LEq v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SNum n2)" using eval_e_leqI by simp + hence "eval_e (i ( y \ s )) v1' (SNum n1) \ eval_e (i ( y \ s )) v2' (SNum n2)" using subst_v_eval_v eval_e_leqI + using * by blast + then show ?case using RCLogic.eval_e_leqI * by meson +next + case (eval_e_eqI i v1 n1 v2 n2) + then obtain v1' and v2' where *:"e = CE_op Eq v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n1 \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) n2" using eval_e_eqI by simp + hence "eval_e (i ( y \ s )) v1' n1 \ eval_e (i ( y \ s )) v2' n2" using subst_v_eval_v eval_e_eqI + using * by blast + then show ?case using RCLogic.eval_e_eqI * by meson +next + case (eval_e_fstI i v1 s1 s2) + then obtain v1' and v2' where *:"e = CE_fst v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_fstI by simp + hence "eval_e (i ( y \ s )) v1' (SPair s1 s2)" using eval_e_fstI * by metis + then show ?case using RCLogic.eval_e_fstI * by meson +next + case (eval_e_sndI i v1 s1 s2) + then obtain v1' and v2' where *:"e = CE_snd v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SPair s1 s2)" using eval_e_sndI by simp + hence "eval_e (i ( y \ s )) v1' (SPair s1 s2)" using subst_v_eval_v eval_e_sndI * by blast + then show ?case using RCLogic.eval_e_sndI * by meson +next + case (eval_e_concatI i v1 bv1 v2 bv2) + then obtain v1' and v2' where *:"e = CE_concat v1' v2' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v \ v2 = v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv1) \ eval_e i (v2'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv2)" using eval_e_concatI by simp + moreover hence "eval_e (i ( y \ s )) v1' (SBitvec bv1) \ eval_e (i ( y \ s )) v2' (SBitvec bv2)" + using subst_v_eval_v eval_e_concatI * by blast + ultimately show ?case using RCLogic.eval_e_concatI * eval_v_uniqueness by (metis eval_e_concatI.hyps(1)) +next + case (eval_e_lenI i v1 bv) + then obtain v1' where *:"e = CE_len v1' \ v1 = v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v" + using assms by(nominal_induct e rule:ce.strong_induct,simp+) + hence "eval_e i (v1'[y::=v]\<^sub>c\<^sub>e\<^sub>v) (SBitvec bv)" using eval_e_lenI by simp + hence "eval_e (i ( y \ s )) v1' (SBitvec bv)" using subst_v_eval_v eval_e_lenI * by blast + then show ?case using RCLogic.eval_e_lenI * by meson +qed + +lemma subst_c_eval_v[simp]: + fixes v::v and c :: c + assumes "i \ v \ ~ s" and "i \ c[x::=v]\<^sub>c\<^sub>v \ ~ s1" and + "(i ( x \ s)) \ c \ ~ s2" + shows "s1 = s2" + using assms proof(nominal_induct c arbitrary: s1 s2 rule: c.strong_induct) + case C_true + hence "s1 = True \ s2 = True" using eval_c_elims subst_cv.simps by auto + then show ?case by auto +next + case C_false + hence "s1 = False \ s2 = False" using eval_c_elims subst_cv.simps by metis + then show ?case by auto +next + case (C_conj c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v AND c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(3) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(3) C_conj by metis + ultimately show ?case using C_conj by (meson eval_c_elims) +next + case (C_disj c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v OR c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(4) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(4) C_disj by metis + ultimately show ?case using C_disj by (meson eval_c_elims) +next + case (C_not c1) + then obtain s11 where "(s1 = (\ s11)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11" using + eval_c_elims(6) by (metis subst_cv.simps(7)) + moreover obtain s21 where "eval_c (i ( x \ s)) c1 s21 \ (s2 = (\s21))" using + eval_c_elims(6) C_not by metis + ultimately show ?case using C_not by (meson eval_c_elims) +next + case (C_imp c1 c2) + hence *:"eval_c i (c1[x::=v]\<^sub>c\<^sub>v IMP c2[x::=v]\<^sub>c\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 \ s12)) \ eval_c i c1[x::=v]\<^sub>c\<^sub>v s11 \ eval_c i c2[x::=v]\<^sub>c\<^sub>v s12" using + eval_c_elims(5) by metis + moreover obtain s21 and s22 where "eval_c (i ( x \ s)) c1 s21 \ eval_c (i ( x \ s)) c2 s22 \ (s2 = (s21 \ s22))" using + eval_c_elims(5) C_imp by metis + ultimately show ?case using C_imp by (meson eval_c_elims) +next + case (C_eq e1 e2) + hence *:"eval_c i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v == e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s1" using subst_cv.simps by auto + then obtain s11 and s12 where "(s1 = (s11 = s12)) \ eval_e i (e1[x::=v]\<^sub>c\<^sub>e\<^sub>v) s11 \ eval_e i (e2[x::=v]\<^sub>c\<^sub>e\<^sub>v) s12" using + eval_c_elims(7) by metis + moreover obtain s21 and s22 where "eval_e (i ( x \ s)) e1 s21 \ eval_e (i ( x \ s)) e2 s22 \ (s2 = (s21 = s22))" using + eval_c_elims(7) C_eq by metis + ultimately show ?case using C_eq subst_e_eval_v by (metis eval_e_uniqueness) +qed + +lemma wfI_upd: + assumes "wfI \ \ i" and "wfRCV \ s b" and "wfG \ B ((x, b, c) #\<^sub>\ \)" + shows "wfI \ ((x, b, c) #\<^sub>\ \) (i(x \ s))" +proof(subst wfI_def,rule) + fix xa + assume as:"xa \ toSet ((x, b, c) #\<^sub>\ \)" + + then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps + using prod_cases3 by blast + + have "\sa. Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" proof(cases "x=x1") + case True + hence "b=b1" using as xa wfG_unique assms by metis + hence "Some s = (i(x \ s)) x1 \ wfRCV \ s b1" using assms True by simp + then show ?thesis by auto + next + case False + hence "(x1,b1,c1) \ toSet \" using xa as by auto + then obtain sa where "Some sa = i x1 \ wfRCV \ sa b1" using assms wfI_def as xa by auto + hence "Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" using False by auto + then show ?thesis by auto + qed + + thus "case xa of (xa, ba, ca) \ \sa. Some sa = (i(x \ s)) xa \ wfRCV \ sa ba" using xa by auto +qed + +lemma wfI_upd_full: + fixes v::v + assumes "wfI \ G i" and "G = ((\'[x::=v]\<^sub>\\<^sub>v)@\)" and "wfRCV \ s b" and "wfG \ B (\'@((x,b,c)#\<^sub>\\))" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "wfI \ (\'@((x,b,c)#\<^sub>\\)) (i(x \ s))" +proof(subst wfI_def,rule) + fix xa + assume as:"xa \ toSet (\'@((x,b,c)#\<^sub>\\))" + + then obtain x1::x and b1::b and c1::c where xa: "xa = (x1,b1,c1)" using toSet.simps + using prod_cases3 by blast + + have "\sa. Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" + proof(cases "x=x1") + case True + hence "b=b1" using as xa wfG_unique_full assms by metis + hence "Some s = (i(x \ s)) x1 \ wfRCV \ s b1" using assms True by simp + then show ?thesis by auto + next + case False + hence "(x1,b1,c1) \ toSet (\'@\)" using as xa by auto + then obtain c1' where "(x1,b1,c1') \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using xa as wfG_member_subst assms False by metis + then obtain sa where "Some sa = i x1 \ wfRCV \ sa b1" using assms wfI_def as xa by blast + hence "Some sa = (i(x \ s)) x1 \ wfRCV \ sa b1" using False by auto + then show ?thesis by auto + qed + + thus "case xa of (xa, ba, ca) \ \sa. Some sa = (i(x \ s)) xa \ wfRCV \ sa ba" using xa by auto +qed + +lemma subst_c_satis[simp]: + fixes v::v + assumes "i \ v \ ~ s" and "wfC \ B ((x,b,c')#\<^sub>\\) c" and "wfI \ \ i" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "i \ (c[x::=v]\<^sub>c\<^sub>v) \ (i ( x \ s)) \ c" +proof - + have "wfI \ ((x, b, c') #\<^sub>\ \) (i(x \ s))" using wfI_upd assms wfC_wf eval_v_base by blast + then obtain s1 where s1:"eval_c (i(x \ s)) c s1" using eval_c_exist[of \ "((x,b,c')#\<^sub>\\)" "(i ( x \ s))" B c ] assms by auto + + have "\ ; B ; \ \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wf_subst1(2)[OF assms(2) _ assms(4) , of GNil x ] subst_gv.simps by simp + then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \ "\" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto + + show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases + using eval_c_uniqueness is_satis.simps by auto +qed + +text \ Key theorem telling us we can replace a substitution with an update to the valuation \ +lemma subst_c_satis_full: + fixes v::v and \'::\ + assumes "i \ v \ ~ s" and "wfC \ B (\'@((x,b,c')#\<^sub>\\)) c" and "wfI \ ((\'[x::=v]\<^sub>\\<^sub>v)@\) i" and "\ ; B ; \ \\<^sub>w\<^sub>f v : b" + shows "i \ (c[x::=v]\<^sub>c\<^sub>v) \ (i ( x \ s)) \ c" +proof - + have "wfI \ (\'@((x, b, c')) #\<^sub>\ \) (i(x \ s))" using wfI_upd_full assms wfC_wf eval_v_base wfI_suffix wfI_def wfV_wf by fast + then obtain s1 where s1:"eval_c (i(x \ s)) c s1" using eval_c_exist[of \ "(\'@(x,b,c')#\<^sub>\\)" "(i ( x \ s))" B c ] assms by auto + + have "\ ; B ; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" using wbc_subst assms by auto + + then obtain s2 where s2:"eval_c i c[x::=v]\<^sub>c\<^sub>v s2" using eval_c_exist[of \ "((\'[x::=v]\<^sub>\\<^sub>v)@\)" "i" B "c[x::=v]\<^sub>c\<^sub>v" ] assms by auto + + show ?thesis using s1 s2 subst_c_eval_v[OF assms(1) s2 s1] is_satis.cases + using eval_c_uniqueness is_satis.simps by auto +qed + +section \Validity\ + +lemma validI[intro]: + fixes c::c + assumes "wfC P B G c" and "\i. P ; G \ i \ i \ G \ i \ c" + shows "P ; B ; G \ c" + using assms valid.simps by presburger + +lemma valid_g_wf: + fixes c::c and G::\ + assumes "P ; B ; G \ c" + shows "P ; B \\<^sub>w\<^sub>f G" + using assms wfC_wf valid.simps by blast + +lemma valid_reflI [intro]: + fixes b::b + assumes "P ; B ; ((x,b,c1)#\<^sub>\G) \\<^sub>w\<^sub>f c1" and "c1 = c2" + shows "P ; B ; ((x,b,c1)#\<^sub>\G) \ c2" + using satis_reflI assms by simp + +subsection \Weakening and Strengthening\ + +text \ Adding to the domain of a valuation doesn't change the result \ + +lemma eval_v_weakening: + fixes c::v and B::"bv fset" + assumes "i = i'|` d" and "supp c \ atom ` d \ supp B " and "i \ c \ ~ s" + shows "i' \ c \ ~ s" + using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_v_litI by metis +next + case (V_var x) + have "atom x \ atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base + proof - + show ?thesis + by (metis UnE V_var.prems(2) \atom x \ supp B\ singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *) + qed + moreover have "Some s = i x" using assms eval_v_elims(2) + using V_var.prems(3) by blast + hence "Some s= i' x" using assms insert_subset restrict_in + proof - + show ?thesis + by (metis (no_types) \i = i' |` d\ \Some s = i x\ atom_eq_iff calculation imageE restrict_in) (* 31 ms *) + qed + thus ?case using eval_v.eval_v_varI by simp + +next + case (V_pair v1 v2) + then show ?case using eval_v_elims(3) eval_v_pairI v.supp + by (metis assms le_sup_iff) +next + case (V_cons dc v1) + then show ?case using eval_v_elims(4) eval_v_consI v.supp + by (metis assms le_sup_iff) +next + case (V_consp tyid dc b1 v1) + + then obtain sv1 where *:"i \ v1 \ ~ sv1 \ s = SConsp tyid dc b1 sv1" using eval_v_elims by metis + hence "i' \ v1 \ ~ sv1" using V_consp by auto + then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis +qed + +lemma eval_v_restrict: + fixes c::v and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B " and "i' \ c \ ~ s" + shows "i \ c \ ~ s" + using assms proof(nominal_induct c arbitrary:s rule: v.strong_induct) + case (V_lit x) + then show ?case using eval_v_elims eval_v_litI by metis +next + case (V_var x) + have "atom x \ atom ` d" using x_not_in_b_set[of x B] assms v.supp(2) supp_at_base + proof - + show ?thesis + by (metis UnE V_var.prems(2) \atom x \ supp B\ singletonI subset_iff supp_at_base v.supp(2)) (* 46 ms *) + qed + moreover have "Some s = i' x" using assms eval_v_elims(2) + using V_var.prems(3) by blast + hence "Some s= i x" using assms insert_subset restrict_in + proof - + show ?thesis + by (metis (no_types) \i = i' |` d\ \Some s = i' x\ atom_eq_iff calculation imageE restrict_in) (* 31 ms *) + qed + thus ?case using eval_v.eval_v_varI by simp +next + case (V_pair v1 v2) + then show ?case using eval_v_elims(3) eval_v_pairI v.supp + by (metis assms le_sup_iff) +next + case (V_cons dc v1) + then show ?case using eval_v_elims(4) eval_v_consI v.supp + by (metis assms le_sup_iff) +next + case (V_consp tyid dc b1 v1) + then obtain sv1 where *:"i' \ v1 \ ~ sv1 \ s = SConsp tyid dc b1 sv1" using eval_v_elims by metis + hence "i \ v1 \ ~ sv1" using V_consp by auto + then show ?case using * eval_v_conspI v.supp eval_v.simps assms le_sup_iff by metis +qed + +lemma eval_e_weakening: + fixes e::ce and B::"bv fset" + assumes "i \ e \ ~ s" and "i = i' |` d" and "supp e \ atom ` d \ supp B " + shows "i' \ e \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_fstI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by metis +next + case (eval_e_concatI i v1 bv2 v2 bv1) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +next + case (eval_e_lenI i v bv) + then show ?case using ce.supp eval_e.intros + using eval_v_weakening by auto +qed + +lemma eval_e_restrict : + fixes e::ce and B::"bv fset" + assumes "i' \ e \ ~ s" and "i = i' |` d" and "supp e \ atom ` d \ supp B " + shows "i \ e \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_fstI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by metis +next + case (eval_e_concatI i v1 bv2 v2 bv1) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +next + case (eval_e_lenI i v bv) + then show ?case using ce.supp eval_e.intros + using eval_v_restrict by auto +qed + +lemma eval_c_i_weakening: + fixes c::c and B::"bv fset" + assumes "i \ c \ ~ s" and "i = i' |` d" and "supp c \ atom ` d \ supp B" + shows "i' \ c \ ~ s" + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_c.intros eval_e_weakening by auto +qed(auto simp add: eval_c.intros)+ + +lemma eval_c_i_restrict: + fixes c::c and B::"bv fset" + assumes "i' \ c \ ~ s" and "i = i' |` d" and "supp c \ atom ` d \ supp B" + shows "i \ c \ ~ s" + using assms proof(induct rule:eval_c.induct) + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_c.intros eval_e_restrict by auto +qed(auto simp add: eval_c.intros)+ + +lemma is_satis_i_weakening: + fixes c::c and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B " and "i \ c" + shows "i' \ c" + using is_satis.simps eval_c_i_weakening[OF _ assms(1) assms(2) ] + using assms(3) by auto + +lemma is_satis_i_restrict: + fixes c::c and B::"bv fset" + assumes "i = i' |` d" and "supp c \ atom ` d \ supp B" and "i' \ c" + shows "i \ c" + using is_satis.simps eval_c_i_restrict[OF _ assms(1) assms(2) ] + using assms(3) by auto + +lemma is_satis_g_restrict1: + fixes \'::\ and \::\ + assumes "toSet \ \ toSet \'" and "i \ \'" + shows "i \ \" + using assms proof(induct \ rule: \.induct) + case GNil + then show ?case by auto +next + case (GCons xbc G) + obtain x and b and c::c where xbc: "xbc=(x,b,c)" + using prod_cases3 by blast + hence "i \ G" using GCons by auto + moreover have "i \ c" using GCons + is_satis_iff toSet.simps subset_iff + using xbc by blast + ultimately show ?case using is_satis_g.simps GCons + by (simp add: xbc) +qed + +lemma is_satis_g_restrict2: + fixes \'::\ and \::\ + assumes "i \ \" and "i' = i |` d" and "atom_dom \ \ atom ` d" and "\ ; B \\<^sub>w\<^sub>f \" + shows "i' \ \" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x b c G) + + hence "i' \ G" proof - + have "i \ G" using GCons by simp + moreover have "atom_dom G \ atom ` d" using GCons by simp + ultimately show ?thesis using GCons wfG_cons2 by blast + qed + + moreover have "i' \ c" proof - + have "i \ c" using GCons by auto + moreover have "\ ; B ; (x, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c" using wfG_wfC GCons by simp + moreover hence "supp c \ atom ` d \ supp B" using wfC_supp GCons atom_dom_eq by blast + ultimately show ?thesis using is_satis_i_restrict[of i' i d c] GCons by simp + qed + + ultimately show ?case by auto +qed + +lemma is_satis_g_restrict: + fixes \'::\ and \::\ + assumes "toSet \ \ toSet \'" and "i' \ \'" and "i = i' |` (fst ` toSet \)" and "\ ; B \\<^sub>w\<^sub>f \" + shows "i \ \" + using assms is_satis_g_restrict1[OF assms(1) assms(2)] is_satis_g_restrict2[OF _ assms(3)] by simp + +subsection \Updating valuation\ + +lemma is_satis_c_i_upd: + fixes c::c + assumes "atom x \ c" and "i \ c" + shows "((i ( x \s))) \ c" + using assms eval_c_i_upd is_satis.simps by simp + +lemma is_satis_g_i_upd: + fixes G::\ + assumes "atom x \ G" and "i \ G" + shows "((i ( x \s))) \ G" + using assms proof(induct G rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' G') + + hence *:"atom x \ G' \ atom x \ c'" + using fresh_def fresh_GCons GCons by force + moreover hence "is_satis ((i ( x \s))) c'" + using is_satis_c_i_upd GCons is_satis_g.simps by auto + moreover have " is_satis_g (i(x \ s)) G'" using GCons * by fastforce + ultimately show ?case + using GCons is_satis_g.simps(2) by metis +qed + +lemma valid_weakening: + assumes "\ ; B ; \ \ c" and "toSet \ \ toSet \'" and "wfG \ B \'" + shows "\ ; B ; \' \ c" +proof - + have "wfC \ B \ c" using assms valid.simps by auto + hence sp: "supp c \ atom `(fst `toSet \) \ supp B" using wfX_wfY wfG_elims + using atom_dom.simps dom.simps wf_supp(2) by metis + have wfg: "wfG \ B \" using assms valid.simps wfC_wf by auto + + moreover have a1: "(\i. wfI \ \' i \ i \ \' \ i \ c)" proof(rule allI, rule impI) + fix i + assume as: "wfI \ \' i \ i \ \'" + hence as1: "fst ` toSet \ \ dom i" using assms wfI_domi by blast + obtain i' where idash: "i' = restrict_map i (fst `toSet \)" by blast + hence as2: "dom i' = (fst `toSet \)" using dom_restrict as1 by auto + + have id2: "\ ; \ \ i' \ i' \ \" proof - + have "wfI \ \ i'" using as2 wfI_restrict_weakening[of \ \' i i' \] as assms + using idash by blast + moreover have "i' \ \" using is_satis_g_restrict[OF assms(2)] wfg as idash by auto + ultimately show ?thesis using idash by auto + qed + hence "i' \ c" using assms valid.simps by auto + thus "i \ c" using assms valid.simps is_satis_i_weakening idash sp by blast + qed + moreover have "wfC \ B \' c" using wf_weakening assms valid.simps + by (meson wfg) + ultimately show ?thesis using assms valid.simps by auto +qed + +lemma is_satis_g_suffix: + fixes G::\ + assumes "i \ (G'@G)" + shows "i \ G" + using assms proof(induct G' rule:\.induct) + case GNil + then show ?case by auto +next + case (GCons xbc x2) + obtain x and b and c::c where xbc: "xbc=(x,b,c)" + using prod_cases3 by blast + hence " i \ (xbc #\<^sub>\ (x2 @ G))" using append_g.simps GCons by fastforce + then show ?case using is_satis_g.simps GCons xbc by blast +qed + +lemma wfG_inside_valid2: + fixes x::x and \::\ and c0::c and c0'::c + assumes "wfG \ B (\'@((x,b0,c0')#\<^sub>\\))" and + "\ ; B ; \'@(x,b0,c0)#\<^sub>\\ \ c0'" + shows "wfG \ B (\'@((x,b0,c0)#\<^sub>\\))" +proof - + have "wfG \ B (\'@(x,b0,c0)#\<^sub>\\)" using valid.simps wfC_wf assms by auto + thus ?thesis using wfG_replace_inside_full assms by auto +qed + +lemma valid_trans: + assumes " \ ; \ ; \ \ c0[z::=v]\<^sub>v" and " \ ; \ ; (z,b,c0)#\<^sub>\\ \ c1" and "atom z \ \" and "wfV \ \ \ v b" + shows " \ ; \ ; \ \ c1[z::=v]\<^sub>v" +proof - + have *:"wfC \ \ ((z,b,c0)#\<^sub>\\) c1" using valid.simps assms by auto + hence "wfC \ \ \ (c1[z::=v]\<^sub>v)" using wf_subst1(2)[OF * , of GNil ] assms subst_gv.simps subst_v_c_def by force + + moreover have "\i. wfI \ \ i \ is_satis_g i \ \ is_satis i (c1[z::=v]\<^sub>v)" + proof(rule,rule) + fix i + assume as: "wfI \ \ i \ is_satis_g i \" + then obtain sv where sv: "eval_v i v sv \ wfRCV \ sv b" using eval_v_exist assms by metis + hence "is_satis i (c0[z::=v]\<^sub>v)" using assms valid.simps as by metis + hence "is_satis (i(z \ sv)) c0" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis + moreover have "is_satis_g (i(z \ sv)) \" + using is_satis_g_i_upd assms by (simp add: as) + ultimately have "is_satis_g (i(z \ sv)) ((z,b,c0)#\<^sub>\\)" + using is_satis_g.simps by simp + moreover have "wfI \ ((z,b,c0)#\<^sub>\\) (i(z \ sv))" using as wfI_upd sv assms valid.simps wfC_wf by metis + ultimately have "is_satis (i(z \ sv)) c1" using assms valid.simps by auto + thus "is_satis i (c1[z::=v]\<^sub>v)" using subst_c_satis sv as assms valid.simps wfC_wf wfG_elim2 subst_v_c_def by metis + qed + + ultimately show ?thesis using valid.simps by auto +qed + +lemma valid_trans_full: + assumes "\ ; \ ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) \ c2[z2::=V_var x]\<^sub>v" and + "\ ; \ ; ((x, b, c2[z2::=V_var x]\<^sub>v) #\<^sub>\ \) \ c3[z3::=V_var x]\<^sub>v" + shows "\ ; \ ; ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) \ c3[z3::=V_var x]\<^sub>v" + unfolding valid.simps proof + show "\ ; \ ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c3[z3::=V_var x]\<^sub>v" using wf_trans valid.simps assms by metis + + show "\i. ( wfI \ ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \) i \ (is_satis_g i ((x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \)) \ (is_satis i (c3[z3::=V_var x]\<^sub>v)) ) " + proof(rule,rule) + fix i + assume as: "\ ; (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ i \ i \ (x, b, c1[z1::=V_var x]\<^sub>v) #\<^sub>\ \" + have "i \ c2[z2::=V_var x]\<^sub>v" using is_satis_g.simps as assms by simp + moreover have "i \ \" using is_satis_g.simps as by simp + ultimately show "i \ c3[z3::=V_var x]\<^sub>v " using assms is_satis_g.simps valid.simps + by (metis append_g.simps(1) as wfI_replace_inside) + qed +qed + +lemma eval_v_weakening_x: + fixes c::v + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_v.induct) + case (eval_v_litI i l) + then show ?case using eval_v.intros by auto +next + case (eval_v_varI sv i1 x1) + hence "x \ x1" using v.fresh fresh_at_base by auto + hence "i x1 = Some sv" using eval_v_varI by simp + then show ?case using eval_v.intros by auto +next + case (eval_v_pairI i v1 s1 v2 s2) + then show ?case using eval_v.intros by auto +next + case (eval_v_consI i v s tyid dc) + then show ?case using eval_v.intros by auto +next + case (eval_v_conspI i v s tyid dc b) + then show ?case using eval_v.intros by auto +qed + +lemma eval_e_weakening_x: + fixes c::ce + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_e.induct) + case (eval_e_valI i v sv) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_plusI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_leqI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_eqI i v1 n1 v2 n2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_fstI i v v1 v2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_sndI i v v1 v2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_concatI i v1 bv1 v2 bv2) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +next + case (eval_e_lenI i v bv) + then show ?case using eval_v_weakening_x eval_e.intros ce.fresh by metis +qed + +lemma eval_c_weakening_x: + fixes c::c + assumes "i' \ c \ ~ s" and "atom x \ c" and "i = i' (x \ s')" + shows "i \ c \ ~ s" + using assms proof(induct rule: eval_c.induct) + case (eval_c_trueI i) + then show ?case using eval_c.intros by auto +next + case (eval_c_falseI i) + then show ?case using eval_c.intros by auto +next + case (eval_c_conjI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_disjI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_impI i c1 b1 c2 b2) + then show ?case using eval_c.intros by auto +next + case (eval_c_notI i c b) + then show ?case using eval_c.intros by auto +next + case (eval_c_eqI i e1 sv1 e2 sv2) + then show ?case using eval_e_weakening_x c.fresh eval_c.intros by metis +qed + +lemma is_satis_weakening_x: + fixes c::c + assumes "i' \ c" and "atom x \ c" and "i = i' (x \ s)" + shows "i \ c" + using eval_c_weakening_x assms is_satis.simps by simp + +lemma is_satis_g_weakening_x: + fixes G::\ + assumes "i' \ G" and "atom x \ G" and "i = i' (x \ s)" + shows "i \ G" + using assms proof(induct G rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + hence "atom x \ c'" using fresh_GCons fresh_prodN by simp + moreover hence "i \ c'" using is_satis_weakening_x is_satis_g.simps(2) GCons by metis + then show ?case using is_satis_g.simps(2)[of i x' b' c' \'] GCons fresh_GCons by simp +qed + +section \Base Type Substitution\ + +text \The idea of boxing is to take an smt val and its base type and at nodes in the smt val that correspond to type variables we +wrap them in an SUt smt val node. Another way of looking at it is that s' where the node for the base type variable is an 'any node'. +It is needed to prove subst\_b\_valid - the base-type variable substitution lemma for validity. + +The first @{text "rcl_val"} is the expanded form (has type with base-variables replaced with base-type terms) + ; the second is its corresponding form + +We only have one variable so we need to ensure that in all of the @{text "bs_boxed_BVarI"} cases, the s has the same + base type. + +For example is an SMT value is (SPair (SInt 1) (SBool true)) and it has sort (BPair (BVar x) BBool)[x::=BInt] then +the boxed version is SPair (SUt (SInt 1)) (SBool true) and is has sort (BPair (BVar x) BBool). We need to do this +so that we can obtain from a valuation i, that gives values like the first smt value, to a valuation i' that gives values like +the second. +\ + +inductive boxed_b :: "\ \ rcl_val \ b \ bv \ b \ rcl_val \ bool" ( " _ \ _ ~ _ [ _ ::= _ ] \ _ " [50,50] 50) where + boxed_b_BVar1I: "\ bv = bv' ; wfRCV P s b \ \ boxed_b P s (B_var bv') bv b (SUt s)" +| boxed_b_BVar2I: "\ bv \ bv'; wfRCV P s (B_var bv') \ \ boxed_b P s (B_var bv') bv b s" +| boxed_b_BIntI:"wfRCV P s B_int \ boxed_b P s B_int _ _ s" +| boxed_b_BBoolI:"wfRCV P s B_bool \ boxed_b P s B_bool _ _ s " +| boxed_b_BUnitI:"wfRCV P s B_unit \ boxed_b P s B_unit _ _ s" +| boxed_b_BPairI:"\ boxed_b P s1 b1 bv b s1' ; boxed_b P s2 b2 bv b s2' \ \ boxed_b P (SPair s1 s2) (B_pair b1 b2) bv b (SPair s1' s2')" + +| boxed_b_BConsI:"\ + AF_typedef tyid dclist \ set P; + (dc, \ x : b | c \) \ set dclist ; + boxed_b P s1 b bv b' s1' + \ \ + boxed_b P (SCons tyid dc s1) (B_id tyid) bv b' (SCons tyid dc s1')" + +| boxed_b_BConspI:"\ AF_typedef_poly tyid bva dclist \ set P; + atom bva \ (b1,bv,b',s1,s1'); + (dc, \ x : b | c \) \ set dclist ; + boxed_b P s1 (b[bva::=b1]\<^sub>b\<^sub>b) bv b' s1' + \ \ + boxed_b P (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1) (B_app tyid b1) bv b' (SConsp tyid dc b1 s1')" + +| boxed_b_Bbitvec: "wfRCV P s B_bitvec \ boxed_b P s B_bitvec bv b s" + +equivariance boxed_b +nominal_inductive boxed_b . + +inductive_cases boxed_b_elims: + "boxed_b P s (B_var bv) bv' b s'" + "boxed_b P s B_int bv b s'" + "boxed_b P s B_bool bv b s'" + "boxed_b P s B_unit bv b s'" + "boxed_b P s (B_pair b1 b2) bv b s'" + "boxed_b P s (B_id dc) bv b s'" + "boxed_b P s B_bitvec bv b s'" + "boxed_b P s (B_app dc b') bv b s'" + +lemma boxed_b_wfRCV: + assumes "boxed_b P s b bv b' s'" (*and "supp s = {}"*) and "\\<^sub>w\<^sub>f P" + shows "wfRCV P s b[bv::=b']\<^sub>b\<^sub>b \ wfRCV P s' b" + using assms proof(induct rule: boxed_b.inducts) + case (boxed_b_BVar1I bv bv' P s b ) + then show ?case using wfRCV.intros by auto +next + case (boxed_b_BVar2I bv bv' P s ) + then show ?case using wfRCV.intros by auto +next + case (boxed_b_BPairI P s1 b1 bv b s1' s2 b2 s2') + then show ?case using wfRCV.intros rcl_val.supp by simp +next + case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1') + hence "supp b = {}" using wfTh_supp_b by metis + hence "b [ bv ::= b' ]\<^sub>b\<^sub>b = b" using fresh_def subst_b_b_def forget_subst[of "bv" b b'] by auto + hence " P \ SCons tyid dc s1 : (B_id tyid)" using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis + moreover have "P \ SCons tyid dc s1' : B_id tyid" using boxed_b_BConsI + using wfRCV.intros rcl_val.supp subst_bb.simps boxed_b_BConsI by metis + ultimately show ?case using subst_bb.simps by metis +next + case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c) + + obtain bva2 and dclist2 where *:"AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ (bv,(P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b))" + using obtain_fresh_bv by metis + + then obtain x2 and b2 and c2 where **:\(dc, \ x2 : b2 | c2 \) \ set dclist2\ + using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff by metis + + have "P \ SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1 : (B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" proof + show 1: \AF_typedef_poly tyid bva2 dclist2 \ set P\ using boxed_b_BConspI * by auto + show 2: \(dc, \ x2 : b2 | c2 \) \ set dclist2\ using boxed_b_BConspI using ** by simp + + hence "atom bv \ b2" proof - + have "supp b2 \ { atom bva2 }" using wfTh_poly_supp_b 1 2 boxed_b_BConspI by auto + moreover have "bv \ bva2" using * fresh_Pair fresh_at_base by metis + ultimately show ?thesis using fresh_def by force + qed + moreover have "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff * 2 boxed_b_BConspI by metis + ultimately show \ P \ s1 : b2[bva2::=b1[bv::=b']\<^sub>b\<^sub>b]\<^sub>b\<^sub>b\ using boxed_b_BConspI subst_b_b_def subst_bb_commute by auto + show "atom bva2 \ (P, SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s1, B_app tyid b1[bv::=b']\<^sub>b\<^sub>b)" using * fresh_Pair by metis + qed + + moreover have "P \ SConsp tyid dc b1 s1' : B_app tyid b1" proof + show \AF_typedef_poly tyid bva dclist \ set P\ using boxed_b_BConspI by auto + show \(dc, \ x : b | c \) \ set dclist\ using boxed_b_BConspI by auto + show \ P \ s1' : b[bva::=b1]\<^sub>b\<^sub>b\ using boxed_b_BConspI by auto + have "atom bva \ P" using boxed_b_BConspI wfTh_fresh by metis + thus "atom bva \ (P, SConsp tyid dc b1 s1', B_app tyid b1)" using boxed_b_BConspI rcl_val.fresh b.fresh pure_fresh fresh_prodN by metis + qed + + ultimately show ?case using subst_bb.simps by simp +qed(auto)+ + +lemma subst_b_var: + assumes "B_var bv2 = b[bv::=b']\<^sub>b\<^sub>b" + shows "(b = B_var bv \ b' = B_var bv2) \ (b=B_var bv2 \ bv \ bv2)" + using assms by(nominal_induct b rule: b.strong_induct,auto+) + +text \Here the valuation i' is the conv wrap version of i. For every x in G, i' x is the conv wrap version of i x. +The next lemma for a clearer explanation of what this is. i produces values of sort b[bv::=b'] and i' produces values of sort b\ + +inductive boxed_i :: "\ \ \ \ b \ bv \ valuation \ valuation \ bool" ( " _ ; _ ; _ , _ \ _ \ _" [50,50] 50) where + boxed_i_GNilI: "\ ; GNil ; b , bv \ i \ i" +| boxed_i_GConsI: "\ Some s = i x; boxed_b \ s b bv b' s' ; \ ; \ ; b' , bv \ i \ i' \ \ \ ; ((x,b,c)#\<^sub>\\) ; b' , bv \ i \ (i'(x \ s'))" +equivariance boxed_i +nominal_inductive boxed_i . + +inductive_cases boxed_i_elims: + "\ ;GNil ; b , bv \ i \ i'" + "\ ; ((x,b,c)#\<^sub>\\) ; b' , bv \ i \ i'" + +lemma wfRCV_poly_elims: + fixes tm::"'a::fs" and b::b + assumes "T \ SConsp typid dc bdc s : b" + obtains bva dclist x1 b1 c1 where "b = B_app typid bdc \ + AF_typedef_poly typid bva dclist \ set T \ (dc, \ x1 : b1 | c1 \) \ set dclist \ T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b \ atom bva \ tm" + using assms proof(nominal_induct "SConsp typid dc bdc s" b avoiding: tm rule:wfRCV.strong_induct) + case (wfRCV_BConsPI bv dclist \ x b c) + then show ?case by simp +qed + +lemma boxed_b_ex: + assumes "wfRCV T s b[bv::=b']\<^sub>b\<^sub>b" and "wfTh T" + shows "\s'. boxed_b T s b bv b' s'" + using assms proof(nominal_induct s arbitrary: b rule: rcl_val.strong_induct) + case (SBitvec x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bitvec" using wfRCV_elims(9)[OF SBitvec(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SBitvec x : B_bitvec" using wfRCV.intros by simp + moreover hence "b' = B_bitvec" using True SBitvec subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_bitvec" using subst_bb_inject * by metis + then show ?thesis using * SBitvec boxed_b.intros by metis + qed +next + case (SNum x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_int" using wfRCV_elims(10)[OF SNum(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SNum x : B_int" using wfRCV.intros by simp + moreover hence "b' = B_int" using True SNum subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis + next + case False + hence "b = B_int" using subst_bb_inject(1) * by metis + then show ?thesis using * SNum boxed_b_BIntI by metis + qed +next + case (SBool x) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_bool" using wfRCV_elims(11)[OF SBool(1)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SBool x : B_bool" using wfRCV.intros by simp + moreover hence "b' = B_bool" using True SBool subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_bool" using subst_bb_inject * by metis + then show ?thesis using * SBool boxed_b.intros by metis + qed +next + case (SPair s1 s2) + then obtain b1 and b2 where *:"b[bv::=b']\<^sub>b\<^sub>b = B_pair b1 b2 \ wfRCV T s1 b1 \ wfRCV T s2 b2" using wfRCV_elims(12) by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SPair s1 s2 : B_pair b1 b2" using wfRCV.intros * by simp + moreover hence "b' = B_pair b1 b2" using True SPair subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I by metis + next + case False + then obtain b1' and b2' where "b = B_pair b1' b2' \ b1=b1'[bv::=b']\<^sub>b\<^sub>b \ b2=b2'[bv::=b']\<^sub>b\<^sub>b" using subst_bb_inject(5)[OF _ False ] * by metis + then show ?thesis using * SPair boxed_b_BPairI by blast + qed +next + case (SCons tyid dc s1) + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_id tyid" using wfRCV_elims(13)[OF SCons(2)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SCons tyid dc s1 : B_id tyid" using wfRCV.intros + using "local.*" SCons.prems by auto + moreover hence "b' = B_id tyid" using True SCons subst_bb.simps(1) * by simp + ultimately show ?thesis using boxed_b_BVar1I wfRCV.intros by metis + next + case False + then obtain b1' where beq: "b = B_id tyid" using subst_bb_inject * by metis + then obtain b2 dclist x c where **:"AF_typedef tyid dclist \ set T \ (dc, \ x : b2 | c \) \ set dclist \ wfRCV T s1 b2" using wfRCV_elims(13) * SCons by metis + then have "atom bv \ b2" using \wfTh T\ wfTh_lookup_supp_empty[of tyid dclist T dc "\ x : b2 | c \"] \.fresh fresh_def by auto + then have "b2 = b2[ bv ::= b']\<^sub>b\<^sub>b" using forget_subst subst_b_b_def by metis + then obtain s1' where s1:"T \ s1 ~ b2 [ bv ::= b' ] \ s1'" using SCons ** by metis + + have "T \ SCons tyid dc s1 ~ (B_id tyid) [ bv ::= b' ] \ SCons tyid dc s1'" proof(rule boxed_b_BConsI) + show "AF_typedef tyid dclist \ set T" using ** by auto + show "(dc, \ x : b2 | c \) \ set dclist" using ** by auto + show "T \ s1 ~ b2 [ bv ::= b' ] \ s1' " using s1 ** by auto + + qed + thus ?thesis using beq by metis + qed +next + case (SConsp typid dc bdc s) + + obtain bva dclist x1 b1 c1 where **:"b[bv::=b']\<^sub>b\<^sub>b = B_app typid bdc \ + AF_typedef_poly typid bva dclist \ set T \ (dc, \ x1 : b1 | c1 \) \ set dclist \ T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b \ atom bva \ bv" + using wfRCV_poly_elims[OF SConsp(2)] by metis + + then have *:"B_app typid bdc = b[bv::=b']\<^sub>b\<^sub>b" using wfRCV_elims(14)[OF SConsp(2)] by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SConsp typid dc bdc s : B_app typid bdc" using wfRCV.intros + using "local.*" SConsp.prems(1) by auto + moreover hence "b' = B_app typid bdc" using True SConsp subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + then obtain bdc' where bdc: "b = B_app typid bdc' \ bdc = bdc'[bv::=b']\<^sub>b\<^sub>b" using * subst_bb_inject(8)[OF *] by metis + (*hence beq:"b = B_app typid bdc" using subst_bb_inject * sory (* going to be like the above as bdc is closed *)*) + have "atom bv \ b1" proof - + have "supp b1 \ { atom bva }" using wfTh_poly_supp_b ** SConsp by metis + moreover have "bv \ bva" using ** by auto + ultimately show ?thesis using fresh_def by force + qed + have "T \ s : b1[bva::=bdc]\<^sub>b\<^sub>b" using ** by auto + moreover have "b1[bva::=bdc']\<^sub>b\<^sub>b[bv::=b']\<^sub>b\<^sub>b = b1[bva::=bdc]\<^sub>b\<^sub>b" using bdc subst_bb_commute \atom bv \ b1\ by auto + ultimately obtain s' where s':"T \ s ~ b1[bva::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \ s'" + using SConsp(1)[of "b1[bva::=bdc']\<^sub>b\<^sub>b"] bdc SConsp by metis + have "T \ SConsp typid dc bdc'[bv::=b']\<^sub>b\<^sub>b s ~ (B_app typid bdc') [ bv ::= b' ] \ SConsp typid dc bdc' s'" + proof - + + obtain bva3 and dclist3 where 3:"AF_typedef_poly typid bva3 dclist3 = AF_typedef_poly typid bva dclist \ + atom bva3 \ (bdc', bv, b', s, s')" using obtain_fresh_bv by metis + then obtain x3 b3 c3 where 4:"(dc, \ x3 : b3 | c3 \) \ set dclist3" + using boxed_b_BConspI td_lookup_eq_iff type_def.eq_iff + by (metis "**") + + show ?thesis proof + show \AF_typedef_poly typid bva3 dclist3 \ set T\ using 3 ** by metis + show \atom bva3 \ (bdc', bv, b', s, s')\ using 3 by metis + show 4:\(dc, \ x3 : b3 | c3 \) \ set dclist3\ using 4 by auto + have "b3[bva3::=bdc']\<^sub>b\<^sub>b = b1[bva::=bdc']\<^sub>b\<^sub>b" proof(rule wfTh_typedef_poly_b_eq_iff) + show \AF_typedef_poly typid bva3 dclist3 \ set T\ using 3 ** by metis + show \(dc, \ x3 : b3 | c3 \) \ set dclist3\ using 4 by auto + show \AF_typedef_poly typid bva dclist \ set T\ using ** by auto + show \(dc, \ x1 : b1 | c1 \) \ set dclist\ using ** by auto + qed(simp add: ** SConsp) + thus \ T \ s ~ b3[bva3::=bdc']\<^sub>b\<^sub>b [ bv ::= b' ] \ s' \ using s' by auto + qed + qed + then show ?thesis using bdc by auto + + qed +next + case SUnit + have *:"b[bv::=b']\<^sub>b\<^sub>b = B_unit" using wfRCV_elims SUnit by metis + show ?case proof (cases "b = B_var bv") + case True + moreover have "T \ SUnit : B_unit" using wfRCV.intros by simp + moreover hence "b' = B_unit" using True SUnit subst_bb.simps * by simp + ultimately show ?thesis using boxed_b.intros wfRCV.intros by metis + next + case False + hence "b = B_unit" using subst_bb_inject * by metis + then show ?thesis using * SUnit boxed_b.intros by metis + qed +next + case (SUt x) + then obtain bv' where *:"b[bv::=b']\<^sub>b\<^sub>b= B_var bv'" using wfRCV_elims by metis + show ?case proof (cases "b = B_var bv") + case True + then show ?thesis using boxed_b_BVar1I + using "local.*" wfRCV_BVarI by fastforce + next + case False + then show ?thesis using boxed_b_BVar1I boxed_b_BVar2I + using "local.*" wfRCV_BVarI by (metis subst_b_var) + qed +qed + +lemma boxed_i_ex: + assumes "wfI T \[bv::=b]\<^sub>\\<^sub>b i" and "wfTh T" + shows "\i'. T ; \ ; b , bv \ i \ i'" + using assms proof(induct \ arbitrary: i rule:\_induct) + case GNil + then show ?case using boxed_i_GNilI by metis +next + case (GCons x' b' c' \') + then obtain s where 1:"Some s = i x' \ wfRCV T s b'[bv::=b]\<^sub>b\<^sub>b" using wfI_def subst_gb.simps by auto + then obtain s' where 2: "boxed_b T s b' bv b s'" using boxed_b_ex GCons by metis + then obtain i' where 3: "boxed_i T \' b bv i i'" using GCons wfI_def subst_gb.simps by force + have "boxed_i T ((x', b', c') #\<^sub>\ \') b bv i (i'(x' \ s'))" proof + show "Some s = i x'" using 1 by auto + show "boxed_b T s b' bv b s'" using 2 by auto + show "T ; \' ; b , bv \ i \ i'" using "3" by auto + qed + thus ?case by auto +qed + +lemma boxed_b_eq: + assumes "boxed_b \ s1 b bv b' s1'" and "\\<^sub>w\<^sub>f \" + shows "wfTh \ \ boxed_b \ s2 b bv b' s2' \ ( s1 = s2 ) = ( s1' = s2' )" + using assms proof(induct arbitrary: s2 s2' rule: boxed_b.inducts ) + case (boxed_b_BVar1I bv bv' P s b ) + then show ?case + using boxed_b_elims(1) rcl_val.eq_iff by metis +next + case (boxed_b_BVar2I bv bv' P s b) + then show ?case using boxed_b_elims(1) by metis +next + case (boxed_b_BIntI P s uu uv) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BBoolI P s uw ux) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BUnitI P s uy uz) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BPairI P s1 b1 bv b s1' s2a b2 s2a') + then show ?case + by (metis boxed_b_elims(5) rcl_val.eq_iff(4)) +next + case (boxed_b_BConsI tyid dclist P dc x b c s1 bv b' s1') + obtain s22 and s22' dclist2 dc2 x2 b2 c2 where *:"s2 = SCons tyid dc2 s22 \ s2' = SCons tyid dc2 s22' \ boxed_b P s22 b2 bv b' s22' + \ AF_typedef tyid dclist2 \ set P \ (dc2, \ x2 : b2 | c2 \) \ set dclist2" using boxed_b_elims(6)[OF boxed_b_BConsI(6)] by metis + show ?case proof(cases "dc = dc2") + case True + hence "b = b2" using wfTh_ctor_unique \.eq_iff wfTh_dclist_unique wf boxed_b_BConsI * by metis + then show ?thesis using boxed_b_BConsI True * by auto + next + case False + then show ?thesis using * boxed_b_BConsI by simp + qed +next + case (boxed_b_Bbitvec P s bv b) + hence "s2 = s2'" using boxed_b_elims by metis + then show ?case by auto +next + case (boxed_b_BConspI tyid bva dclist P b1 bv b' s1 s1' dc x b c) + obtain bva2 s22 s22' dclist2 dc2 x2 b2 c2 where *:" + s2 = SConsp tyid dc2 b1[bv::=b']\<^sub>b\<^sub>b s22 \ + s2' = SConsp tyid dc2 b1 s22' \ + boxed_b P s22 b2[bva2::=b1]\<^sub>b\<^sub>b bv b' s22' \ + AF_typedef_poly tyid bva2 dclist2 \ set P \ (dc2, \ x2 : b2 | c2 \) \ set dclist2" using boxed_b_elims(8)[OF boxed_b_BConspI(7)] by metis + show ?case proof(cases "dc = dc2") + case True + hence "AF_typedef_poly tyid bva2 dclist2 \ set P \ (dc, \ x2 : b2 | c2 \) \ set dclist2" using * by auto + hence "b[bva::=b1]\<^sub>b\<^sub>b = b2[bva2::=b1]\<^sub>b\<^sub>b" using wfTh_typedef_poly_b_eq_iff[OF boxed_b_BConspI(1) boxed_b_BConspI(3)] * boxed_b_BConspI by metis + then show ?thesis using boxed_b_BConspI True * by auto + next + case False + then show ?thesis using * boxed_b_BConspI by simp + qed +qed + +lemma bs_boxed_var: + assumes "boxed_i \ \ b' bv i i'" + shows "Some (b,c) = lookup \ x \ Some s = i x \ Some s' = i' x \ boxed_b \ s b bv b' s'" + using assms proof(induct rule: boxed_i.inducts) + case (boxed_i_GNilI T i) + then show ?case using lookup.simps by auto +next + case (boxed_i_GConsI s i x1 \ b1 bv b' s' \ i' c) + show ?case proof (cases "x=x1") + case True + then show ?thesis using boxed_i_GConsI + fun_upd_same lookup.simps(2) option.inject prod.inject by metis + next + case False + then show ?thesis using boxed_i_GConsI + fun_upd_same lookup.simps option.inject prod.inject by auto + qed +qed + +lemma eval_l_boxed_b: + assumes "\ l \ = s" + shows "boxed_b \ s (base_for_lit l) bv b' s" + using assms proof(nominal_induct l arbitrary: s rule:l.strong_induct) +qed(auto simp add: boxed_b.intros wfRCV.intros )+ + +lemma boxed_i_eval_v_boxed_b: + fixes v::v + assumes "boxed_i \ \ b' bv i i'" and "i \ v[bv::=b']\<^sub>v\<^sub>b \ ~ s" and "i' \ v \ ~ s'" and "wfV \ B \ v b" and "wfI \ \ i'" + shows "boxed_b \ s b bv b' s'" + using assms proof(nominal_induct v arbitrary: s s' b rule:v.strong_induct) + case (V_lit l) + hence "\ l \ = s \ \ l \ = s'" using eval_v_elims by auto + moreover have "b = base_for_lit l" using wfV_elims(2) V_lit by metis + ultimately show ?case using V_lit using eval_l_boxed_b subst_b_base_for_lit by metis +next + case (V_var x) (* look at defn of bs_boxed *) + hence "Some s = i x \ Some s' = i' x" using eval_v_elims subst_vb.simps by metis + moreover obtain c1 where bc:"Some (b,c1) = lookup \ x" using wfV_elims V_var by metis + ultimately show ?case using bs_boxed_var V_var by metis + +next + case (V_pair v1 v2) + then obtain b1 and b2 where b:"b=B_pair b1 b2" using wfV_elims subst_vb.simps by metis + obtain s1 and s2 where s: "eval_v i (v1[bv::=b']\<^sub>v\<^sub>b) s1 \ eval_v i (v2[bv::=b']\<^sub>v\<^sub>b) s2 \ s = SPair s1 s2" using eval_v_elims V_pair subst_vb.simps by metis + obtain s1' and s2' where s': "eval_v i' v1 s1' \ eval_v i' v2 s2' \ s' = SPair s1' s2'" using eval_v_elims V_pair by metis + have "boxed_b \ (SPair s1 s2) (B_pair b1 b2) bv b' (SPair s1' s2') " proof(rule boxed_b_BPairI) + show "boxed_b \ s1 b1 bv b' s1'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis + show "boxed_b \ s2 b2 bv b' s2'" using V_pair eval_v_elims wfV_elims b s s' b.eq_iff by metis + qed + then show ?case using s s' b by auto +next + case (V_cons tyid dc v1) + + obtain dclist x b1 c where *: "b = B_id tyid \ AF_typedef tyid dclist \ set \ \ (dc, \ x : b1 | c \) \ set dclist \ \ ; B ; \ \\<^sub>w\<^sub>f v1 : b1" + using wfV_elims(4)[OF V_cons(5)] V_cons by metis + obtain s2 where s2: "s = SCons tyid dc s2 \ i \ (v1[bv::=b']\<^sub>v\<^sub>b) \ ~ s2" using eval_v_elims V_cons subst_vb.simps by metis + obtain s2' where s2': "s' = SCons tyid dc s2' \ i' \ v1 \ ~ s2'" using eval_v_elims V_cons by metis + have sp: "supp \ x : b1 | c \ = {}" using wfTh_lookup_supp_empty * wfX_wfY by metis + + have "boxed_b \ (SCons tyid dc s2) (B_id tyid) bv b' (SCons tyid dc s2')" + proof(rule boxed_b_BConsI) + show 1:"AF_typedef tyid dclist \ set \" using * by auto + show 2:"(dc, \ x : b1 | c \) \ set dclist" using * by auto + have bvf:"atom bv \ b1" using sp \.fresh fresh_def by auto + show "\ \ s2 ~ b1 [ bv ::= b' ] \ s2'" using V_cons s2 s2' * by metis + qed + then show ?case using * s2 s2' by simp +next + case (V_consp tyid dc b1 v1) + + obtain bv2 dclist x2 b2 c2 where *: "b = B_app tyid b1 \ AF_typedef_poly tyid bv2 dclist \ set \ \ + (dc, \ x2 : b2 | c2 \) \ set dclist \ \ ; B ; \ \\<^sub>w\<^sub>f v1 : b2[bv2::=b1]\<^sub>b\<^sub>b" + using wf_strong_elim(1)[OF V_consp (5)] by metis + + obtain s2 where s2: "s = SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2 \ i \ (v1[bv::=b']\<^sub>v\<^sub>b) \ ~ s2" + using eval_v_elims V_consp subst_vb.simps by metis + + obtain s2' where s2': "s' = SConsp tyid dc b1 s2' \ i' \ v1 \ ~ s2'" + using eval_v_elims V_consp by metis + + have "\\<^sub>w\<^sub>f \" using V_consp wfX_wfY by metis + then obtain bv3::bv and dclist3 x3 b3 c3 where **:" AF_typedef_poly tyid bv2 dclist = AF_typedef_poly tyid bv3 dclist3 \ + (dc, \ x3 : b3 | c3 \) \ set dclist3 \ atom bv3 \ (b1, bv, b', s2, s2') \ b2[bv2::=b1]\<^sub>b\<^sub>b = b3[bv3::=b1]\<^sub>b\<^sub>b" + using * obtain_fresh_bv_dclist_b_iff[where tm="(b1, bv, b', s2, s2')"] by metis + + have "boxed_b \ (SConsp tyid dc b1[bv::=b']\<^sub>b\<^sub>b s2) (B_app tyid b1) bv b' (SConsp tyid dc b1 s2')" + proof(rule boxed_b_BConspI[of tyid bv3 dclist3 \, where x=x3 and b=b3 and c=c3]) + show 1:"AF_typedef_poly tyid bv3 dclist3 \ set \" using * ** by auto + show 2:"(dc, \ x3 : b3 | c3 \) \ set dclist3" using ** by auto + show "atom bv3 \ (b1, bv, b', s2, s2')" using ** by auto + show " \ \ s2 ~ b3[bv3::=b1]\<^sub>b\<^sub>b [ bv ::= b' ] \ s2'" using V_consp s2 s2' * ** by metis + qed + then show ?case using * s2 s2' by simp +qed + +lemma boxed_b_eq_eq: + assumes "boxed_b \ n1 b1 bv b' n1'" and "boxed_b \ n2 b1 bv b' n2'" and "s = SBool (n1 = n2)" and "\\<^sub>w\<^sub>f \" + "s' = SBool (n1' = n2')" + shows "s=s'" + using boxed_b_eq assms by auto + +lemma boxed_i_eval_ce_boxed_b: + fixes e::ce + assumes "i' \ e \ ~ s'" and "i \ e[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ s" and "wfCE \ B \ e b" and "boxed_i \ \ b' bv i i'" and "wfI \ \ i'" + shows "boxed_b \ s b bv b' s'" + using assms proof(nominal_induct e arbitrary: s s' b b' rule: ce.strong_induct) + case (CE_val x) + then show ?case using boxed_i_eval_v_boxed_b eval_e_elims wfCE_elims subst_ceb.simps by metis +next + case (CE_op opp v1 v2) + + show ?case proof(rule opp.exhaust) + assume \opp = Plus\ + have 1:"wfCE \ B \ v1 (B_int)" using wfCE_elims CE_op \opp = Plus\ by metis + have 2:"wfCE \ B \ v2 (B_int)" using wfCE_elims CE_op \opp = Plus\ by metis + have *:"b = B_int" using CE_op wfCE_elims + by (metis \opp = plus\) + + obtain n1 and n2 where n:"s = SNum (n1 + n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n2" using eval_e_elims CE_op subst_ceb.simps \opp = plus\ by metis + obtain n1' and n2' where n':"s' = SNum (n1' + n2') \ i' \ v1 \ ~ SNum n1' \ i' \ v2 \ ~ SNum n2'" using eval_e_elims Plus CE_op \opp = plus\ by metis + + have "boxed_b \ (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op \opp = plus\ by metis + moreover have "boxed_b \ (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + ultimately have "s=s'" using n' n boxed_b_elims(2) + by (metis rcl_val.eq_iff(2)) + thus ?thesis using * n n' boxed_b_BIntI CE_op wfRCV.intros Plus by simp + next + assume \opp = LEq\ + have 1:"wfCE \ B \ v1 (B_int)" using wfCE_elims CE_op \opp = LEq\ by metis + have 2:"wfCE \ B \ v2 (B_int)" using wfCE_elims CE_op \opp = LEq\ by metis + hence *:"b = B_bool" using CE_op wfCE_elims \opp = LEq\ by metis + obtain n1 and n2 where n:"s = SBool (n1 \ n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SNum n2" using eval_e_elims subst_ceb.simps CE_op \opp = LEq\ by metis + obtain n1' and n2' where n':"s' = SBool (n1' \ n2') \ i' \ v1 \ ~ SNum n1' \ i' \ v2 \ ~ SNum n2'" using eval_e_elims CE_op \opp = LEq\ by metis + + have "boxed_b \ (SNum n1) B_int bv b' (SNum n1')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + moreover have "boxed_b \ (SNum n2) B_int bv b' (SNum n2')" using boxed_i_eval_v_boxed_b 1 2 n n' CE_op by metis + ultimately have "s=s'" using n' n boxed_b_elims(2) + by (metis rcl_val.eq_iff(2)) + thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \opp = LEq\ by simp + next + assume \opp = Eq\ + obtain b1 where b1:"wfCE \ B \ v1 b1 \ wfCE \ B \ v2 b1" using wfCE_elims CE_op \opp = Eq\ by metis + + hence *:"b = B_bool" using CE_op wfCE_elims \opp = Eq\ by metis + obtain n1 and n2 where n:"s = SBool (n1 = n2) \ i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ n1 \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ n2" using eval_e_elims subst_ceb.simps CE_op \opp = Eq\ by metis + obtain n1' and n2' where n':"s' = SBool (n1' = n2') \ i' \ v1 \ ~ n1' \ i' \ v2 \ ~ n2'" using eval_e_elims CE_op \opp = Eq\ by metis + + have "boxed_b \ n1 b1 bv b' n1'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis + moreover have "boxed_b \ n2 b1 bv b' n2'" using boxed_i_eval_v_boxed_b b1 n n' CE_op by metis + moreover have "\\<^sub>w\<^sub>f \" using b1 wfX_wfY by metis + ultimately have "s=s'" using n' n boxed_b_elims + boxed_b_eq_eq by metis + thus ?thesis using * n n' boxed_b_BBoolI CE_op wfRCV.intros \opp = Eq\ by simp + qed + +next + case (CE_concat v1 v2) + + obtain bv1 and bv2 where s : "s = SBitvec (bv1 @ bv2) \ (i \ v1[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec bv1) \ i \ v2[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec bv2" + using eval_e_elims(7) subst_ceb.simps CE_concat.prems(2) eval_e_elims(6) subst_ceb.simps(6) by metis + obtain bv1' and bv2' where s' : "s' = SBitvec (bv1' @ bv2') \ i' \ v1 \ ~ SBitvec bv1' \ i' \ v2 \ ~ SBitvec bv2'" + using eval_e_elims(7) CE_concat by metis + + then show ?case using boxed_i_eval_v_boxed_b wfCE_elims s s' CE_concat + by (metis CE_concat.prems(3) assms assms(5) wfRCV_BBitvecI boxed_b_Bbitvec boxed_b_elims(7) eval_e_concatI eval_e_uniqueness) +next + case (CE_fst ce) + obtain s2 where 1:"i \ ce[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SPair s s2" using CE_fst eval_e_elims subst_ceb.simps by metis + obtain s2' where 2:"i' \ ce \ ~ SPair s' s2'" using CE_fst eval_e_elims by metis + obtain b2 where 3:"wfCE \ B \ ce (B_pair b b2)" using wfCE_elims(4) CE_fst by metis + + have "boxed_b \ (SPair s s2) (B_pair b b2) bv b' (SPair s' s2')" + using 1 2 3 CE_fst boxed_i_eval_v_boxed_b boxed_b_BPairI by auto + thus ?case using boxed_b_elims(5) by force +next + case (CE_snd v) + obtain s1 where 1:"i \ v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SPair s1 s" using CE_snd eval_e_elims subst_ceb.simps by metis + obtain s1' where 2:"i' \ v \ ~ SPair s1' s'" using CE_snd eval_e_elims by metis + obtain b1 where 3:"wfCE \ B \ v (B_pair b1 b )" using wfCE_elims(5) CE_snd by metis + + have "boxed_b \ (SPair s1 s ) (B_pair b1 b ) bv b' (SPair s1' s')" using 1 2 3 CE_snd boxed_i_eval_v_boxed_b by simp + thus ?case using boxed_b_elims(5) by force +next + case (CE_len v) + obtain s1 where s: "i \ v[bv::=b']\<^sub>c\<^sub>e\<^sub>b \ ~ SBitvec s1" using CE_len eval_e_elims subst_ceb.simps by metis + obtain s1' where s': "i' \ v \ ~ SBitvec s1'" using CE_len eval_e_elims by metis + + have "\ ; B ; \ \\<^sub>w\<^sub>f v : B_bitvec \ b = B_int" using wfCE_elims CE_len by metis + then show ?case using boxed_i_eval_v_boxed_b s s' CE_len + by (metis boxed_b_BIntI boxed_b_elims(7) eval_e_lenI eval_e_uniqueness subst_ceb.simps(5) wfI_wfCE_eval_e) +qed + +lemma eval_c_eq_bs_boxed: + fixes c::c + assumes "i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s" and "i' \ c \ ~ s'" and "wfC \ B \ c" and "wfI \ \ i'" and "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i " + and "boxed_i \ \ b bv i i'" + shows "s = s'" + using assms proof(nominal_induct c arbitrary: s s' rule:c.strong_induct) + case C_true + then show ?case using eval_c_elims subst_cb.simps by metis +next + case C_false + then show ?case using eval_c_elims subst_cb.simps by metis +next + case (C_conj c1 c2) + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1\s2)" using C_conj eval_c_elims(3) subst_cb.simps(3) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1'\s2')" using C_conj eval_c_elims(3) by metis + then show ?case using 1 2 wfC_elims C_conj by metis +next + case (C_disj c1 c2) + + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1\s2)" using C_disj eval_c_elims(4) subst_cb.simps(4) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1'\s2')" using C_disj eval_c_elims(4) by metis + then show ?case using 1 2 wfC_elims C_disj by metis +next + case (C_not c) + obtain s1::bool where 1: "(i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s1) \ (s = (\ s1))" using C_not eval_c_elims(6) subst_cb.simps(7) by metis + obtain s1'::bool where 2: "(i' \ c \ ~ s1') \ (s' = (\ s1'))" using C_not eval_c_elims(6) by metis + then show ?case using 1 2 wfC_elims C_not by metis +next + case (C_imp c1 c2) + obtain s1 and s2 where 1: "eval_c i (c1[bv::=b]\<^sub>c\<^sub>b) s1 \ eval_c i (c2[bv::=b]\<^sub>c\<^sub>b) s2 \ s = (s1 \ s2)" using C_imp eval_c_elims(5) subst_cb.simps(5) by metis + obtain s1' and s2' where 2:"eval_c i' c1 s1' \ eval_c i' c2 s2' \ s' = (s1' \ s2')" using C_imp eval_c_elims(5) by metis + then show ?case using 1 2 wfC_elims C_imp by metis +next + case (C_eq e1 e2) + obtain be where be: "wfCE \ B \ e1 be \ wfCE \ B \ e2 be" using C_eq wfC_elims by metis + obtain s1 and s2 where 1: "eval_e i (e1[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s1 \ eval_e i (e2[bv::=b]\<^sub>c\<^sub>e\<^sub>b) s2 \ s = (s1 = s2)" using C_eq eval_c_elims(7) subst_cb.simps(6) by metis + obtain s1' and s2' where 2:"eval_e i' e1 s1' \ eval_e i' e2 s2' \ s' = (s1' = s2' )" using C_eq eval_c_elims(7) by metis + have "\\<^sub>w\<^sub>f \" using C_eq wfX_wfY by metis + moreover have "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i " using C_eq by auto + ultimately show ?case using boxed_b_eq[of \ s1 be bv b s1' s2 s2'] 1 2 boxed_i_eval_ce_boxed_b C_eq wfC_elims subst_cb.simps 1 2 be by auto +qed + +lemma is_satis_bs_boxed: + fixes c::c + assumes "boxed_i \ \ b bv i i'" and "wfC \ B \ c" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\ ; \ \ i'" + and "(i \ c[bv::=b]\<^sub>c\<^sub>b)" + shows "(i' \ c)" +proof - + have "eval_c i (c[bv::=b]\<^sub>c\<^sub>b) True" using is_satis.simps assms by auto + moreover obtain s where "i' \ c \ ~ s" using eval_c_exist assms by metis + ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis +qed + +lemma is_satis_bs_boxed_rev: + fixes c::c + assumes "boxed_i \ \ b bv i i'" and "wfC \ B \ c" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\ ; \ \ i'" and "wfC \ {||} \[bv::=b]\<^sub>\\<^sub>b (c[bv::=b]\<^sub>c\<^sub>b)" + and "(i' \ c)" + shows "(i \ c[bv::=b]\<^sub>c\<^sub>b)" +proof - + have "eval_c i' c True" using is_satis.simps assms by auto + moreover obtain s where "i \ c[bv::=b]\<^sub>c\<^sub>b \ ~ s" using eval_c_exist assms by metis + ultimately show ?thesis using eval_c_eq_bs_boxed assms is_satis.simps by metis +qed + +lemma bs_boxed_wfi_aux: + fixes b::b and bv::bv and \::\ and B::\ + assumes "boxed_i \ \ b bv i i'" and "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" and "\\<^sub>w\<^sub>f \" and "wfG \ B \" + shows "\ ; \ \ i'" + using assms proof(induct rule: boxed_i.inducts) + case (boxed_i_GNilI T i) + then show ?case using wfI_def by auto +next + case (boxed_i_GConsI s i x1 T b1 bv b s' G i' c1) + { + fix x2 b2 c2 + assume as : "(x2,b2,c2) \ toSet ((x1, b1, c1) #\<^sub>\ G)" + + then consider (hd) "(x2,b2,c2) = (x1, b1, c1)" | (tail) "(x2,b2,c2) \ toSet G" using toSet.simps by auto + hence "\s. Some s = (i'(x1 \ s')) x2 \ wfRCV T s b2" proof(cases) + case hd + hence "b1=b2" by auto + moreover have "(x2,b2[bv::=b]\<^sub>b\<^sub>b,c2[bv::=b]\<^sub>c\<^sub>b) \ toSet ((x1, b1, c1) #\<^sub>\ G)[bv::=b]\<^sub>\\<^sub>b" using hd subst_gb.simps by simp + moreover hence "wfRCV T s b2[bv::=b]\<^sub>b\<^sub>b" using wfI_def boxed_i_GConsI hd + proof - + obtain ss :: "b \ x \ (x \ rcl_val option) \ type_def list \ rcl_val" where + "\x1a x2a x3 x4. (\v5. Some v5 = x3 x2a \ wfRCV x4 v5 x1a) = (Some (ss x1a x2a x3 x4) = x3 x2a \ wfRCV x4 (ss x1a x2a x3 x4) x1a)" + by moura (* 0.0 ms *) + then have f1: "Some (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) = i x1 \ wfRCV T (ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T) b2[bv::=b]\<^sub>b\<^sub>b" + using boxed_i_GConsI.prems(1) hd wfI_def by auto (* 31 ms *) + then have "ss b2[bv::=b]\<^sub>b\<^sub>b x1 i T = s" + by (metis (no_types) boxed_i_GConsI.hyps(1) option.inject) (* 0.0 ms *) + then show ?thesis + using f1 by blast (* 0.0 ms *) + qed + ultimately have "wfRCV T s' b2" using boxed_i_GConsI boxed_b_wfRCV by metis + + then show ?thesis using hd by simp + next + case tail + hence "wfI T G i'" using boxed_i_GConsI wfI_suffix wfG_suffix subst_gb.simps + by (metis (no_types, lifting) Un_iff toSet.simps(2) wfG_cons2 wfI_def) + then show ?thesis using wfI_def[of T G i'] tail + using boxed_i_GConsI.prems(3) split_G wfG_cons_fresh2 by fastforce + qed + } + thus ?case using wfI_def by fast + +qed + +lemma is_satis_g_bs_boxed_aux: + fixes G::\ + assumes "boxed_i \ G1 b bv i i'" and "wfI \ G1[bv::=b]\<^sub>\\<^sub>b i" and "wfI \ G1 i'" and "G1 = (G2@G)" and "wfG \ B G1" + and "(i \ G[bv::=b]\<^sub>\\<^sub>b) " + shows "(i' \ G)" + using assms proof(induct G arbitrary: G2 rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \' G2) + show ?case proof(subst is_satis_g.simps,rule) + have *:"wfC \ B G1 c'" using GCons wfG_wfC_inside by force + show "i' \ c'" using is_satis_bs_boxed[OF assms(1) * ] GCons by auto + obtain G3 where "G1 = G3 @ \'" using GCons append_g.simps + by (metis append_g_assoc) + then show "i' \ \'" using GCons append_g.simps by simp + qed +qed + +lemma is_satis_g_bs_boxed: + fixes G::\ + assumes "boxed_i \ G b bv i i'" and "wfI \ G[bv::=b]\<^sub>\\<^sub>b i" and "wfI \ G i'" and "wfG \ B G" + and "(i \ G[bv::=b]\<^sub>\\<^sub>b) " + shows "(i' \ G)" + using is_satis_g_bs_boxed_aux assms + by (metis (full_types) append_g.simps(1)) + +lemma subst_b_valid: + fixes s::s and b::b + assumes "\ ; {||} \\<^sub>w\<^sub>f b" and "B = {|bv|}" and "\ ; {|bv|} ;\ \ c" + shows "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \ c[bv::=b]\<^sub>c\<^sub>b " +proof(rule validI) + + show **:"\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using assms valid.simps wf_b_subst subst_gb.simps by metis + show "\i. (wfI \ \[bv::=b]\<^sub>\\<^sub>b i \ i \ \[bv::=b]\<^sub>\\<^sub>b) \ i \ c[bv::=b]\<^sub>c\<^sub>b " + proof(rule,rule) + fix i + assume *:"wfI \ \[bv::=b]\<^sub>\\<^sub>b i \ i \ \[bv::=b]\<^sub>\\<^sub>b" + + obtain i' where idash: "boxed_i \ \ b bv i i'" using boxed_i_ex wfX_wfY assms * by fastforce + + have wfc: "\ ; {|bv|} ; \ \\<^sub>w\<^sub>f c" using valid.simps assms by simp + have wfg: "\ ; {|bv|} \\<^sub>w\<^sub>f \" using valid.simps wfX_wfY assms by metis + hence wfi: "wfI \ \ i'" using idash * bs_boxed_wfi_aux subst_gb.simps wfX_wfY by metis + moreover have "i' \ \" proof (rule is_satis_g_bs_boxed[OF idash ] wfX_wfY(2)[OF wfc]) + show "wfI \ \[bv::=b]\<^sub>\\<^sub>b i" using subst_gb.simps * by simp + show "wfI \ \ i'" using wfi by auto + show "\ ; B \\<^sub>w\<^sub>f \ " using wfg assms by auto + show "i \ \[bv::=b]\<^sub>\\<^sub>b" using subst_gb.simps * by simp + qed + ultimately have ic:"i' \ c" using assms valid_def using valid.simps by blast + + show "i \ c[bv::=b]\<^sub>c\<^sub>b" proof(rule is_satis_bs_boxed_rev) + show "\ ; \ ; b , bv \ i \ i'" using idash by auto + show "\ ; B ; \ \\<^sub>w\<^sub>f c " using wfc assms by auto + show "\ ; \[bv::=b]\<^sub>\\<^sub>b \ i" using subst_gb.simps * by simp + show "\ ; \ \ i'" using wfi by auto + show "\ ; {||} ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using ** by auto + show "i' \ c" using ic by auto + qed + + qed +qed + +section \Expression Operator Lemmas\ + +lemma is_satis_len_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (L_num (int (length v)))) )" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit (L_num (int (length v))))) (SNum (int (length v)))" + using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SNum (int (length v)))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_len [V_lit (L_bitvec v)]\<^sup>c\<^sup>e) (SNum (int (length v)))" + using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma is_satis_plus_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (L_num (n1+n2))))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e))" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit (L_num (n1+n2)))) (SNum (n1+n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps by (metis eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SNum (n1+n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)) (SNum (n1+n2))" + using eval_e_elims(7) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma is_satis_leq_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (if (n1 \ n2) then L_true else L_false)))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit ((if (n1 \ n2) then L_true else L_false)))) (SBool (n1\n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps + by (metis (full_types) eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SBool (n1\n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op LEq [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2) )]\<^sup>c\<^sup>e) (SBool (n1\n2))" + using eval_e_elims(3) eval_v_elims eval_l.simps by (metis eval_e.intros eval_v_litI) + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma eval_lit_inj: + fixes n1::l and n2::l + assumes "\ n1 \ = s" and "\ n2 \ = s" + shows "n1=n2" + using assms proof(nominal_induct s rule: rcl_val.strong_induct) + case (SBitvec x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SNum x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SBool x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SPair x1a x2a) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SCons x1a x2a x3a) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SConsp x1a x2a x3a x4) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case SUnit + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +next + case (SUt x) + then show ?case using eval_l.simps + by (metis l.strong_exhaust rcl_val.distinct rcl_val.eq_iff) +qed + +lemma eval_e_lit_inj: + fixes n1::l and n2::l + assumes "i \ [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" and "i \ [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" + shows "n1=n2" + using eval_lit_inj assms eval_e_elims eval_v_elims by metis + +lemma is_satis_eq_imp: + assumes "i \ (CE_val (V_var x) == CE_val (V_lit (if (n1 = n2) then L_true else L_false)))" (is "is_satis i ?c1") + shows "i \ (CE_val (V_var x) == CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2))]\<^sup>c\<^sup>e)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + then have "eval_e i (CE_val (V_lit ((if (n1=n2) then L_true else L_false)))) (SBool (n1=n2))" + using eval_e_elims(1) eval_v_elims eval_l.simps + by (metis (full_types) eval_e.intros(1) eval_v_litI) + hence "eval_e i (CE_val (V_var x)) (SBool (n1=n2))" using eval_c_elims(7)[OF *] + by (metis eval_e_elims(1) eval_v_elims(1)) + moreover have "eval_e i (CE_op Eq [(V_lit (n1))]\<^sup>c\<^sup>e [(V_lit (n2) )]\<^sup>c\<^sup>e) (SBool (n1=n2))" + proof - + obtain s1 and s2 where *:"i \ [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1 \ i \ [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s2" using eval_l.simps eval_e.intros eval_v_litI by metis + moreover have " SBool (n1 = n2) = SBool (s1 = s2)" proof(cases "n1=n2") + case True + then show ?thesis using * + by (simp add: calculation eval_e_uniqueness) + next + case False + then show ?thesis using * eval_e_lit_inj by auto + qed + ultimately show ?thesis using eval_e_eqI[of i "[(V_lit (n1))]\<^sup>c\<^sup>e" s1 "[(V_lit (n2))]\<^sup>c\<^sup>e" s2 ] by auto + qed + ultimately show ?thesis using eval_c.intros is_satis.simps by fastforce +qed + +lemma valid_eq_e: + assumes "\i s1 s2. wfG P \ GNil \ wfI P GNil i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" + and "wfCE P \ GNil e1 b" and "wfCE P \ GNil e2 b" + shows "P ; \ ; (x, b , CE_val (V_var x) == e1 )#\<^sub>\ GNil \ CE_val (V_var x) == e2" + unfolding valid.simps +proof(intro conjI) + show \ P ; \ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \ + using assms wf_intros wfX_wfY b.eq_iff fresh_GNil wfC_e_eq2 wfV_elims by meson + show \\i. ((P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \ i) \ (i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil) \ + (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2)) \ proof(rule+) + fix i + assume as:"P ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil \ i \ i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1 ) #\<^sub>\ GNil" + + have *: "P ; GNil \ i " using wfI_def by auto + + then obtain s1 where s1:"eval_e i e1 s1" using assms eval_e_exist by metis + obtain s2 where s2:"eval_e i e2 s2" using assms eval_e_exist * by metis + moreover have "i x = Some s1" proof - + have "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e1" using as is_satis_g.simps by auto + thus ?thesis using s1 + by (metis eval_c_elims(7) eval_e_elims(1) eval_e_uniqueness eval_v_elims(2) is_satis.cases) + qed + moreover have "s1 = s2" using s1 s2 * assms wfG_nilI wfX_wfY by metis + + ultimately show "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == e2 \ ~ True" + using eval_c.intros eval_e.intros eval_v.intros + proof - + have "i \ e2 \ ~ s1" + by (metis \s1 = s2\ s2) (* 0.0 ms *) + then show ?thesis + by (metis (full_types) \i x = Some s1\ eval_c_eqI eval_e_valI eval_v_varI) (* 31 ms *) + qed + qed +qed + +lemma valid_len: + assumes " \\<^sub>w\<^sub>f \" + shows "\ ; \ ; (x, B_int, [[x]\<^sup>v]\<^sup>c\<^sup>e == [[ L_num (int (length v)) ]\<^sup>v]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == CE_len [[ L_bitvec v ]\<^sup>v]\<^sup>c\<^sup>e" (is "\ ; \ ; ?G \ ?c" ) +proof - + have *:"\ \\<^sub>w\<^sub>f ([]::\) \ \ ; \ ; GNil \\<^sub>w\<^sub>f []\<^sub>\ " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto + + moreover hence "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int" + using wfCE_valI * wfV_litI base_for_lit.simps + by (metis wfE_valI wfX_wfY) + + moreover have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int" + using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY wfCE_valI + by (metis wfCE_lenI) + moreover have "atom x \ GNil" by auto + ultimately have "\ ; \ ; ?G \\<^sub>w\<^sub>f ?c" using wfC_e_eq2 assms by simp + moreover have "(\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c)" using is_satis_len_imp by auto + ultimately show ?thesis using valid.simps by auto +qed + +lemma valid_arith_bop: + assumes "wfG \ \ \" and "opp = Plus \ ll = (L_num (n1+n2)) \ (opp = LEq \ ll = ( if n1\n2 then L_true else L_false))" + and "(opp = Plus \ b = B_int) \ (opp = LEq \ b = B_bool)" and + "atom x \ \" + shows "\; \ ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e ))" (is "\ ; \ ; ?G \ ?c") +proof - + have "wfC \ \ ?G ?c" proof(rule wfC_e_eq2) + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI wfV_litI assms base_for_lit.simps by metis + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) : b " + using wfCE_plusI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis + show "\\<^sub>w\<^sub>f \" using assms wfX_wfY by auto + show "atom x \ \" using assms by auto + qed + + moreover have "\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume "wfI \ ?G i \ is_satis_g i ?G" + + hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (ll)) ))" by auto + thus "is_satis i ((CE_val (V_var x) == CE_op opp ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e)))" + using is_satis_plus_imp assms opp.exhaust is_satis_leq_imp by auto + qed + ultimately show ?thesis using valid.simps by metis +qed + +lemma valid_eq_bop: + assumes "wfG \ \ \" and "atom x \ \" and "base_for_lit l1 = base_for_lit l2" + shows "\; \ ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e ))" (is "\ ; \ ; ?G \ ?c") +proof - + let ?ll = "(if l1 = l2 then L_true else L_false)" + have "wfC \ \ ?G ?c" proof(rule wfC_e_eq2) + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI wfV_litI assms base_for_lit.simps by metis + show "\ ; \ ; \ \\<^sub>w\<^sub>f CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e) : B_bool " + using wfCE_eqI wfCE_leqI wfCE_eqI wfV_litI wfCE_valI base_for_lit.simps assms by metis + show "\\<^sub>w\<^sub>f \" using assms wfX_wfY by auto + show "atom x \ \" using assms by auto + qed + + moreover have "\i. wfI \ ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume "wfI \ ?G i \ is_satis_g i ?G" + + hence "is_satis i ((CE_val (V_var x) == CE_val (V_lit (?ll)) ))" by auto + thus "is_satis i ((CE_val (V_var x) == CE_op Eq ([V_lit (l1)]\<^sup>c\<^sup>e) ([V_lit (l2)]\<^sup>c\<^sup>e)))" + using is_satis_eq_imp assms by auto + qed + ultimately show ?thesis using valid.simps by metis +qed + +lemma valid_fst: + fixes x::x and v\<^sub>1::v and v\<^sub>2::v + assumes "wfTh \" and "wfV \ \ GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)" + shows "\ ; \ ; (x, b\<^sub>1, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e" +proof(rule valid_eq_e) + show \\i s1 s2. (\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ v\<^sub>1 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) \ s1 = s2\ + proof(rule+) + fix i s1 s2 + assume as:"\ ; \ \\<^sub>w\<^sub>f GNil \ \ ; GNil \ i \ (i \ [ v\<^sub>1 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2)" + then obtain s2' where *:"i \ [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \ ~ SPair s2 s2'" + using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims + by meson + then have " i \ v\<^sub>1 \ ~ s2" using eval_v_elims(3)[OF *] by auto + then show "s1 = s2" using eval_v_uniqueness as + using eval_e_uniqueness eval_e_valI by blast + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ v\<^sub>1 ]\<^sup>c\<^sup>e : b\<^sub>1 \ using assms + by (metis b.eq_iff(4) wfV_elims(3) wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [#1[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>1 \ using assms using wfCE_fstI + using wfCE_valI by blast +qed + +lemma valid_snd: + fixes x::x and v\<^sub>1::v and v\<^sub>2::v + assumes "wfTh \" and "wfV \ \ GNil (V_pair v\<^sub>1 v\<^sub>2) (B_pair b\<^sub>1 b\<^sub>2)" + shows "\ ; \ ; (x, b\<^sub>2, [[x]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>2]\<^sup>c\<^sup>e) #\<^sub>\ GNil \ [[x]\<^sup>v]\<^sup>c\<^sup>e == [#2[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e" +proof(rule valid_eq_e) + show \\i s1 s2. (\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ v\<^sub>2 ]\<^sup>c\<^sup>e \ ~ s1) \ +(i \ [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) \ s1 = s2\ + proof(rule+) + fix i s1 s2 + assume as:"\ ; \ \\<^sub>w\<^sub>f GNil \ \ ; GNil \ i \ (i \ [ v\<^sub>2 ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2)" + then obtain s2' where *:"i \ [ v\<^sub>1 , v\<^sub>2 ]\<^sup>v \ ~ SPair s2' s2" + using eval_e_elims(5)[of i "[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e" s2] eval_e_elims + by meson + then have " i \ v\<^sub>2 \ ~ s2" using eval_v_elims(3)[OF *] by auto + then show "s1 = s2" using eval_v_uniqueness as + using eval_e_uniqueness eval_e_valI by blast + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ v\<^sub>2 ]\<^sup>c\<^sup>e : b\<^sub>2 \ using assms + by (metis b.eq_iff wfV_elims wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [#2[[ v\<^sub>1 , v\<^sub>2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b\<^sub>2 \ using assms using wfCE_sndI wfCE_valI by blast +qed + +lemma valid_concat: + fixes v1::"bit list" and v2::"bit list" + assumes " \\<^sub>w\<^sub>f \" + shows "\ ; \ ; (x, B_bitvec, (CE_val (V_var x) == CE_val (V_lit (L_bitvec (v1@ v2))))) #\<^sub>\ GNil \ + (CE_val (V_var x) == CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e ) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e) )" +proof(rule valid_eq_e) + show \\i s1 s2. ((\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ + (i \ [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1) \ (i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ s2) \ + s1 = s2)\ + proof(rule+) + fix i s1 s2 + assume as: "(\ ; \ \\<^sub>w\<^sub>f GNil) \ (\ ; GNil \ i) \ (i \ [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s1) \ + (i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2) " + + hence *: "i \ [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ s2" by auto + obtain bv1 bv2 where s2:"s2 = SBitvec (bv1 @ bv2) \ i \ [ L_bitvec v1 ]\<^sup>v \ ~ SBitvec bv1 \ (i \ [ L_bitvec v2 ]\<^sup>v \ ~ SBitvec bv2)" + using eval_e_elims(7)[OF *] eval_e_elims(1) by metis + hence "v1 = bv1 \ v2 = bv2" using eval_v_elims(1) eval_l.simps(5) by force + moreover then have "s1 = SBitvec (bv1 @ bv2)" using s2 using eval_v_elims(1) eval_l.simps(5) + by (metis as eval_e_elims(1)) + + then show "s1 = s2" using s2 by auto + qed + + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [ [ L_bitvec (v1 @ v2) ]\<^sup>v ]\<^sup>c\<^sup>e : B_bitvec \ + by (metis assms base_for_lit.simps(5) wfG_nilI wfV_litI wfV_wfCE) + show \ \ ; \ ; GNil \\<^sub>w\<^sub>f [[[ L_bitvec v1 ]\<^sup>v]\<^sup>c\<^sup>e @@ [[ L_bitvec v2 ]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : B_bitvec \ + by (metis assms base_for_lit.simps(5) wfCE_concatI wfG_nilI wfV_litI wfCE_valI) +qed + +lemma valid_ce_eq: + fixes ce::ce + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" + shows \\ ; \ ; \ \ ce == ce \ + unfolding valid.simps proof + show \ \ ; \ ; \ \\<^sub>w\<^sub>f ce == ce \ using assms wfC_eqI by auto + show \\i. \ ; \ \ i \ i \ \ \ i \ ce == ce \ proof(rule+) + fix i + assume "\ ; \ \ i \ i \ \" + then obtain s where "i\ ce \ ~ s" using assms eval_e_exist by metis + then show "i \ ce == ce \ ~ True " using eval_c_eqI by metis + qed +qed + +lemma valid_eq_imp: + fixes c1::c and c2::c + assumes " \ ; \ ; (x, b, c2) #\<^sub>\ \ \\<^sub>w\<^sub>f c1 IMP c2" + shows " \ ; \ ; (x, b, c2) #\<^sub>\ \ \ c1 IMP c2 " +proof - + have "\i. (\ ; (x, b, c2) #\<^sub>\ \ \ i \ i \ (x, b, c2) #\<^sub>\ \) \ i \ ( c1 IMP c2 )" + proof(rule,rule) + fix i + assume as:"\ ; (x, b, c2) #\<^sub>\ \ \ i \ i \ (x, b, c2) #\<^sub>\ \" + + have "\ ; \ ; (x, b, c2) #\<^sub>\ \ \\<^sub>w\<^sub>f c1" using wfC_elims assms by metis + + then obtain sc where "i \ c1 \ ~ sc" using eval_c_exist assms as by metis + moreover have "i \ c2 \ ~ True" using as is_satis_g.simps is_satis.simps by auto + + ultimately have "i \ c1 IMP c2 \ ~ True" using eval_c_impI by metis + + thus "i \ c1 IMP c2" using is_satis.simps by auto + qed + thus ?thesis using assms by auto +qed + +lemma valid_range: + assumes "0 \ n \ n \ m" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ GNil \ + (C_eq (CE_op LEq (CE_val (V_var x)) (CE_val (V_lit (L_num m)))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)" + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") +proof(rule validI) + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil " + using assms base_for_lit.simps wfG_nilI wfV_litI fresh_GNil wfB_intI wfC_v_eq wfG_cons1I wfG_cons2I by metis + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI + by metis + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence *:"i \ V_var x \ ~ SNum n" + proof - + obtain sv where sv: "i x = Some sv \ \ \ sv : B_int" using a wfI_def by force + have "i \ (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \ ~ True" + using a is_satis_g.simps + using is_satis.cases by blast + hence "i x = Some(SNum n)" using sv + by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2)) + thus ?thesis using eval_v_varI by auto + qed + + show "i \ ?c1 AND ?c2" + proof - + have "i \ ?c1 \ ~ True" + proof - + have "i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num m ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + + moreover have "i \ ?c2 \ ~ True" + proof - + have "i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + ultimately show ?thesis using eval_c_conjI is_satis.simps by metis + qed + qed +qed + +lemma valid_range_length: + fixes \::\ + assumes "0 \ n \ n \ int (length v)" and " \ ; {||} \\<^sub>w\<^sub>f \" and "atom x \ \" + shows "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ \ \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") +proof(rule validI) + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (x, B_int, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ \ " apply(rule wfG_cons1I) + apply simp + using assms apply simp+ + using assms base_for_lit.simps wfG_nilI wfV_litI wfB_intI wfC_v_eq wfB_intI wfX_wfY assms by metis+ + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + using wfC_conjI wfC_eqI wfCE_leqI wfCE_valI wfV_varI wfg lookup.simps base_for_lit.simps wfV_litI wfB_intI wfB_boolI + by (metis (full_types) wfCE_lenI) + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence *:"i \ V_var x \ ~ SNum n" + proof - + obtain sv where sv: "i x = Some sv \ \ \ sv : B_int" using a wfI_def by force + have "i \ (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n)))) \ ~ True" + using a is_satis_g.simps + using is_satis.cases by blast + hence "i x = Some(SNum n)" using sv + by (metis eval_c_elims(7) eval_e_elims(1) eval_l.simps(3) eval_v_elims(1) eval_v_elims(2)) + thus ?thesis using eval_v_varI by auto + qed + + show "i \ ?c1 AND ?c2" + proof - + have "i \ ?c2 \ ~ True" + proof - + have "i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e\ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_lenI eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + + moreover have "i \ ?c1 \ ~ True" + proof - + have "i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_e_leqI assms eval_v_litI eval_l.simps * + by (metis (full_types) eval_e_valI) + moreover have "i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBool True" + using eval_v_litI eval_e_valI eval_l.simps by metis + ultimately show ?thesis using eval_c_eqI by metis + qed + ultimately show ?thesis using eval_c_conjI is_satis.simps by metis + qed + qed +qed + +lemma valid_range_length_inv_gnil: + fixes \::\ + assumes "\\<^sub>w\<^sub>f \ " + and "\ ; {||} ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ GNil \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") + shows "0 \ n \ n \ int (length v)" +proof - + have *:"\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" using assms valid.simps by simp + + obtain i where i: "i x = Some (SNum n)" by auto + have "\ ; ?G \ i \ i \ ?G" proof + show "\ ; ?G \ i" unfolding wfI_def using wfRCV_BIntI i * by auto + have "i \ ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \ ~ True" + using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps + by (metis (full_types) i) + thus "i \ ?G" unfolding is_satis_g.simps is_satis.simps by auto + qed + hence **:"i \ ?c1 AND ?c2" using * by auto + + hence 1: "i \ ?c1 \ ~ True" using eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where "(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where "SBool True = SBool (n1 \ n2) \ (i \ [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"] + using \(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2\ \sv1 = SBool True\ by fastforce + moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i + apply (metis eval_l.simps(3) rcl_val.eq_iff(2)) + using eval_e_elims eval_v_elims i + by (metis calculation option.inject rcl_val.eq_iff(2)) + ultimately have le1: "0 \ n " by simp + + hence 2: "i \ ?c2 \ ~ True" using ** eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \ i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where ***:"SBool True = SBool (n1 \ n2) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3) + using sv \sv1 = SBool True\ by metis + moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto + moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i + by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2)) + ultimately have le2: "n \ int (length v) " by simp + + show ?thesis using le1 le2 by auto +qed + +lemma wfI_cons: + fixes i::valuation and \::\ + assumes "i' \ \" and "\ ; \ \ i'" and "i = i' ( x \ s)" and "\ \ s : b" and "atom x \ \" + shows "\ ; (x,b,c) #\<^sub>\ \ \ i" + unfolding wfI_def proof - + { + fix x' b' c' + assume "(x',b',c') \ toSet ((x, b, c) #\<^sub>\ \)" + then consider "(x',b',c') = (x,b,c)" | "(x',b',c') \ toSet \" using toSet.simps by auto + then have "\s. Some s = i x' \ \ \ s : b'" proof(cases) + case 1 + then show ?thesis using assms by auto + next + case 2 + then obtain s where s:"Some s = i' x' \ \ \ s : b'" using assms wfI_def by auto + moreover have "x' \ x" using assms 2 fresh_dom_free by auto + ultimately have "Some s = i x'" using assms by auto + then show ?thesis using s wfI_def by auto + qed + } + thus "\(x, b, c)\toSet ((x, b, c) #\<^sub>\ \). \s. Some s = i x \ \ \ s : b" by auto +qed + +lemma valid_range_length_inv: + fixes \::\ + assumes "\ ; B \\<^sub>w\<^sub>f \ " and "atom x \ \" and "\i. i \ \ \ \ ; \ \ i" + and "\ ; B ; (x, B_int , (C_eq (CE_val (V_var x)) (CE_val (V_lit (L_num n))))) #\<^sub>\ \ \ + (C_eq (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var x))) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + (C_eq (CE_op LEq (CE_val (V_var x)) ([| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e )) [[ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + " + (is "\ ; ?B ; ?G \ ?c1 AND ?c2") + shows "0 \ n \ n \ int (length v)" +proof - + have *:"\i. \ ; ?G \ i \ i \ ?G \ i \ ?c1 AND ?c2" using assms valid.simps by simp + + obtain i' where idash: "is_satis_g i' \ \ \ ; \ \ i'" using assms by auto + obtain i where i: "i = i' ( x \ SNum n)" by auto + hence ix: "i x = Some (SNum n)" by auto + have "\ ; ?G \ i \ i \ ?G" proof + show "\ ; ?G \ i" using wfI_cons i idash ix wfRCV_BIntI assms by simp + + have **:"i \ ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) \ ~ True" + using * eval_c.intros(7) eval_e.intros eval_v.intros eval_l.simps i + by (metis (full_types) ix) + + show "i \ ?G" unfolding is_satis_g.simps proof + show \ i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ using ** is_satis.simps by auto + show \ i \ \ \ using idash i assms is_satis_g_i_upd by metis + qed + qed + hence **:"i \ ?c1 AND ?c2" using * by auto + + hence 1: "i \ ?c1 \ ~ True" using eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where "(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where "SBool True = SBool (n1 \ n2) \ (i \ [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3)[of i " [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e" "[ [ x ]\<^sup>v ]\<^sup>c\<^sup>e " "SBool True"] + using \(sv1 = sv2) = True \ i \ [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2\ \sv1 = SBool True\ by fastforce + moreover hence "n1 = 0" and "n2 = n" using eval_e_elims eval_v_elims i + apply (metis eval_l.simps(3) rcl_val.eq_iff(2)) + using eval_e_elims eval_v_elims i + calculation option.inject rcl_val.eq_iff(2) + by (metis ix) + ultimately have le1: "0 \ n " by simp + + hence 2: "i \ ?c2 \ ~ True" using ** eval_c_elims(3) is_satis.simps + by fastforce + then obtain sv1 and sv2 where sv: "(sv1 = sv2) = True \ i \ [ leq [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ sv1 \ i \ [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv2" + using eval_c_elims(7) by metis + hence "sv1 = SBool True" using eval_e_elims eval_v_elims eval_l.simps i by metis + obtain n1 and n2 where ***:"SBool True = SBool (n1 \ n2) \ (i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n1) \ (i \ [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum n2)" + using eval_e_elims(3) + using sv \sv1 = SBool True\ by metis + moreover hence "n1 = n" using eval_e_elims(1)[of i] eval_v_elims(2)[of i x "SNum n1"] i by auto + moreover have "n2 = int (length v)" using eval_e_elims(8) eval_v_elims(1) eval_l.simps i + by (metis "***" eval_e_elims(1) rcl_val.eq_iff(1) rcl_val.eq_iff(2)) + ultimately have le2: "n \ int (length v) " by simp + + show ?thesis using le1 le2 by auto +qed + +lemma eval_c_conj2I[intro]: + assumes "i \ c1 \ ~ True" and "i \ c2 \ ~ True" + shows "i \ (C_conj c1 c2) \ ~ True" + using assms eval_c_conjI by metis + +lemma valid_split: + assumes "split n v (v1,v2)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; (z , [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\ GNil +\ ([ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) AND ([| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e)" + (is "\ ; {||} ; ?G \ ?c1 AND ?c2") + unfolding valid.simps proof + + have wfg: " \ ; {||} \\<^sub>w\<^sub>f (z, [B_bitvec , B_bitvec ]\<^sup>b , [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e) #\<^sub>\ GNil" + using wf_intros assms base_for_lit.simps fresh_GNil wfC_v_eq wfG_intros2 by metis + + show "\ ; {||} ; ?G \\<^sub>w\<^sub>f ?c1 AND ?c2" + apply(rule wfC_conjI) + apply(rule wfC_eqI) + apply(rule wfCE_valI) + apply(rule wfV_litI) + using wf_intros wfg lookup.simps base_for_lit.simps wfC_v_eq + apply (metis )+ + done + + have len:"int (length v1) = n" using assms split_length by auto + + show "\i. \ ; ?G \ i \ i \ ?G \ i \ (?c1 AND ?c2)" + proof(rule,rule) + fix i + assume a:"\ ; ?G \ i \ i \ ?G" + hence "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ True" + using is_satis_g.simps is_satis.simps by simp + then obtain sv where "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv \ i \ [ [ [ L_bitvec v1 ]\<^sup>v , [ L_bitvec v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ sv" + using eval_c_elims by metis + hence "i \ [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ (SPair (SBitvec v1) (SBitvec v2))" using eval_c_eqI eval_v.intros eval_l.simps + by (metis eval_e_elims(1) eval_v_uniqueness) + hence b:"i z = Some (SPair (SBitvec v1) (SBitvec v2))" using a eval_e_elims eval_v_elims by metis + + have v1: "i \ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ SBitvec v1 " + using eval_e_fstI eval_e_valI eval_v_varI b by metis + have v2: "i \ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \ ~ SBitvec v2" + using eval_e_sndI eval_e_valI eval_v_varI b by metis + + have "i \ [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SBitvec v" using eval_e.intros eval_v.intros eval_l.simps by metis + moreover have "i \ [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ ~ SBitvec v" + using assms split_concat v1 v2 eval_e_concatI by metis + moreover have "i \ [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e \ ~ SNum (int (length v1))" + using v1 eval_e_lenI by auto + moreover have "i \ [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ SNum n" using eval_e.intros eval_v.intros eval_l.simps by metis + ultimately show "i \ ?c1 AND ?c2" using is_satis.intros eval_c_conj2I eval_c_eqI len by metis + qed +qed + + +lemma is_satis_eq: + assumes "wfI \ G i" and "wfCE \ \ G e b" + shows "is_satis i (e == e)" +proof(rule) + obtain s where "eval_e i e s" using eval_e_exist assms by metis + thus "eval_c i (e == e ) True" using eval_c_eqI by metis +qed + +lemma is_satis_g_i_upd2: + assumes "eval_v i v s" and "is_satis ((i ( x \ s))) c0" and "atom x \ G" and "wfG \ \ (G3@((x,b,c0)#\<^sub>\G))" and "wfV \ \ G v b" and "wfI \ (G3[x::=v]\<^sub>\\<^sub>v@G) i" + and "is_satis_g i (G3[x::=v]\<^sub>\\<^sub>v@G)" + shows "is_satis_g (i ( x \ s)) (G3@((x,b,c0)#\<^sub>\G))" + using assms proof(induct G3 rule: \_induct) + case GNil + hence "is_satis_g (i(x \ s)) G" using is_satis_g_i_upd by auto + then show ?case using GNil using is_satis_g.simps append_g.simps by metis +next + case (GCons x' b' c' \') + hence "x\x'" using wfG_cons_append by metis + hence "is_satis_g i (((x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (\'[x::=v]\<^sub>\\<^sub>v) @ G))" using subst_gv.simps GCons by auto + hence *:"is_satis i c'[x::=v]\<^sub>c\<^sub>v \ is_satis_g i ((\'[x::=v]\<^sub>\\<^sub>v) @ G)" using subst_gv.simps by auto + + have "is_satis_g (i(x \ s)) ((x', b', c') #\<^sub>\ (\'@ (x, b, c0) #\<^sub>\ G))" proof(subst is_satis_g.simps,rule) + show "is_satis (i(x \ s)) c'" proof(subst subst_c_satis_full[symmetric]) + show \eval_v i v s\ using GCons by auto + show \ \ ; \ ; ((x', b', c') #\<^sub>\\')@(x, b, c0) #\<^sub>\ G \\<^sub>w\<^sub>f c' \ using GCons wfC_refl by auto + show \wfI \ ((((x', b', c') #\<^sub>\ \')[x::=v]\<^sub>\\<^sub>v) @ G) i\ using GCons by auto + show \ \ ; \ ; G \\<^sub>w\<^sub>f v : b \ using GCons by auto + show \is_satis i c'[x::=v]\<^sub>c\<^sub>v\ using * by auto + qed + show "is_satis_g (i(x \ s)) (\' @ (x, b, c0) #\<^sub>\ G)" proof(rule GCons(1)) + show \eval_v i v s\ using GCons by auto + show \is_satis (i(x \ s)) c0\ using GCons by metis + show \atom x \ G\ using GCons by auto + show \ \ ; \\\<^sub>w\<^sub>f \' @ (x, b, c0) #\<^sub>\ G \ using GCons wfG_elims append_g.simps by metis + show \is_satis_g i (\'[x::=v]\<^sub>\\<^sub>v @ G)\ using * by auto + show "wfI \ (\'[x::=v]\<^sub>\\<^sub>v @ G) i" using GCons wfI_def subst_g_assoc_cons \x\x'\ by auto + show "\ ; \ ; G \\<^sub>w\<^sub>f v : b " using GCons by auto + qed + qed + moreover have "((x', b', c') #\<^sub>\ \' @ (x, b, c0) #\<^sub>\ G) = (((x', b', c') #\<^sub>\ \') @ (x, b, c0) #\<^sub>\ G)" by auto + ultimately show ?case using GCons by metis +qed + + +end \ No newline at end of file diff --git a/thys/MiniSail/ROOT b/thys/MiniSail/ROOT new file mode 100644 --- /dev/null +++ b/thys/MiniSail/ROOT @@ -0,0 +1,15 @@ +chapter AFP + +session "MiniSail" (AFP) = "HOL-Library" + + description \Formalisation of MiniSail\ + options [timeout = 1200] + sessions + "Nominal2" + "HOL-Eisbach" + "Native_Word" + "Show" + theories + "MiniSail" + document_files + "root.tex" + "root.bib" diff --git a/thys/MiniSail/Safety.thy b/thys/MiniSail/Safety.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Safety.thy @@ -0,0 +1,1812 @@ +(*<*) +theory Safety + imports Operational IVSubstTypingL BTVSubstTypingL +begin + (*>*) + +method supp_calc = (metis (mono_tags, hide_lams) pure_supp c.supp e.supp v.supp supp_l_empty opp.supp sup_bot.right_neutral supp_at_base) +declare infer_e.intros[simp] +declare infer_e.intros[intro] + +chapter \Safety\ + +text \Lemmas about the operational semantics leading up to progress and preservation and then +safety.\ + +section \Store Lemmas\ + +abbreviation delta_ext (" _ \ _ ") where + "delta_ext \ \' \ (setD \ \ setD \')" + +nominal_function dc_of :: "branch_s \ string" where + "dc_of (AS_branch dc _ _) = dc" + apply(auto,simp add: eqvt_def dc_of_graph_aux_def) + using s_branch_s_branch_list.exhaust by metis +nominal_termination (eqvt) by lexicographic_order + +lemma delta_sim_fresh: + assumes "\ \ \ \ \" and "atom u \ \" + shows "atom u \ \" + using assms proof(induct rule : delta_sim.inducts) + case (delta_sim_nilI \) + then show ?case using fresh_def supp_DNil by blast +next + case (delta_sim_consI \ \ \ v \ u') + hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using check_v_wf by meson + hence "supp \ = {}" using wfT_supp by fastforce + moreover have "atom u \ u'" using delta_sim_consI fresh_Cons fresh_Pair by blast + moreover have "atom u \ \" using delta_sim_consI fresh_Cons by blast + ultimately show ?case using fresh_Pair fresh_DCons fresh_def by blast +qed + +lemma delta_sim_v: + fixes \::\ + assumes "\ \ \ \ \ " and "(u,v) \ set \" and "(u,\) \ setD \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; GNil \ v \ \" + using assms proof(induct \ arbitrary: \) + case Nil + then show ?case by auto +next + case (Cons uv \) + obtain u' and v' where uv : "uv=(u',v')" by fastforce + show ?case proof(cases "u'=u") + case True + hence *:"\ \ ((u,v')#\) \ \" using uv Cons by blast + then obtain \' and \' where tt: "\ ; {||} ; GNil \ v' \ \' \ u \ fst ` set \ \ \ = (u,\')#\<^sub>\\'" using delta_sim_elims(3)[OF *] by metis + moreover hence "v'=v" using Cons True + by (metis Pair_inject fst_conv image_eqI set_ConsD uv) + moreover have "\=\'" using wfD_unique tt Cons + setD.simps list.set_intros by blast + ultimately show ?thesis by metis + next + case False + hence *:"\ \ ((u',v')#\) \ \" using uv Cons by blast + then obtain \' and \' where tt: "\ \ \ \ \' \ \ ; {||} ; GNil \ v' \ \' \ u' \ fst ` set \ \ \ = (u',\')#\<^sub>\\'" using delta_sim_elims(3)[OF *] by metis + + moreover hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using wfD_elims Cons delta_sim_elims by metis + ultimately show ?thesis using Cons + using False by auto + qed +qed + +lemma delta_sim_delta_lookup: + assumes "\ \ \ \ \ " and "(u, \ z : b | c \) \ setD \" + shows "\v. (u,v) \ set \" + using assms by(induct rule: delta_sim.inducts,auto+) + +lemma update_d_stable: + "fst ` set \ = fst ` set (update_d \ u v)" +proof(induct \) + case Nil + then show ?case by auto +next + case (Cons a \) + then show ?case using update_d.simps + by (metis (no_types, lifting) eq_fst_iff image_cong image_insert list.simps(15) prod.exhaust_sel) +qed + +lemma update_d_sim: + fixes \::\ + assumes "\ \ \ \ \" and "\ ; {||} ; GNil \ v \ \" and "(u,\) \ setD \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" + shows "\ \ (update_d \ u v) \ \" + using assms proof(induct \ arbitrary: \) + case Nil + then show ?case using delta_sim_consI by simp +next + case (Cons uv \) + obtain u' and v' where uv : "uv=(u',v')" by fastforce + + hence *:"\ \ ((u',v')#\) \ \" using uv Cons by blast + then obtain \' and \' where tt: "\ \ \ \ \' \ \ ; {||} ; GNil \ v' \ \' \ u' \ fst ` set \ \ \ = (u',\')#\<^sub>\\'" using delta_sim_elims * by metis + + show ?case proof(cases "u=u'") + case True + then have "(u,\') \ setD \" using tt by auto + then have "\ = \'" using Cons wfD_unique by metis + moreover have "update_d ((u',v')#\) u v = ((u',v)#\)" using update_d.simps True by presburger + ultimately show ?thesis using delta_sim_consI tt Cons True + by (simp add: tt uv) + next + case False + have "\ \ (u',v') # (update_d \ u v) \ (u',\')#\<^sub>\\'" + proof(rule delta_sim_consI) + show "\ \ update_d \ u v \ \' " using Cons using delta_sim_consI + delta_sim.simps update_d.simps Cons delta_sim_elims uv tt + False fst_conv set_ConsD wfG_elims wfD_elims by (metis setD_ConsD) + show "\ ; {||} ; GNil \ v' \ \'" using tt by auto + show "u' \ fst ` set (update_d \ u v)" using update_d.simps Cons update_d_stable tt by auto + qed + thus ?thesis using False update_d.simps uv + by (simp add: tt) + qed +qed + +section \Preservation\ +text \Types are preserved under reduction step. Broken down into lemmas about different operations\ + +subsection \Function Application\ + +lemma check_s_x_fresh: + fixes x::x and s::s + assumes "\ ; \ ; B ; GNil ; D \ s \ \" + shows "atom x \ s \ atom x \ \ \ atom x \ D" +proof - + have "\ ; \ ; B ; GNil ; D \\<^sub>w\<^sub>f s : b_of \" using check_s_wf[OF assms] by auto + moreover have "\ ; B ; GNil \\<^sub>w\<^sub>f \ " using check_s_wf assms by auto + moreover have "\ ; B ; GNil \\<^sub>w\<^sub>f D" using check_s_wf assms by auto + ultimately show ?thesis using wf_supp x_fresh_u + by (meson fresh_GNil wfS_x_fresh wfT_x_fresh wfD_x_fresh) +qed + +lemma check_funtyp_subst_b: + fixes b'::b + assumes "check_funtyp \ \ {|bv|} (AF_fun_typ x b c \ s)" and \ \ ; {||} \\<^sub>w\<^sub>f b' \ + shows "check_funtyp \ \ {||} (AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \[bv::=b']\<^sub>\\<^sub>b s[bv::=b']\<^sub>s\<^sub>b)" + using assms proof (nominal_induct "{|bv|}" "AF_fun_typ x b c \ s" rule: check_funtyp.strong_induct) + case (check_funtypI x' \ \ c' s' \') + have "check_funtyp \ \ {||} (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \'[bv::=b']\<^sub>\\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" proof + show \atom x' \ (\, \, {||}::bv fset, b[bv::=b']\<^sub>b\<^sub>b)\ using check_funtypI fresh_prodN x_fresh_b fresh_empty_fset by metis + + have \ \ ; \ ; {||} ; ((x', b, c') #\<^sub>\ GNil)[bv::=b']\<^sub>\\<^sub>b ; []\<^sub>\[bv::=b']\<^sub>\\<^sub>b \ s'[bv::=b']\<^sub>s\<^sub>b \ \'[bv::=b']\<^sub>\\<^sub>b\ proof(rule subst_b_check_s) + show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using check_funtypI by metis + show \{|bv|} = {|bv|}\ by auto + show \ \ ; \ ; {|bv|} ; (x', b, c') #\<^sub>\ GNil ; []\<^sub>\ \ s' \ \'\ using check_funtypI by metis + qed + + thus \ \ ; \ ; {||} ; (x', b[bv::=b']\<^sub>b\<^sub>b, c'[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ GNil ; []\<^sub>\ \ s'[bv::=b']\<^sub>s\<^sub>b \ \'[bv::=b']\<^sub>\\<^sub>b\ + using subst_gb.simps subst_db.simps by simp + qed + + moreover have "(AF_fun_typ x b c \ s) = (AF_fun_typ x' b c' \' s')" using fun_typ.eq_iff check_funtypI by metis + moreover hence "(AF_fun_typ x b[bv::=b']\<^sub>b\<^sub>b (c[bv::=b']\<^sub>c\<^sub>b) \[bv::=b']\<^sub>\\<^sub>b s[bv::=b']\<^sub>s\<^sub>b) = (AF_fun_typ x' b[bv::=b']\<^sub>b\<^sub>b (c'[bv::=b']\<^sub>c\<^sub>b) \'[bv::=b']\<^sub>\\<^sub>b s'[bv::=b']\<^sub>s\<^sub>b)" + using subst_ft_b.simps by metis + ultimately show ?case by metis +qed + +lemma funtyp_simple_check: + fixes s::s and \::\ and \::\ and v::v + assumes "check_funtyp \ \ ({||}::bv fset) (AF_fun_typ x b c \ s)" and + "\ ; {||} ; GNil \ v \ \ x : b | c \" + shows "\ ; \ ; {||} ; GNil ; DNil \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct " ({||}::bv fset)" "AF_fun_typ x b c \ s" avoiding: v x rule: check_funtyp.strong_induct) + case (check_funtypI x' \ \ c' s' \') + + hence eq1: "\ x' : b | c' \ = \ x : b | c \" using funtyp_eq_iff_equalities by metis + + obtain x'' and c'' where xf:"\ x : b | c \ = \ x'' : b | c'' \ \ atom x'' \ (x',v) \ atom x'' \ (x,c)" using obtain_fresh_z3 by metis + moreover have "atom x' \ c''" proof - + have "supp \ x'' : b | c'' \ = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis + hence "supp c'' \ { atom x'' }" using \.supp eq1 xf by (auto simp add: freshers) + moreover have "atom x' \ atom x''" using xf fresh_Pair fresh_x_neq by metis + ultimately show ?thesis using xf fresh_Pair fresh_x_neq fresh_def fresh_at_base by blast + qed + ultimately have eq2: "c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c'" using eq1 type_eq_subst_eq3(1)[of x' b c' x'' b c''] by metis + + have "atom x' \ c" proof - + have "supp \ x : b | c \ = {}" using eq1 check_funtypI xf check_v_wf wfT_nil_supp by metis + hence "supp c \ { atom x }" using \.supp by auto + moreover have "atom x \ atom x'" using check_funtypI fresh_Pair fresh_x_neq by metis + ultimately show ?thesis using fresh_def by force + qed + hence eq: " c[x::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v = c' \ s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>v" + using funtyp_eq_iff_equalities type_eq_subst_eq3 check_funtypI by metis + + have " \ ; \ ; {||} ; ((x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil)[x'::=v]\<^sub>\\<^sub>v ; []\<^sub>\[x'::=v]\<^sub>\\<^sub>v \ s'[x'::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v" + proof(rule subst_check_check_s ) + show \\ ; {||} ; GNil \ v \ \ x'' : b | c'' \\ using check_funtypI eq1 xf by metis + show \atom x'' \ (x', v)\ using check_funtypI fresh_x_neq fresh_Pair xf by metis + show \ \ ; \ ; {||} ; (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; []\<^sub>\ \ s' \ \'\ using check_funtypI eq2 by metis + show \ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil = GNil @ (x', b, c''[x''::=[ x' ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil\ using append_g.simps by auto + qed + hence " \; \; {||}; GNil; []\<^sub>\ \ s'[x'::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v" using subst_gv.simps subst_dv.simps by auto + thus ?case using eq by auto +qed + +lemma funtypq_simple_check: + fixes s::s and \::\ and \::\ and v::v + assumes "check_funtypq \ \ (AF_fun_typ_none (AF_fun_typ x b c t s))" and + "\ ; {||} ; GNil \ v \ \ x : b | c \" + shows "\; \; {||}; GNil; DNil \ s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "(AF_fun_typ_none (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct) + case (check_fundefq_simpleI \ \ x' c' t' s') + hence eq: "\ x : b | c \ = \ x' : b | c' \ \ s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v = t'[x'::=v]\<^sub>\\<^sub>v" + using funtyp_eq_iff_equalities by metis + hence "\; \; {||}; GNil; []\<^sub>\ \ s'[x'::=v]\<^sub>s\<^sub>v \ t'[x'::=v]\<^sub>\\<^sub>v" + using funtyp_simple_check[OF check_fundefq_simpleI(1)] check_fundefq_simpleI by metis + thus ?case using eq by metis +qed + +lemma funtyp_poly_eq_iff_equalities: + assumes "[[atom bv']]lst. AF_fun_typ x' b'' c' t' s' = [[atom bv]]lst. AF_fun_typ x b c t s" + shows "\ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \ = \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ + s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v = t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" +proof - + have "subst_ft_b (AF_fun_typ x' b'' c' t' s') bv' b' = subst_ft_b (AF_fun_typ x b c t s) bv b'" + using subst_b_flip_eq_two subst_b_fun_typ_def assms by metis + thus ?thesis using fun_typ.eq_iff subst_ft_b.simps funtyp_eq_iff_equalities subst_tb.simps + by (metis (full_types) assms fun_poly_arg_unique) + +qed + +lemma funtypq_poly_check: + fixes s::s and \::\ and \::\ and v::v and b'::b + assumes "check_funtypq \ \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" and + "\ ; {||} ; GNil \ v \ \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" and + "\ ; {||} \\<^sub>w\<^sub>f b'" + shows "\; \; {||}; GNil; DNil \ s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "(AF_fun_typ_some bv (AF_fun_typ x b c t s))" avoiding: v rule: check_funtypq.strong_induct) + case (check_funtypq_polyI bv' \ \ x' b'' c' t' s') + + hence **:"\ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \ = \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \ \ + s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v = s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v = t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" + using funtyp_poly_eq_iff_equalities by metis + + have *:"check_funtyp \ \ {||} (AF_fun_typ x' b''[bv'::=b']\<^sub>b\<^sub>b (c'[bv'::=b']\<^sub>c\<^sub>b) (t'[bv'::=b']\<^sub>\\<^sub>b) s'[bv'::=b']\<^sub>s\<^sub>b)" + using check_funtyp_subst_b[OF check_funtypq_polyI(5) check_funtypq_polyI(8)] by metis + moreover have "\ ; {||} ; GNil \ v \ \ x' : b''[bv'::=b']\<^sub>b\<^sub>b | c'[bv'::=b']\<^sub>c\<^sub>b \" using ** check_funtypq_polyI by metis + ultimately have "\; \; {||}; GNil; []\<^sub>\ \ s'[bv'::=b']\<^sub>s\<^sub>b[x'::=v]\<^sub>s\<^sub>v \ t'[bv'::=b']\<^sub>\\<^sub>b[x'::=v]\<^sub>\\<^sub>v" + using funtyp_simple_check[OF *] check_funtypq_polyI by metis + thus ?case using ** by metis + +qed + +lemma fundef_simple_check: + fixes s::s and \::\ and \::\ and v::v + assumes "check_fundef \ \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" and + "\ ; {||} ; GNil \ v \ \ x : b | c \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" + shows "\; \; {||}; GNil; \ \ s[x::=v]\<^sub>s\<^sub>v \ t[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct) + case (check_fundefI \ \) + then show ?case using funtypq_simple_check[THEN check_s_d_weakening(1) ] setD.simps by auto +qed + +lemma fundef_poly_check: + fixes s::s and \::\ and \::\ and v::v and b'::b + assumes "check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" and + "\ ; {||} ; GNil \ v \ \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\ ; {||} \\<^sub>w\<^sub>f b'" + shows "\; \; {||}; GNil; \ \ s[bv::=b']\<^sub>s\<^sub>b[x::=v]\<^sub>s\<^sub>v \ t[bv::=b']\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v" + using assms proof(nominal_induct "(AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c t s)))" avoiding: v rule: check_fundef.strong_induct) + case (check_fundefI \ \) + then show ?case using funtypq_poly_check[THEN check_s_d_weakening(1) ] setD.simps by auto +qed + +lemma preservation_app: + assumes + "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "(\fd\set \. check_fundef \ \ fd)" + shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ ss = LET x = (AE_app f v) IN s \ + \; \; {||}; GNil; \ \ LET x : (\1'[x1::=v]\<^sub>\\<^sub>v) = (s1'[x1::=v]\<^sub>s\<^sub>v) IN s \ \" and + "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" + using assms proof(nominal_induct \ and \ and \ avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_letI x2 \ \ \ \ \ e \ z s2 b c) + + hence eq: "e = (AE_app f v)" by simp + hence *:"\ ; \ ; {||} ;GNil ; \ \ (AE_app f v) \ \ z : b | c \" using check_letI by auto + + then obtain x3 b3 c3 \3 s3 where + **:"\ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \3 s3))) = lookup_fun \ f \ + \ ; {||} ; GNil \ v \ \ x3 : b3 | c3 \ \ atom x3 \ (\, \, ({||}::bv fset), GNil, \, v, \ z : b | c \) \ \3[x3::=v]\<^sub>\\<^sub>v = \ z : b | c \" + using infer_e_elims(6)[OF *] subst_defs by metis + + obtain z3 where z3:"\ x3 : b3 | c3 \ = \ z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \ \ atom z3 \ (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis + + have seq:"[[atom x3]]lst. s3 = [[atom x1]]lst. s1'" using fun_def_eq check_letI ** option.inject by metis + + let ?ft = "AF_fun_typ x3 b3 c3 \3 s3" + + + have sup: "supp \3 \ { atom x3} \ supp s3 \ { atom x3}" using wfPhi_f_supp ** by metis + + have "\; \; {||}; GNil; \ \ AS_let2 x2 \3[x3::=v]\<^sub>\\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 \ \" proof + show \atom x2 \ (\, \, {||}::bv fset, GNil, \, \3[x3::=v]\<^sub>\\<^sub>v, s3[x3::=v]\<^sub>s\<^sub>v, \)\ + unfolding fresh_prodN using check_letI fresh_subst_v_if subst_v_\_def sup + by (metis all_not_in_conv fresh_def fresh_empty_fset fresh_subst_sv_if fresh_subst_tv_if singleton_iff subset_singleton_iff) + + show \ \; \; {||}; GNil; \ \ s3[x3::=v]\<^sub>s\<^sub>v \ \3[x3::=v]\<^sub>\\<^sub>v\ proof(rule fundef_simple_check) + show \check_fundef \ \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x3 b3 c3 \3 s3)))\ using ** check_letI lookup_fun_member by metis + show \\ ; {||} ; GNil \ v \ \ x3 : b3 | c3 \\ using ** by auto + show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using ** by auto + qed + show \ \ ; \ ; {||} ; (x2, b_of \3[x3::=v]\<^sub>\\<^sub>v, c_of \3[x3::=v]\<^sub>\\<^sub>v x2) #\<^sub>\ GNil ; \ \ s2 \ \\ + using check_letI ** b_of.simps c_of.simps subst_defs by metis + qed + + moreover have "AS_let2 x2 \3[x3::=v]\<^sub>\\<^sub>v (s3[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\1'[x1::=v]\<^sub>\\<^sub>v) (s1'[x1::=v]\<^sub>s\<^sub>v) s" proof - + have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto + moreover have " \3[x3::=v]\<^sub>\\<^sub>v = \1'[x1::=v]\<^sub>\\<^sub>v" using fun_ret_unique ** check_letI by metis + moreover have "s3[x3::=v]\<^sub>s\<^sub>v = (s1'[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def seq by metis + ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis + qed + + ultimately show ?case using check_letI by auto +qed(auto+) + +lemma fresh_subst_v_subst_b: + fixes x2::x and tm::"'a::{has_subst_v,has_subst_b}" and x::x + assumes "supp tm \ { atom bv, atom x }" and "atom x2 \ v" + shows "atom x2 \ tm[bv::=b]\<^sub>b[x::=v]\<^sub>v" + using assms proof(cases "x2=x") + case True + then show ?thesis using fresh_subst_v_if assms by blast +next + case False + hence "atom x2 \ tm" using assms fresh_def fresh_at_base by force + hence "atom x2 \ tm[bv::=b]\<^sub>b" using assms fresh_subst_if x_fresh_b False by force + then show ?thesis using fresh_subst_v_if assms by auto +qed + +lemma preservation_poly_app: + assumes + "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "(\fd\set \. check_fundef \ \ fd)" + shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ ss = LET x = (AE_appP f b' v) IN s \ \ ; {||} \\<^sub>w\<^sub>f b' \ + \; \; {||}; GNil; \ \ LET x : (\1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v) = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) IN s \ \" and + "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" + using assms proof(nominal_induct \ and \ and \ avoiding: v x1 rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_letI x2 \ \ \ \ \ e \ z s2 b c) + + hence eq: "e = (AE_appP f b' v)" by simp + hence *:"\ ; \ ; {||} ;GNil ; \ \ (AE_appP f b' v) \ \ z : b | c \" using check_letI by auto + + then obtain x3 b3 c3 \3 s3 bv3 where + **:"\ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ Some (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3))) = lookup_fun \ f \ + \ ; {||} ; GNil \ v \ \ x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \ \ atom x3 \ GNil \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v = \ z : b | c \ + \ \ ; {||} \\<^sub>w\<^sub>f b'" + using infer_e_elims(21)[OF *] subst_defs by metis + + obtain z3 where z3:"\ x3 : b3 | c3 \ = \ z3 : b3 | c3[x3::=V_var z3]\<^sub>c\<^sub>v \ \ atom z3 \ (x3, v,c3,x1,c1)" using obtain_fresh_z3 by metis + + let ?ft = "(AF_fun_typ x3 (b3[bv3::=b']\<^sub>b\<^sub>b) (c3[bv3::=b']\<^sub>c\<^sub>b) (\3[bv3::=b']\<^sub>\\<^sub>b) (s3[bv3::=b']\<^sub>s\<^sub>b))" + + have *:"check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)))" using ** check_letI lookup_fun_member by metis + + hence ftq:"check_funtypq \ \ (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3))" using check_fundef_elims by auto + + let ?ft = "AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)" + + have sup: "supp \3 \ { atom x3, atom bv3} \ supp s3 \ { atom x3, atom bv3 }" + using wfPhi_f_poly_supp_t wfPhi_f_poly_supp_s ** by metis + + have "\; \; {||}; GNil; \ \ AS_let2 x2 \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 \ \" + proof + show \atom x2 \ (\, \, {||}::bv fset, GNil, \, \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v, s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v, \)\ + proof - + + have "atom x2 \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v" + using fresh_subst_v_subst_b subst_v_\_def subst_b_\_def \ atom x2 \ v\ sup by fastforce + moreover have "atom x2 \ s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v" + using fresh_subst_v_subst_b subst_v_s_def subst_b_s_def \ atom x2 \ v\ sup + proof - + have "\b. atom x2 = atom x3 \ atom x2 \ s3[bv3::=b]\<^sub>b" + by (metis (no_types) check_letI.hyps(1) fresh_subst_sv_if(1) fresh_subst_v_subst_b insert_commute subst_v_s_def sup) (* 140 ms *) + then show ?thesis + by (metis check_letI.hyps(1) fresh_subst_sb_if fresh_subst_sv_if(1) has_subst_b_class.subst_b_fresh_x x_fresh_b) (* 78 ms *) + qed + ultimately show ?thesis using fresh_prodN check_letI by metis + qed + + show \ \; \; {||}; GNil; \ \ s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v \ \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v\ proof( rule fundef_poly_check) + show \check_fundef \ \ (AF_fundef f (AF_fun_typ_some bv3 (AF_fun_typ x3 b3 c3 \3 s3)))\ + using ** lookup_fun_member check_letI by metis + show \\ ; {||} ; GNil \ v \ \ x3 : b3[bv3::=b']\<^sub>b\<^sub>b | c3[bv3::=b']\<^sub>c\<^sub>b \\ using ** by metis + show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using ** by metis + show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using ** by metis + qed + show \ \ ; \ ; {||} ; (x2, b_of \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v, c_of \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v x2) #\<^sub>\ GNil ; \ \ s2 \ \\ + using check_letI ** b_of.simps c_of.simps subst_defs by metis + qed + + moreover have "AS_let2 x2 \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v (s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v) s2 = AS_let2 x (\1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v) (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v) s" proof - + have *: "[[atom x2]]lst. s2 = [[atom x]]lst. s" using check_letI s_branch_s_branch_list.eq_iff by auto + moreover have " \3[bv3::=b']\<^sub>\\<^sub>b[x3::=v]\<^sub>\\<^sub>v = \1'[bv1::=b']\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v" using fun_poly_ret_unique ** check_letI by metis + moreover have "s3[bv3::=b']\<^sub>s\<^sub>b[x3::=v]\<^sub>s\<^sub>v = (s1'[bv1::=b']\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v)" using subst_v_flip_eq_two subst_v_s_def fun_poly_body_unique ** check_letI by metis + ultimately show ?thesis using s_branch_s_branch_list.eq_iff by metis + qed + + ultimately show ?case using check_letI by auto +qed(auto+) + +lemma check_s_plus: + assumes "\; \; {||}; GNil; \ \ LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \" + shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \ \" +proof - + obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)) \ t1" + using assms check_s_elims by metis + then obtain z1 where 2: "t1 = \ z1 : B_int | CE_val (V_var z1) == CE_op Plus ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \" + using infer_e_plus by metis + + obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit (L_num (n1+n2))) \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \\ + using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 + by (simp add: fresh_GNil) + + let ?e = " (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2)))" + + show ?thesis proof(rule subtype_let) + show "\ ; \ ; {||} ; GNil ; \ \ LET x = ?e IN s' \ \" using assms by auto + show "\ ; \ ; {||} ; GNil ; \ \ ?e \ t1" using 1 by auto + show "\ ; \ ; {||} ; GNil ; \ \ [ [ L_num (n1 + n2) ]\<^sup>v ]\<^sup>e \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \" using 3 by auto + show "\ ; {||} ; GNil \ \ z2 : B_int | CE_val (V_var z2) == CE_val (V_lit (L_num (n1+n2))) \ \ t1" using subtype_bop_arith + by (metis "1" \\thesis. (\z1. t1 = \ z1 : B_int | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ plus [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ \ thesis) \ thesis\ infer_e_wf(2) opp.distinct(1) type_for_lit.simps(3)) + qed + +qed + +lemma check_s_leq: + assumes "\ ; \ ; {||} ; GNil ; \ \ LET x = (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \" + shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (if (n1 \ n2) then L_true else L_false))) IN s' \ \" +proof - + obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2)) \ t1" + using assms check_s_elims by metis + then obtain z1 where 2: "t1 = \ z1 : B_bool | CE_val (V_var z1) == CE_op LEq ([V_lit (L_num n1)]\<^sup>c\<^sup>e) ([V_lit (L_num n2)]\<^sup>c\<^sup>e) \" + using infer_e_leq by auto + + obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \\ + using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 + fresh_GNil + by simp + + show ?thesis proof(rule subtype_let) + show \ \; \; {||}; GNil; \ \ AS_let x (AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v) s' \ \\ using assms by auto + show \\; \; {||}; GNil; \ \ AE_op LEq [ L_num n1 ]\<^sup>v [ L_num n2 ]\<^sup>v \ t1\ using 1 by auto + show \\; \; {||}; GNil; \ \ [ [ if n1 \ n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \\ using 3 by auto + show \\ ; {||} ; GNil \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 \ n2) then L_true else L_false))) \ \ t1\ + using subtype_bop_arith[where opp=LEq] check_s_wf assms 2 + by (metis opp.distinct(1) subtype_bop_arith type_l_eq) + qed +qed + +lemma check_s_eq: + assumes "\ ; \ ; {||} ; GNil ; \ \ LET x = (AE_op Eq (V_lit (n1)) (V_lit ( n2))) IN s' \ \" + shows "\; \; {||}; GNil; \ \ LET x = (AE_val (V_lit (if (n1 = n2) then L_true else L_false))) IN s' \ \" +proof - + obtain t1 where 1: "\; \; {||}; GNil; \ \ AE_op Eq (V_lit (n1)) (V_lit (n2)) \ t1" + using assms check_s_elims by metis + then obtain z1 where 2: "t1 = \ z1 : B_bool | CE_val (V_var z1) == CE_op Eq ([V_lit (n1)]\<^sup>c\<^sup>e) ([V_lit (n2)]\<^sup>c\<^sup>e) \" + using infer_e_leq by auto + + obtain z2 where 3: \\ ; \ ; {||} ; GNil ; \ \ AE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \\ + using infer_v_form infer_e_valI infer_v_litI infer_l.intros infer_e_wf 1 + fresh_GNil + by simp + + show ?thesis proof(rule subtype_let) + show \ \; \; {||}; GNil; \ \ AS_let x (AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v) s' \ \\ using assms by auto + show \\; \; {||}; GNil; \ \ AE_op Eq [ n1 ]\<^sup>v [ n2 ]\<^sup>v \ t1\ using 1 by auto + show \\; \; {||}; GNil; \ \ [ [ if n1 = n2 then L_true else L_false ]\<^sup>v ]\<^sup>e \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \\ using 3 by auto + show \\ ; {||} ; GNil \ \ z2 : B_bool | CE_val (V_var z2) == CE_val (V_lit ((if (n1 = n2) then L_true else L_false))) \ \ t1\ + proof - + have " \ z2 : B_bool | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ = t1" using 2 + by (metis \_fresh_c fresh_opp_all infer_l_form2 infer_l_fresh ms_fresh_all(31) ms_fresh_all(33) obtain_fresh_z type_e_eq type_l_eq) + moreover have "\ ; {||} \\<^sub>w\<^sub>f GNil" using assms wfX_wfY by fastforce + moreover have "base_for_lit n1 = base_for_lit n2" using 1 infer_e_wf wfE_elims(12) wfV_elims + by metis + ultimately show ?thesis using subtype_bop_eq[OF \\ ; {||} \\<^sub>w\<^sub>f GNil\, of n1 n2 z2] by auto + qed + qed +qed + +subsection \Operators\ + +lemma preservation_plus: + assumes "\; \; \ \ \ \ , LET x = (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) IN s' \ \ \" + shows "\; \; \ \ \ \ , LET x = (AE_val (V_lit (L_num (n1+n2)))) IN s' \ \ \" +proof - + + have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op Plus (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" + using assms config_type_elims by blast+ + + hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit (L_num (n1+n2)))) s' \ \" using check_s_plus assms by auto + + hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (L_num (n1+n2))))) s' \ \ \" using dsim config_typeI fd by presburger + then show ?thesis using dsim config_typeI + by (meson order_refl) +qed + +lemma preservation_leq: + assumes "\; \; \ \ \ \ , AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \ \" + shows "\; \; \ \ \ \ , AS_let x (AE_val (V_lit (((if (n1 \ n2) then L_true else L_false))))) s' \ \ \" +proof - + + have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op LEq (V_lit (L_num n1)) (V_lit (L_num n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" + using assms config_type_elims by blast+ + + hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit ( ((if (n1 \ n2) then L_true else L_false))))) s' \ \" using check_s_leq assms by auto + + hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (((if (n1 \ n2) then L_true else L_false)))))) s' \ \ \" using dsim config_typeI fd by presburger + then show ?thesis using dsim config_typeI + by (meson order_refl) +qed + +lemma preservation_eq: + assumes "\; \; \ \ \ \ , AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \ \ \" + shows "\; \; \ \ \ \ , AS_let x (AE_val (V_lit (((if (n1 = n2) then L_true else L_false))))) s' \ \ \" +proof - + + have tt: "\; \; {||}; GNil; \ \ AS_let x (AE_op Eq (V_lit (n1)) (V_lit (n2))) s' \ \" and dsim: "\ \ \ \ \" and fd:"(\fd\set \. check_fundef \ \ fd)" + using assms config_type_elims by blast+ + + hence "\; \; {||}; GNil; \ \AS_let x (AE_val (V_lit ( ((if (n1 = n2) then L_true else L_false))))) s' \ \" using check_s_eq assms by auto + + hence "\; \; \ \ \ \ , AS_let x (AE_val (V_lit ( (((if (n1 = n2) then L_true else L_false)))))) s' \ \ \" using dsim config_typeI fd by presburger + then show ?thesis using dsim config_typeI + by (meson order_refl) +qed + +subsection \Let Statements\ + +lemma subst_s_abs_lst: + fixes s::s and sa::s and v'::v + assumes "[[atom x]]lst. s = [[atom xa]]lst. sa" and "atom xa \ v \ atom x \ v" + shows "s[x::=v]\<^sub>s\<^sub>v = sa[xa::=v]\<^sub>s\<^sub>v" +proof - + obtain c'::x where cdash: "atom c' \ (v, x, xa, s, sa)" using obtain_fresh by blast + moreover have " (x \ c') \ s = (xa \ c') \ sa" proof - + have "atom c' \ (s, sa) \ atom c' \ (x, xa, s, sa)" using cdash by auto + thus ?thesis using assms by auto + qed + ultimately show ?thesis using assms + using subst_sv_flip by auto +qed + +lemma check_let_val: + fixes v::v and s::s + shows "\ ; \ ; B ; G ; \ \ ss \ \ \ B = {||} \ G = GNil \ + ss = AS_let x (AE_val v) s \ ss = AS_let2 x t (AS_val v) s \ \; \; {||}; GNil; \ \ (s[x::=v]\<^sub>s\<^sub>v) \ \" and + "check_branch_s \ \ \ GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ \ \ \ tid dclist v css \ \ True" +proof(nominal_induct \ and \ and \ avoiding: v rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_letI x1 \ \ \ \ \ e \ z s1 b c) + hence *:"e = AE_val v" by auto + let ?G = "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" + have "\ ; \ ; \ ; ?G[x1::=v]\<^sub>\\<^sub>v ; \[x1::=v]\<^sub>\\<^sub>v \ s1[x1::=v]\<^sub>s\<^sub>v \ \[x1::=v]\<^sub>\\<^sub>v" + proof(rule subst_infer_check_s(1)) + show **:\ \ ; \ ; \ \ v \ \ z : b | c \\ using infer_e_elims check_letI * by fast + thus \\ ; \ ; \ \ \ z : b | c \ \ \ z : b | c \\ using subtype_reflI infer_v_wf by metis + show \atom z \ (x1, v)\ using check_letI fresh_Pair by auto + show \\ ; \ ; \ ; (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \ ; \ \ s1 \ \\ using check_letI subst_defs by auto + show "(x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \ = GNil @ (x1, b, c[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" by auto + qed + + hence "\ ; \ ; \ ; \ ; \ \ s1[x1::=v]\<^sub>s\<^sub>v \ \" using check_letI by auto + moreover have " s1[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" + by (metis (full_types) check_letI fresh_GNil infer_e_elims(7) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(2) + subst_s_abs_lst wfG_x_fresh_in_v_simple) + + ultimately show ?case using check_letI by simp +next + case (check_let2I x1 \ \ \ \ \ t s1 \ s2 ) + + hence s1eq:"s1 = AS_val v" by auto + let ?G = "(x1, b_of t, c_of t x1) #\<^sub>\ \" + obtain z::x where *:"atom z \ (x1 , v,t)" using obtain_fresh_z by metis + hence teq:"t = \ z : b_of t | c_of t z \" using b_of_c_of_eq by auto + have "\ ; \ ; \ ; ?G[x1::=v]\<^sub>\\<^sub>v ; \[x1::=v]\<^sub>\\<^sub>v \ s2[x1::=v]\<^sub>s\<^sub>v \ \[x1::=v]\<^sub>\\<^sub>v" + proof(rule subst_check_check_s(1)) + obtain t' where "\ ; \ ; \ \ v \ t' \ \ ; \ ; \ \ t' \ t" using check_s_elims(1) check_let2I(10) s1eq by auto + thus **:\ \ ; \ ; \ \ v \ \ z : b_of t | c_of t z \\ using check_v.intros teq by auto + show "atom z \ (x1, v)" using * by auto + show \ \ ; \ ; \ ; (x1, b_of t, c_of t x1) #\<^sub>\ \ ; \ \ s2 \ \\ using check_let2I by auto + show "(x1, b_of t , c_of t x1) #\<^sub>\ \ = GNil @ (x1, b_of t, (c_of t z)[z::=V_var x1]\<^sub>c\<^sub>v) #\<^sub>\ \" using append_g.simps c_of_switch * fresh_prodN by metis + qed + + hence "\ ; \ ; \ ; \ ; \ \ s2[x1::=v]\<^sub>s\<^sub>v \ \" using check_let2I by auto + moreover have " s2[x1::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v" using + check_let2I fresh_GNil check_s_elims s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff + subst_s_abs_lst wfG_x_fresh_in_v_simple + proof - + have "AS_let2 x t (AS_val v) s = AS_let2 x1 t s1 s2" + by (metis check_let2I.prems(3) s_branch_s_branch_list.distinct s_branch_s_branch_list.eq_iff(3)) (* 62 ms *) + then show ?thesis + by (metis (no_types) check_let2I check_let2I.prems(2) check_s_elims(1) fresh_GNil s_branch_s_branch_list.eq_iff(3) subst_s_abs_lst wfG_x_fresh_in_v_simple) (* 375 ms *) + qed + + ultimately show ?case using check_let2I by simp +qed(auto+) + +lemma preservation_let_val: + assumes "\; \; \ \ \ \ , AS_let x (AE_val v) s \ \ \ \ \; \; \ \ \ \ , AS_let2 x t (AS_val v) s \ \ \" (is "?A \ ?B") + shows " \\'. \; \; \' \ \ \ , s[x::=v]\<^sub>s\<^sub>v \ \ \ \ \ \ \'" +proof - + have tt: "\ \ \ \ \" and fd: "(\fd\set \. check_fundef \ \ fd)" + using assms by blast+ + + have "?A \ ?B" using assms by auto + then have "\; \; {||}; GNil; \ \ s[x::=v]\<^sub>s\<^sub>v \ \" + proof + assume "\; \; \ \ \ \ , AS_let x (AE_val v) s \ \ \" + hence * : " \; \; {||}; GNil; \ \ AS_let x (AE_val v) s \ \" by blast + thus ?thesis using check_let_val by simp + next + assume "\; \; \ \ \ \ , AS_let2 x t (AS_val v) s \ \ \" + hence * : "\; \; {||}; GNil; \ \ AS_let2 x t (AS_val v) s \ \" by blast + thus ?thesis using check_let_val by simp + qed + + thus ?thesis using tt config_typeI fd + order_refl by metis +qed + +lemma check_s_fst_snd: + assumes "fst_snd = AE_fst \ v=v1 \ fst_snd = AE_snd \ v=v2" + and "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \" + shows "\; \; {||}; GNil; \ \ AS_let x ( AE_val v) s' \ \" +proof - + have 1: \ \; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \\ using assms by auto + + then obtain t1 where 2:"\; \; {||}; GNil; \ \ (fst_snd (V_pair v1 v2)) \ t1" using check_s_elims by auto + + show ?thesis using subtype_let 1 2 assms + by (meson infer_e_fst_pair infer_e_snd_pair) +qed + +lemma preservation_fst_snd: + assumes "\; \; \ \ \ \ , LET x = (fst_snd (V_pair v1 v2)) IN s' \ \ \" and + "fst_snd = AE_fst \ v=v1 \ fst_snd = AE_snd \ v=v2" + shows "\\'. \; \; \ \ \ \ , LET x = (AE_val v) IN s' \ \ \ \ \ \ \'" +proof - + have tt: "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \ \ \ \ \ \ \" using assms by blast + hence t2: "\; \; {||}; GNil; \ \ AS_let x (fst_snd (V_pair v1 v2)) s' \ \" by auto + + moreover have "\fd\set \. check_fundef \ \ fd" using assms config_type_elims by auto + ultimately show ?thesis using config_typeI order_refl tt assms check_s_fst_snd by metis +qed + +inductive_cases check_branch_s_elims2[elim!]: + "\ ; \ ; \ ; \ ; \; tid ; cons ; const ; v \ cs \ \" + +lemmas freshers = freshers atom_dom.simps toSet.simps fresh_def x_not_in_b_set +declare freshers [simp] + +lemma subtype_eq_if: + fixes t::\ and va::v + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c_of t z \" and "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c IMP c_of t z \ " + shows "\ ; \ ; \ \ \ z : b_of t | c_of t z \ \ \ z : b_of t | c IMP c_of t z \ " +proof - + obtain x::x where xf:"atom x \ ((\, \, \, z, c_of t z, z, c IMP c_of t z ),c)" using obtain_fresh by metis + + moreover have "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ (c IMP c_of t z )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v " + unfolding subst_cv.simps + proof(rule valid_eq_imp) + + have "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f (c IMP (c_of t z))[z::=[ x ]\<^sup>v]\<^sub>v " + apply(rule wfT_wfC_cons) + apply(simp add: assms, simp add: assms, unfold fresh_prodN ) + using xf fresh_prodN by metis+ + thus "\ ; \ ; (x, b_of t, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v IMP (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v " + using subst_cv.simps subst_defs by auto + qed + + ultimately show ?thesis using subtype_baseI assms fresh_Pair subst_defs by metis +qed + +lemma subtype_eq_if_\: + fixes t::\ and va::v + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f t" and "\ ; \ ; \ \\<^sub>w\<^sub>f \ z : b_of t | c IMP c_of t z \ " and "atom z \ t" + shows "\ ; \ ; \ \ t \ \ z : b_of t | c IMP c_of t z \ " +proof - + have "t = \ z : b_of t | c_of t z \" using b_of_c_of_eq assms by auto + thus ?thesis using subtype_eq_if assms c_of.simps b_of.simps by metis +qed + +lemma valid_conj: + assumes "\ ; \ ; \ \ c1" and "\ ; \ ; \ \ c2" + shows "\ ; \ ; \ \ c1 AND c2" +proof + show \ \ ; \ ; \ \\<^sub>w\<^sub>f c1 AND c2 \ using valid.simps wfC_conjI assms by auto + show \\i. \ ; \ \ i \ i \ \ \ i \ c1 AND c2 \ + proof(rule+) + fix i + assume *:"\ ; \ \ i \ i \ \ " + thus "i \ c1 \ ~ True" using assms valid.simps + using is_satis.cases by blast + show "i \ c2 \ ~ True" using assms valid.simps + using is_satis.cases * by blast + qed +qed + +subsection \Other Statements\ + +lemma check_if: + fixes s'::s and cs::branch_s and css::branch_list and v::v + shows "\; \; B; G; \ \ s' \ \ \ s' = IF (V_lit ll) THEN s1 ELSE s2 \ + \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ G = GNil \ B = {||} \ ll = L_true \ s = s1 \ ll = L_false \ s = s2 \ + \; \; {||}; GNil; \ \ s \ \" and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" +proof(nominal_induct \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_ifI z \ \ \ \ \ v s1 s2 \) + obtain z' where teq: "\ = \ z' : b_of \ | c_of \ z' \ \ atom z' \ (z,\)" using obtain_fresh_z_c_of by metis + hence ceq: "(c_of \ z')[z'::=[ z ]\<^sup>v]\<^sub>c\<^sub>v = (c_of \ z)" using c_of_switch fresh_Pair by metis + have zf: " atom z \ c_of \ z'" + using c_of_fresh check_ifI teq fresh_Pair fresh_at_base by metis + hence 1:"\; \; {||}; GNil; \ \ s \ \ z : b_of \ | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \ z \" using check_ifI by auto + moreover have 2:"\ ; {||} ; GNil \ (\ z : b_of \ | CE_val (V_lit ll) == CE_val (V_lit ll) IMP c_of \ z \) \ \" + proof - + have "\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ z : b_of \ | CE_val (V_lit ll ) == CE_val (V_lit ll) IMP c_of \ z \)" using check_ifI check_s_wf by auto + moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using check_s_wf check_ifI by auto + ultimately show ?thesis using subtype_if_simp[of \ " {||}" z "b_of \" ll "c_of \ z'" z'] using teq ceq zf subst_defs by metis + qed + ultimately show ?case using check_s_supertype(1) check_ifI by metis +qed(auto+) + +lemma preservation_if: + assumes "\; \; \ \ \ \ , IF (V_lit ll) THEN s1 ELSE s2 \ \ \" and + "ll = L_true \ s = s1 \ ll = L_false \ s = s2" + shows "\; \; \ \ \\, s\ \ \ \ setD \ \ setD \" +proof - + have *: "\; \; {||}; GNil; \ \ AS_if (V_lit ll) s1 s2 \ \ \ (\fd\set \. check_fundef \ \ fd)" + using assms config_type_elims by metis + hence "\; \; {||}; GNil; \ \ s \ \" using check_s_wf check_if assms by metis + hence "\; \; \ \ \\, s\ \ \ \ setD \ \ setD \" using config_typeI * + using assms(1) by blast + thus ?thesis by blast +qed + +lemma wfT_conj: + assumes "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 \" and "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c2 \" + shows "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 AND c2\" +proof + show \atom z \ (\, \, GNil)\ + apply(unfold fresh_prodN, intro conjI) + using wfTh_fresh wfT_wf assms apply metis + using fresh_GNil x_not_in_b_set fresh_def by metis+ + show \ \ ; \ \\<^sub>w\<^sub>f b \ using wfT_elims assms by metis + have "\ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c1" using wfT_wfC fresh_GNil assms by auto + moreover have "\ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c2" using wfT_wfC fresh_GNil assms by auto + ultimately show \ \ ; \ ; (z, b, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f c1 AND c2 \ using wfC_conjI by auto +qed + +lemma subtype_conj: + assumes "\ ; \ ; GNil \ t \ \ z : b | c1 \" and "\ ; \ ; GNil \ t \ \ z : b | c2 \" + shows "\ ; \ ; GNil \ \ z : b | c_of t z \ \ \ z : b | c1 AND c2 \" +proof - + have beq: "b_of t = b" using subtype_eq_base2 b_of.simps assms by metis + obtain x::x where x:\atom x \ (\, \, GNil, z, c_of t z, z, c1 AND c2 )\ using obtain_fresh by metis + thus ?thesis proof + have "atom z \ t" using subtype_wf wfT_supp fresh_def x_not_in_b_set atom_dom.simps toSet.simps assms dom.simps by fastforce + hence t:"t = \ z : b_of t | c_of t z \" using b_of_c_of_eq by auto + thus \ \ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c_of t z \ \ using subtype_wf beq assms by auto + + show \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ (c1 AND c2 )[z::=[ x ]\<^sup>v]\<^sub>v \ + proof - + have \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ c1[z::=[ x ]\<^sup>v]\<^sub>v \ + proof(rule subtype_valid) + show \\ ; \ ; GNil \ t \ \ z : b | c1 \\ using assms by auto + show \atom x \ GNil\ using fresh_GNil by auto + show \t = \ z : b | c_of t z \\ using t beq by auto + show \\ z : b | c1 \ = \ z : b | c1 \\ by auto + qed + moreover have \\ ; \ ; (x, b, (c_of t z)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ c2[z::=[ x ]\<^sup>v]\<^sub>v \ + proof(rule subtype_valid) + show \\ ; \ ; GNil \ t \ \ z : b | c2 \\ using assms by auto + show \atom x \ GNil\ using fresh_GNil by auto + show \t = \ z : b | c_of t z \\ using t beq by auto + show \\ z : b | c2 \ = \ z : b | c2 \\ by auto + qed + ultimately show ?thesis unfolding subst_cv.simps subst_v_c_def using valid_conj by metis + qed + thus \ \ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b | c1 AND c2 \ \ using subtype_wf wfT_conj assms by auto + qed +qed + +lemma infer_v_conj: + assumes "\ ; \ ; GNil \ v \ \ z : b | c1 \" and "\ ; \ ; GNil \ v \ \ z : b | c2 \" + shows "\ ; \ ; GNil \ v \ \ z : b | c1 AND c2 \" +proof - + obtain t1 where t1: "\ ; \ ; GNil \ v \ t1 \ \ ; \ ; GNil \ t1 \ \ z : b | c1 \" + using assms check_v_elims by metis + obtain t2 where t2: "\ ; \ ; GNil \ v \ t2 \ \ ; \ ; GNil \ t2 \ \ z : b | c2 \" + using assms check_v_elims by metis + have teq: "t1 = \ z : b | c_of t1 z \" using subtype_eq_base2 b_of.simps + by (metis (full_types) b_of_c_of_eq fresh_GNil infer_v_t_wf t1 wfT_x_fresh) + have "t1 = t2" using infer_v_uniqueness t1 t2 by auto + hence " \ ; \ ; GNil \ \ z : b | c_of t1 z \ \ \ z : b | c1 AND c2 \" using subtype_conj t1 t2 by simp + hence " \ ; \ ; GNil \ t1 \ \ z : b | c1 AND c2 \" using teq by auto + thus ?thesis using t1 using check_v.intros by auto +qed + +lemma wfG_conj: + fixes c1::c + assumes "\ ; \ \\<^sub>w\<^sub>f (x, b, c1 AND c2) #\<^sub>\ \" + shows "\ ; \ \\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\ \" +proof(cases "c1 \ {TRUE, FALSE}") + case True + then show ?thesis using assms wfG_cons2I wfG_elims wfX_wfY by metis +next + case False + then show ?thesis using assms wfG_cons1I wfG_elims wfX_wfY wfC_elims + by (metis wfG_elim2) +qed + +lemma check_match: + fixes s'::s and s::s and css::branch_list and cs::branch_s + shows "\ ; \ ; \ ; \ ; \ \ s \ \ \ True" and + "\ ; \ ; B ; G ; \ ; tid ; dc ; const ; vcons \ cs \ \ \ + vcons = V_cons tid dc v \ B = {||} \ G = GNil \ cs = (dc x' \ s') \ + \ ; {||} ; GNil \ v \ const \ + \; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" and + "\ ; \ ; B ; G ; \ ; tid ; dclist ; vcons \ css \ \ \ distinct (map fst dclist) \ + vcons = V_cons tid dc v \ B = {||} \ (dc, const) \ set dclist \ G = GNil \ + Some (AS_branch dc x' s') = lookup_branch dc css \ \ ; {||} ; GNil \ v \ const \ + \; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" +proof(nominal_induct \ and \ and \ avoiding: x' v rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_branch_list_consI \ \ \ \ \ tid consa consta va cs \ dclist cssa) + + then obtain xa and sa where cseq:"cs = AS_branch consa xa sa" using check_branch_s_elims2[OF check_branch_list_consI(1)] by metis + + show ?case proof(cases "dc = consa") + case True + hence "cs = AS_branch consa x' s'" using check_branch_list_consI cseq + by (metis lookup_branch.simps(2) option.inject) + moreover have "const = consta" using check_branch_list_consI distinct.simps + by (metis True dclist_distinct_unique list.set_intros(1)) + moreover have "va = V_cons tid consa v" using check_branch_list_consI True by auto + ultimately show ?thesis using check_branch_list_consI by auto + next + case False + hence "Some (AS_branch dc x' s') = lookup_branch dc cssa" using lookup_branch.simps(2) check_branch_list_consI(10) cseq by auto + moreover have "(dc, const) \ set dclist" using check_branch_list_consI False by simp + ultimately show ?thesis using check_branch_list_consI by auto + qed + +next + case (check_branch_list_finalI \ \ \ \ \ tid cons const va cs \) + hence " cs = AS_branch cons x' s'" using lookup.simps check_branch_list_finalI lookup_branch.simps option.inject + by (metis map_of.simps(1) map_of_Cons_code(2) option.distinct(1) s_branch_s_branch_list.exhaust(2) weak_map_of_SomeI) + then show ?case using check_branch_list_finalI by auto +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons va s) + + text \Supporting facts here to make the main body of the proof concise\ + have xf:"atom x \ \" proof - + have "supp \ \ supp \" using wf_supp(4) check_branch_s_branchI atom_dom.simps toSet.simps dom.simps by fastforce + thus ?thesis using fresh_def x_not_in_b_set by blast + qed + + hence "\[x::=v]\<^sub>\\<^sub>v = \" using forget_subst_v subst_v_\_def by auto + have "\[x::=v]\<^sub>\\<^sub>v = \" using forget_subst_dv wfD_x_fresh fresh_GNil check_branch_s_branchI by metis + + have "supp v = {}" using check_branch_s_branchI check_v_wf wfV_supp_nil by metis + hence "supp va = {}" using \ va = V_cons tid cons v\ v.supp pure_supp by auto + + let ?G = "(x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\ \" + obtain z::x where z: "const = \ z : b_of const | c_of const z \ \ atom z \ (x', v,x,const,va)" + using obtain_fresh_z_c_of by metis + + have vt: \\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \\ + proof(rule infer_v_conj) + obtain t' where t: "\ ; \ ; GNil \ v \ t' \ \ ; \ ; GNil \ t' \ const" + using check_v_elims check_branch_s_branchI by metis + show "\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \" + proof(rule check_v_top) + + show "\ ; \ ; GNil \\<^sub>w\<^sub>f \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e \ " + + proof(rule wfG_wfT) + show \ \ ; \ \\<^sub>w\<^sub>f (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil \ + proof - + have 1: "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis + moreover have 2: "\ ; \ \\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const x ) #\<^sub>\ GNil " + using check_branch_s_branchI(17)[THEN check_s_wf] \\ = GNil\ by auto + moreover hence "\ ; \ \\<^sub>w\<^sub>f (x, b_of const, [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ x ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" + using wfG_conj by metis + ultimately show ?thesis + unfolding subst_cv.simps subst_cev.simps subst_vv.simps by auto + qed + show \atom x \ ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e )\ unfolding c.fresh ce.fresh v.fresh + apply(intro conjI ) + using check_branch_s_branchI fresh_at_base z freshers apply simp + using check_branch_s_branchI fresh_at_base z freshers apply simp + using pure_supp apply force + using z fresh_x_neq fresh_prod5 by metis + qed + show \[ va ]\<^sup>c\<^sup>e = [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e[z::=v]\<^sub>c\<^sub>e\<^sub>v\ + using \ va = V_cons tid cons v\ subst_cev.simps subst_vv.simps by auto + show \ \ ; \ ; GNil \ v \ const \ using check_branch_s_branchI by auto + show "supp [ va ]\<^sup>c\<^sup>e \ supp \" using \supp va = {}\ ce.supp by simp + qed + show "\ ; \ ; GNil \ v \ \ z : b_of const | c_of const z \" + using check_branch_s_branchI z by auto + qed + + text \Main body of proof for this case\ + have "\ ; \ ; \ ; (?G)[x::=v]\<^sub>\\<^sub>v ; \[x::=v]\<^sub>\\<^sub>v \ s[x::=v]\<^sub>s\<^sub>v \ \[x::=v]\<^sub>\\<^sub>v" + proof(rule subst_check_check_s) + show \\ ; \ ; GNil \ v \ \ z : b_of const | [ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z \\ using vt by auto + show \atom z \ (x, v)\ using z fresh_prodN by auto + show \ \ ; \ ; \ ; ?G ; \ \ s \ \\ + using check_branch_s_branchI by auto + + show \ ?G = GNil @ (x, b_of const, ([ va ]\<^sup>c\<^sup>e == [ V_cons tid cons [ z ]\<^sup>v ]\<^sup>c\<^sup>e AND c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ GNil\ + proof - + have "va[z::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = va" using forget_subst_v subst_v_v_def z fresh_prodN by metis + moreover have "(c_of const z)[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c_of const x" + using c_of_switch[of z const x] z fresh_prodN by metis + ultimately show ?thesis + unfolding subst_cv.simps subst_cev.simps subst_vv.simps append_g.simps + using c_of_switch[of z const x] z fresh_prodN z fresh_prodN check_branch_s_branchI by argo + qed + qed + moreover have "s[x::=v]\<^sub>s\<^sub>v = s'[x'::=v]\<^sub>s\<^sub>v" + using check_branch_s_branchI subst_v_flip_eq_two subst_v_s_def s_branch_s_branch_list.eq_iff by metis + ultimately show ?case using check_branch_s_branchI \\[x::=v]\<^sub>\\<^sub>v = \\ \\[x::=v]\<^sub>\\<^sub>v = \\ by auto +qed(auto+) + +text \Lemmas for while reduction. Making these separate lemmas allows flexibility in wiring them into the main proof and robustness +if we change it\ + +lemma check_unit: + fixes \::\ and \::\ and \::\ and G::\ + assumes "\ ; {||} ; GNil \ \ z : B_unit | TRUE \ \ \'" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>f \" and "\ ; {||} \\<^sub>w\<^sub>f G" + shows \ \ ; \ ; {||} ; G ; \ \ [[ L_unit ]\<^sup>v]\<^sup>s \ \'\ +proof - + have *:"\ ; {||} ; GNil \ [L_unit]\<^sup>v \ \ z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \" + using infer_l.intros(4) infer_v_litI fresh_GNil assms wfX_wfY by (meson subtype_g_wf) + moreover have "\ ; {||} ; GNil \ \ z : B_unit | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_unit ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \'" + using subtype_top subtype_trans * infer_v_wf + by (meson assms(1)) + ultimately show ?thesis (*apply(rule check_valI, auto simp add: assms,rule * )*) + using subtype_top subtype_trans fresh_GNil assms check_valI assms check_s_g_weakening assms toSet.simps + by (metis bot.extremum infer_v_g_weakening subtype_weakening wfD_wf) +qed + +lemma preservation_var: + shows "\; \; {||}; GNil; \ \ VAR u : \' = v IN s \ \ \ \ \ \ \ \ \ atom u \ \ \ atom u \ \ \ + \; \; {||}; GNil; (u,\')#\<^sub>\\ \ s \ \ \ \ \ (u,v)#\ \ (u,\')#\<^sub>\\" + and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" +proof(nominal_induct "{||}::bv fset" GNil \ "VAR u : \' = v IN s" \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_varI u' \ \ \ \ s') + hence "\; \; {||}; GNil; (u, \') #\<^sub>\ \ \ s \ \" using check_s_abs_u check_v_wf by metis + + moreover have "\ \ ((u,v)#\) \ ((u,\')#\<^sub>\\)" proof + show \\ \ \ \ \ \ using check_varI by auto + show \\ ; {||} ; GNil \ v \ \'\ using check_varI by auto + show \u \ fst ` set \\ using check_varI fresh_d_fst_d by auto + qed + + ultimately show ?case by simp +qed(auto+) + +lemma check_while: + shows "\; \; {||}; GNil; \ \ WHILE s1 DO { s2 } \ \ \ atom x \ (s1, s2) \ atom z' \ x \ + \; \; {||}; GNil; \ \ LET x : (\ z' : B_bool | TRUE \) = s1 IN (IF (V_var x) THEN (s2 ;; (WHILE s1 DO {s2})) + ELSE ([ V_lit L_unit]\<^sup>s)) \ \" and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "check_branch_list \ \ {||} \ \ tid dclist v css \ \ True" +proof(nominal_induct "{||}::bv fset" GNil \ "AS_while s1 s2" \ and \ and \ avoiding: s1 s2 x z' rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_whileI \ \ \ s1 z s2 \') + have teq:"\ z' : B_bool | TRUE \ = \ z : B_bool | TRUE \" using \.eq_iff by auto + + show ?case proof + have " atom x \ \' " using wfT_nil_supp fresh_def subtype_wfT check_whileI(5) by fast + moreover have "atom x \ \ z' : B_bool | TRUE \ " using \.fresh c.fresh b.fresh by metis + ultimately show \atom x \ (\, \, {||}, GNil, \, \ z' : B_bool | TRUE \, s1, \')\ + apply(unfold fresh_prodN) + using check_whileI wb_x_fresh check_s_wf wfD_x_fresh fresh_empty_fset fresh_GNil fresh_Pair \atom x \ \'\ by metis + + show \ \ ; \ ; {||} ; GNil ; \ \ s1 \ \ z' : B_bool | TRUE \\ using check_whileI teq by metis + + let ?G = "(x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil" + + have cof:"(c_of \ z' : B_bool | TRUE \ x) = C_true" using c_of.simps check_whileI subst_cv.simps by metis + have wfg: "\ ; {||} \\<^sub>w\<^sub>f ?G" proof + show "c_of \ z' : B_bool | TRUE \ x \ {TRUE, FALSE}" using subst_cv.simps cof by auto + show "\ ; {||} \\<^sub>w\<^sub>f GNil " using wfG_nilI check_whileI wfX_wfY check_s_wf by metis + show "atom x \ GNil" using fresh_GNil by auto + show "\ ; {||} \\<^sub>w\<^sub>f b_of \ z' : B_bool | TRUE \ " using wfB_boolI wfX_wfY check_s_wf b_of.simps + by (metis \\ ; {||} \\<^sub>w\<^sub>f GNil\) + qed + + obtain zz::x where zf:\atom zz \ ((\, \, {||}::bv fset, ?G , \, [ x ]\<^sup>v, + AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \'),x,?G)\ + using obtain_fresh by blast + show \ \ ; \ ; {||} ; ?G ; \ \ + AS_if [ x ]\<^sup>v (AS_seq s2 (AS_while s1 s2)) (AS_val [ L_unit ]\<^sup>v) \ \'\ + proof + show "atom zz \ (\, \, {||}::bv fset, ?G , \, [ x ]\<^sup>v, AS_seq s2 (AS_while s1 s2), AS_val [ L_unit ]\<^sup>v, \')" using zf by auto + show \\ ; {||} ; ?G \ [ x ]\<^sup>v \ \ zz : B_bool | TRUE \\ proof + have "atom zz \ x \ atom zz \ (\, {||}::bv fset, ?G)" using zf fresh_prodN by metis + thus \ \ ; {||} ; ?G \ [ x ]\<^sup>v \\ zz : B_bool | [[zz]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \\ + using infer_v_varI lookup.simps wfg b_of.simps by metis + thus \\ ; {||} ; ?G \ \ zz : B_bool | [[ zz ]\<^sup>v]\<^sup>c\<^sup>e == [[ x ]\<^sup>v]\<^sup>c\<^sup>e \ \ \ zz : B_bool | TRUE \\ + using subtype_top infer_v_wf by metis + qed + show \ \ ; \ ; {||} ; ?G ; \ \ AS_seq s2 (AS_while s1 s2) \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ + proof + have "\ zz : B_unit | TRUE \ = \ z : B_unit | TRUE \" using \.eq_iff by auto + thus \ \ ; \ ; {||} ; ?G ; \ \ s2 \ \ zz : B_unit | TRUE \\ using check_s_g_weakening(1) [OF check_whileI(3) _ wfg] toSet.simps + by (simp add: \\ zz : B_unit | TRUE \ = \ z : B_unit | TRUE \\) + + show \ \ ; \ ; {||} ; ?G ; \ \ AS_while s1 s2 \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ + proof(rule check_s_supertype(1)) + have \ \; \; {||}; GNil; \ \ AS_while s1 s2 \ \'\ using check_whileI by auto + thus *:\ \ ; \ ; {||} ; ?G ; \ \ AS_while s1 s2 \ \' \ using check_s_g_weakening(1)[OF _ _ wfg] toSet.simps by auto + + show \\ ; {||} ; ?G \ \' \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ + proof(rule subtype_eq_if_\) + show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \' \ using * check_s_wf by auto + show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \ \ + apply(rule wfT_eq_imp, simp add: base_for_lit.simps) + using subtype_wf check_whileI wfg zf fresh_prodN by metis+ + show \atom zz \ \'\ using zf fresh_prodN by metis + qed + qed + + qed + show \ \ ; \ ; {||} ; ?G ; \ \ AS_val [ L_unit ]\<^sup>v \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ + proof(rule check_s_supertype(1)) + + show *:\ \ ; \ ; {||} ; ?G ; \ \ AS_val [ L_unit ]\<^sup>v \ \'\ + using check_unit[OF check_whileI(5) _ _ wfg] using check_whileI wfg wfX_wfY check_s_wf by metis + show \\ ; {||} ; ?G \ \' \ \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \\ + proof(rule subtype_eq_if_\) + show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \' \ using * check_s_wf by metis + show \ \ ; {||} ; ?G \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \ \ + apply(rule wfT_eq_imp, simp add: base_for_lit.simps) + using subtype_wf check_whileI wfg zf fresh_prodN by metis+ + show \atom zz \ \'\ using zf fresh_prodN by metis + qed + qed + qed + qed +qed(auto+) + +lemma check_s_narrow: + fixes s::s and x::x + assumes "atom x \ (\, \, \, \, \, c, \, s)" and "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ \" and + "\ ; \ ; \ \ c " + shows "\ ; \ ; \ ; \ ; \ \ s \ \" +proof - + let ?B = "({||}::bv fset)" + let ?v = "V_lit L_true" + + obtain z::x where z: "atom z \ (x, [ L_true ]\<^sup>v,c)" using obtain_fresh by metis + + have "atom z \ c" using z fresh_prodN by auto + hence c:" c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v = c" using subst_v_c_def by simp + + have "\ ; \ ; \ ; ((x,B_bool, c) #\<^sub>\ \)[x::=?v]\<^sub>\\<^sub>v ; \[x::=?v]\<^sub>\\<^sub>v \ s[x::=?v]\<^sub>s\<^sub>v \ \[x::=?v]\<^sub>\\<^sub>v" proof(rule subst_infer_check_s(1) ) + show vt: \ \ ; \ ; \ \ [ L_true ]\<^sup>v \ \ z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \ \ + using infer_v_litI check_s_wf wfG_elims(2) infer_trueI assms by metis + show \\ ; \ ; \ \ \ z : B_bool | (CE_val (V_var z)) == (CE_val (V_lit L_true )) \ \ \ z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \\ proof + show \atom x \ (\, \, \, z, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e , z, c[x::=[ z ]\<^sup>v]\<^sub>v)\ + apply(unfold fresh_prodN, intro conjI) + prefer 5 + using c.fresh ce.fresh v.fresh z fresh_prodN apply auto[1] + prefer 6 + using fresh_subst_v_if[of "atom x" c x] assms fresh_prodN apply simp + using z assms fresh_prodN apply metis + using z assms fresh_prodN apply metis + using z assms fresh_prodN apply metis + using z fresh_prodN assms fresh_at_base by metis+ + show \ \ ; \ ; \ \\<^sub>w\<^sub>f \ z : B_bool | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ \ using vt infer_v_wf by simp + show \ \ ; \ ; \ \\<^sub>w\<^sub>f \ z : B_bool | c[x::=[ z ]\<^sup>v]\<^sub>v \ \ proof(rule wfG_wfT) + show \ \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \ \ using c check_s_wf assms by metis + have " atom x \ [ z ]\<^sup>v" using v.fresh z fresh_at_base by auto + thus \atom x \ c[x::=[ z ]\<^sup>v]\<^sub>v\ using fresh_subst_v_if[of "atom x" c ] by auto + qed + have wfg: "\ ; \ \\<^sub>w\<^sub>f(x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \" + using wfT_wfG vt infer_v_wf fresh_prodN assms by simp + show \\ ; \ ; (x, B_bool, ([ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e )[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>v \ + using c valid_weakening[OF assms(3) _ wfg ] toSet.simps + using subst_v_c_def by auto + qed + show \atom z \ (x, [ L_true ]\<^sup>v)\ using z fresh_prodN by metis + show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \ s \ \\ using assms by auto + + thus \(x, B_bool, c) #\<^sub>\ \ = GNil @ (x, B_bool, c[x::=[ z ]\<^sup>v]\<^sub>v[z::=[ x ]\<^sup>v]\<^sub>c\<^sub>v) #\<^sub>\ \\ using append_g.simps c by auto + qed + + moreover have "((x,B_bool, c) #\<^sub>\ \)[x::=?v]\<^sub>\\<^sub>v = \ " using subst_gv.simps by auto + ultimately show ?thesis using assms forget_subst_dv forget_subst_sv forget_subst_tv fresh_prodN by metis +qed + +lemma check_assert_s: + fixes s::s and x::x + assumes "\; \; {||}; GNil; \ \ AS_assert c s \ \" + shows "\; \; {||}; GNil; \ \ s \ \ \ \ ; {||} ; GNil \ c " +proof - + let ?B = "({||}::bv fset)" + let ?v = "V_lit L_true" + + obtain x where x: "\ ; \ ; ?B ; (x,B_bool, c) #\<^sub>\ GNil ; \ \ s \ \ \ atom x \ (\, \, ?B, GNil, \, c, \, s ) \ \ ; ?B ; GNil \ c " + using check_s_elims(10)[OF \\ ; \ ; ?B ; GNil ; \ \ AS_assert c s \ \\] valid.simps by metis + + show ?thesis using assms check_s_narrow x by metis +qed + +lemma infer_v_pair2I: + "atom z \ (v1, v2) \ + atom z \ (\, \, \) \ + \ ; \ ; \ \ v1 \ t1 \ + \ ; \ ; \ \ v2 \ t2 \ + b1 = b_of t1 \ b2 = b_of t2 \ + \ ; \ ; \ \ [ v1 , v2 ]\<^sup>v \ \ z : [ b1 , b2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \" + using infer_v_pairI by simp + +subsection \Main Lemma\ + +lemma preservation: + assumes "\ \ \\, s\ \ \\', s'\" and "\; \; \ \ \\, s\ \ \" + shows "\\'. \; \; \' \ \\', s'\ \ \ \ \ \ \'" + using assms +proof(induct arbitrary: \ rule: reduce_stmt.induct) + case (reduce_let_plusI \ x n1 n2 s') + then show ?case using preservation_plus + by (metis order_refl) +next + case (reduce_let_leqI b n1 n2 \ x s) + then show ?case using preservation_leq by (metis order_refl) +next + case (reduce_let_eqI b n1 n2 \ \ x s) + then show ?case using preservation_eq[OF reduce_let_eqI(2)] order_refl by metis +next + case (reduce_let_appI f z b c \' s' \ \ x v s) + hence tt: "\; \; {||}; GNil; \ \ AS_let x (AE_app f v) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let_appI(2)] by metis + hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_app f v) s \ \" by auto + + hence "\; \; {||}; GNil; \ \ AS_let2 x (\'[z::=v]\<^sub>\\<^sub>v) (s'[z::=v]\<^sub>s\<^sub>v) s \ \" + using preservation_app reduce_let_appI tt by auto + + hence "\; \; \ \ \ \ , AS_let2 x (\'[z::=v]\<^sub>\\<^sub>v) s'[z::=v]\<^sub>s\<^sub>v s \ \ \" using config_typeI tt by auto + then show ?case using tt order_refl reduce_let_appI by metis + +next + case (reduce_let_appPI f bv z b c \' s' \ \ x b' v s) + + hence tt: "\; \; {||}; GNil; \ \ AS_let x (AE_appP f b' v) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let_appPI(2)] by metis + hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_appP f b' v) s \ \" by auto + + have "\; \; {||}; GNil; \ \ AS_let2 x (\'[bv::=b']\<^sub>\\<^sub>b[z::=v]\<^sub>\\<^sub>v) (s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v) s \ \" + proof(rule preservation_poly_app) + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ z b c \' s'))) = lookup_fun \ f\ using reduce_let_appPI by metis + show \\fd\set \. check_fundef \ \ fd\ using tt lookup_fun_member by metis + show \ \ ; \ ;{||} ; GNil ; \ \ AS_let x (AE_appP f b' v) s \ \\ using * by auto + show \ \ ; {||} \\<^sub>w\<^sub>f b' \ using check_s_elims infer_e_wf wfE_elims * by metis + qed(auto+) + + hence "\; \; \ \ \ \ , AS_let2 x (\'[bv::=b']\<^sub>\\<^sub>b[z::=v]\<^sub>\\<^sub>v) s'[bv::=b']\<^sub>s\<^sub>b[z::=v]\<^sub>s\<^sub>v s \ \ \" using config_typeI tt by auto + then show ?case using tt order_refl reduce_let_appI by metis + +next + case (reduce_if_trueI \ s1 s2) + then show ?case using preservation_if by metis +next + case (reduce_if_falseI uw \ s1 s2) + then show ?case using preservation_if by metis +next + case (reduce_let_valI \ x v s) + then show ?case using preservation_let_val by presburger +next + case (reduce_let_mvar u v \ \ x s) + hence *:"\; \; {||}; GNil; \ \ AS_let x (AE_mvar u) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by blast + + hence **: "\; \; {||}; GNil; \ \ AS_let x (AE_mvar u) s \ \" by auto + obtain xa::x and za::x and ca::c and ba::b and sa::s where + sa1: "atom xa \ (\, \, {||}::bv fset, GNil, \, AE_mvar u, \) \ atom za \ (xa, \, \, {||}::bv fset, GNil, \, AE_mvar u, \, sa) \ + \; \; {||}; GNil; \ \ AE_mvar u \ \ za : ba | ca \ \ + \ ; \ ; {||} ; (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \ \ + (\c. atom c \ (s, sa) \ atom c \ (x, xa, s, sa) \ (x \ c) \ s = (xa \ c) \ sa)" + using check_s_elims(2)[OF **] subst_defs by metis + + have "\ ; {||} ; GNil \ v \ \ za : ba | ca \" proof - + have " (u , \ za : ba | ca \) \ setD \" using infer_e_elims(11) sa1 by fast + thus ?thesis using delta_sim_v reduce_let_mvar config_type_elims check_s_wf by metis + qed + + then obtain \' where vst: "\ ; {||} ; GNil \ v \ \' \ + \ ; {||} ; GNil \ \' \ \ za : ba | ca \" using check_v_elims by blast + + obtain za2 and ba2 and ca2 where zbc: "\' = (\ za2 : ba2 | ca2 \) \ atom za2 \ (xa, (xa, \, \, {||}::bv fset, GNil, \, AE_val v, \, sa))" + using obtain_fresh_z by blast + have beq: "ba=ba2" using subtype_eq_base vst zbc by blast + + moreover have xaf: "atom xa \ (za, za2)" + apply(unfold fresh_prodN, intro conjI) + using sa1 zbc fresh_prodN fresh_x_neq by metis+ + + have sat2: " \ ; \ ; {||} ; GNil@(xa, ba, ca2[za2::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" proof(rule ctx_subtype_s) + show "\ ; \ ; {||} ; GNil @ (xa, ba, ca[za::=V_var xa]\<^sub>c\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" using sa1 by auto + show "\ ; {||} ; GNil \ \ za2 : ba | ca2 \ \ \ za : ba | ca \" using beq zbc vst by fast + show "atom xa \ (za, za2, ca, ca2)" proof - + have *:"\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ za2 : ba2 | ca2 \) " using zbc vst subtype_wf by auto + hence "supp ca2 \ { atom za2 }" using wfT_supp_c[OF *] supp_GNil by simp + moreover have "atom za2 \ xa" using zbc fresh_Pair fresh_x_neq by metis + ultimately have "atom xa \ ca2" using zbc supp_at_base fresh_def + by (metis empty_iff singleton_iff subset_singletonD) + moreover have "atom xa \ ca" proof - + have *:"\ ; {||} ; GNil \\<^sub>w\<^sub>f (\ za : ba | ca \) " using zbc vst subtype_wf by auto + hence "supp ca \ { atom za }" using wfT_supp \.supp by force + moreover have "xa \ za" using fresh_def fresh_x_neq xaf fresh_Pair by metis + ultimately show ?thesis using fresh_def by auto + qed + ultimately show ?thesis using xaf sa1 fresh_prod4 fresh_Pair by metis + qed + qed + hence dwf: "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" using sa1 infer_e_wf by meson + + have "\; \; {||}; GNil; \ \ AS_let xa (AE_val v) sa \ \" proof + have "atom xa \ (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce + thus "atom xa \ (\, \, {||}::bv fset, GNil, \, AE_val v, \)" using sa1 freshers by simp + have "atom za2 \ (AE_val v)" using infer_v_wf(1) wfV_supp fresh_def e.fresh x_not_in_b_set vst by fastforce + thus "atom za2 \ (xa, \, \, {||}::bv fset, GNil, \, AE_val v, \, sa)" using zbc freshers fresh_prodN by auto + have " \ \\<^sub>w\<^sub>f \" using sa1 infer_e_wf by auto + thus "\; \; {||}; GNil; \ \ AE_val v \ \ za2 : ba | ca2 \" + using zbc vst beq dwf infer_e_valI by blast + show "\ ; \ ; {||} ; (xa, ba, ca2[za2::=V_var xa]\<^sub>v) #\<^sub>\ GNil ; \ \ sa \ \" using sat2 append_g.simps subst_defs by metis + qed + moreover have "AS_let xa (AE_val v) sa = AS_let x (AE_val v) s" proof - + have "[[atom x]]lst. s = [[atom xa]]lst. sa" + using sa1 Abs1_eq_iff_all(3)[where z=" (s, sa)"] by metis + thus ?thesis using s_branch_s_branch_list.eq_iff(2) by metis + qed + ultimately have "\; \; {||}; GNil; \ \ AS_let x (AE_val v) s \ \" by auto + + then show ?case using reduce_let_mvar * config_typeI + by (meson order_refl) +next + case (reduce_let2I \ \ s1 \' s1' x t s2) + hence **: "\; \; {||}; GNil; \ \ AS_let2 x t s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" using config_type_elims[OF reduce_let2I(3)] by blast + hence *:"\; \; {||}; GNil; \ \ AS_let2 x t s1 s2 \ \" by auto + + obtain xa::x and z::x and c and b and s2a::s where st: "atom xa \ (\, \, {||}::bv fset, GNil, \, t, s1, \) \ + \; \; {||}; GNil; \ \ s1 \ t \ + \ ; \ ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\ GNil ; \ \ s2a \ \ \ ([[atom x]]lst. s2 = [[atom xa]]lst. s2a) " + using check_s_elims(4)[OF *] Abs1_eq_iff_all(3) by metis + + hence "\; \; \ \ \ \ , s1 \ \ t" using config_typeI ** by auto + then obtain \' where s1r: "\; \; \' \ \ \' , s1' \ \ t \ \ \ \'" using reduce_let2I by presburger + + have "\; \; {||}; GNil; \' \ AS_let2 xa t s1' s2a \ \" + proof(rule check_let2I) + show *:"\; \; {||}; GNil; \' \ s1' \ t" using config_type_elims st s1r by metis + show "atom xa \ (\, \, {||}::bv fset, GNil, \',t, s1', \)" proof - + have "atom xa \ s1'" using check_s_x_fresh * by auto + moreover have "atom xa \ \'" using check_s_x_fresh * by auto + ultimately show ?thesis using st fresh_prodN by metis + qed + + show "\ ; \ ; {||} ; (xa, b_of t, c_of t xa) #\<^sub>\ GNil ; \' \ s2a \ \" proof - + have "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using * check_s_wf by auto + moreover have "\ ; {||} \\<^sub>w\<^sub>f ((xa, b_of t, c_of t xa) #\<^sub>\ GNil)" using st check_s_wf by auto + ultimately have "\ ; {||} ; ((xa, b_of t , c_of t xa) #\<^sub>\ GNil) \\<^sub>w\<^sub>f \'" using wf_weakening by auto + thus ?thesis using check_s_d_weakening check_s_wf st s1r by metis + qed + qed + moreover have "AS_let2 xa t s1' s2a = AS_let2 x t s1' s2" using st s_branch_s_branch_list.eq_iff by metis + ultimately have "\; \; {||}; GNil; \' \ AS_let2 x t s1' s2 \ \" using st by argo + moreover have "\ \ \' \ \'" using config_type_elims s1r by fast + ultimately show ?case using config_typeI ** + by (meson s1r) +next + case (reduce_let2_valI vb \ x t v s) + then show ?case using preservation_let_val by meson +next + case (reduce_varI u \ \ \' v s) + + hence ** : "\; \; {||}; GNil; \ \ AS_var u \' v s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by meson + have uf: "atom u \ \" using reduce_varI delta_sim_fresh by force + hence *: "\; \; {||}; GNil; \ \ AS_var u \' v s \ \" and " \ \ \ \ \" using ** by auto + + thus ?case using preservation_var reduce_varI config_typeI ** set_subset_Cons + setD_ConsD subsetI by (metis delta_sim_fresh) + +next + case (reduce_assignI \ \ u v ) + hence *: "\; \; {||}; GNil; \ \ AS_assign u v \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by meson + then obtain z and \' where zt: "\ ; {||} ; GNil \ (\ z : B_unit | TRUE \) \ \ \ (u,\') \ setD \ \ \ ; {||} ; GNil \ v \ \' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \" + using check_s_elims(8) by metis + hence "\ \ update_d \ u v \ \" using update_d_sim * by metis + moreover have "\; \; {||}; GNil; \ \ AS_val (V_lit L_unit ) \ \" using zt * check_s_v_unit check_s_wf + by auto + ultimately show ?case using config_typeI * by (meson order_refl) +next + case (reduce_seq1I \ \ s) + hence "\ ; \ ; {||} ; GNil ; \ \ s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using check_s_elims config_type_elims by force + then show ?case using config_typeI by blast +next + case (reduce_seq2I s1 v \ \ \' s1' s2) + hence tt: "\; \; {||}; GNil; \ \ AS_seq s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by blast + then obtain z where zz: "\ ; \ ; {||} ; GNil; \ \ s1 \ (\ z : B_unit | TRUE \) \ \; \; {||}; GNil; \ \ s2 \ \" + using check_s_elims by blast + hence "\; \; \ \ \ \ , s1 \ \ (\ z : B_unit | TRUE \)" + using tt config_typeI tt by simp + then obtain \' where *: "\; \; \' \ \ \' , s1' \ \ (\ z : B_unit | TRUE \) \ \ \ \'" + using reduce_seq2I by meson + moreover hence s't: "\; \; {||}; GNil; \' \ s1' \ (\ z : B_unit | TRUE \) \ \ \ \' \ \'" + using config_type_elims by force + moreover hence "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" using check_s_wf by meson + moreover hence "\; \; {||}; GNil; \' \ s2 \ \" + using calculation(1) zz check_s_d_weakening * by metis + moreover hence "\; \; {||}; GNil; \' \ (AS_seq s1' s2) \ \" + using check_seqI zz s't by meson + ultimately have "\; \; \' \ \ \' , AS_seq s1' s2 \ \ \ \ \ \ \'" + using zz config_typeI tt by meson + then show ?case by meson +next + case (reduce_whileI x s1 s2 z' \ \ ) + + hence *: "\; \; {||}; GNil; \ \ AS_while s1 s2 \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by meson + + hence **:"\; \; {||}; GNil; \ \ AS_while s1 s2 \ \" by auto + hence "\; \; {||}; GNil; \ \ AS_let2 x (\ z' : B_bool | TRUE \) s1 (AS_if (V_var x) (AS_seq s2 (AS_while s1 s2)) (AS_val (V_lit L_unit)) ) \ \" + using check_while reduce_whileI by auto + thus ?case using config_typeI * by (meson subset_refl) + +next + case (reduce_caseI dc x' s' css \ \ tyid v) + + hence **: "\; \; {||}; GNil; \ \ AS_match (V_cons tyid dc v) css \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims[OF reduce_caseI(2)] by metis + hence ***: "\; \; {||}; GNil; \ \ AS_match (V_cons tyid dc v) css \ \" by auto + + let ?vcons = "V_cons tyid dc v" + + obtain dclist tid and z::x where cv: "\; {||} ; GNil \ (V_cons tyid dc v) \ (\ z : B_id tid | TRUE \) \ + \; \; {||}; GNil; \ ; tid ; dclist ; (V_cons tyid dc v) \ css \ \ \ AF_typedef tid dclist \ set \ \ + \ ; {||} ; GNil \ V_cons tyid dc v \ \ z : B_id tid | TRUE \" + using check_s_elims(9)[OF ***] by metis + + hence vi:" \ ; {||} ; GNil \ V_cons tyid dc v \ \ z : B_id tid | TRUE \" by auto + obtain tcons where vi2:" \ ; {||} ; GNil \ V_cons tyid dc v \ tcons \ \ ; {||} ; GNil \ tcons \ \ z : B_id tid | TRUE \" + using check_v_elims(1)[OF vi] by metis + hence vi1: "\ ; {||} ; GNil \ V_cons tyid dc v \ tcons" by auto + + show ?case proof(rule infer_v_elims(4)[OF vi1],goal_cases) + case (1 dclist2 tc tv z2) + have "tyid = tid" using \.eq_iff using subtype_eq_base vi2 1 by fastforce + hence deq:"dclist = dclist2" using check_v_wf wfX_wfY cv 1 wfTh_dclist_unique by metis + have "\; \; {||}; GNil; \ \ s'[x'::=v]\<^sub>s\<^sub>v \ \" proof(rule check_match(3)) + show \ \ ; \ ; {||} ; GNil ; \ ; tyid ; dclist ; ?vcons \ css \ \\ using \tyid = tid\ cv by auto + show "distinct (map fst dclist)" using wfTh_dclist_distinct check_v_wf wfX_wfY cv by metis + show \?vcons = V_cons tyid dc v\ by auto + show \{||} = {||}\ by auto + show \(dc, tc) \ set dclist\ using 1 deq by auto + show \GNil = GNil\ by auto + show \Some (AS_branch dc x' s') = lookup_branch dc css\ using reduce_caseI by auto + show \\ ; {||} ; GNil \ v \ tc\ using 1 check_v.intros by auto + qed + thus ?case using config_typeI ** by blast + qed + +next + case (reduce_let_fstI \ \ x v1 v2 s) + thus ?case using preservation_fst_snd order_refl by metis +next + case (reduce_let_sndI \ \ x v1 v2 s) + thus ?case using preservation_fst_snd order_refl by metis +next + case (reduce_let_concatI \ \ x v1 v2 s) + hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \ \ \ + \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by metis + + obtain z::x where z: "atom z \ (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))))" + using obtain_fresh by metis + + have "\ ; {||} \\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto + + have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_lit (L_bitvec (v1 @ v2)))) s \ \" proof(rule subtype_let) + show \ \; \; {||}; GNil; \ \ AS_let x (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) s \ \\ using elim by auto + show \\; \; {||}; GNil; \ \ (AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \ \ z : B_bitvec | CE_val (V_var z) == (CE_concat ([V_lit (L_bitvec v1)]\<^sup>c\<^sup>e) ([V_lit (L_bitvec v2)]\<^sup>c\<^sup>e))\ \ + (is "\; \; {||}; GNil; \ \ ?e1 \ ?t1") + proof + show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto + show \ \ \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto + show \ \ ; {||} ; GNil \ V_lit (L_bitvec v1) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v1)) \\ + using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto + show \ \ ; {||} ; GNil \ V_lit (L_bitvec v2) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v2)) \\ + using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto + show \atom z \ AE_concat (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))\ using z fresh_Pair by metis + show \atom z \ GNil\ using z fresh_Pair by auto + qed + show \\; \; {||}; GNil; \ \ AE_val (V_lit (L_bitvec (v1 @ v2))) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \ \ + (is "\; \; {||}; GNil; \ \ ?e2 \ ?t2") + using infer_e_valI infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil check_s_wf elim by metis + show \\ ; {||} ; GNil \ ?t2 \ ?t1\ using subtype_concat check_s_wf elim by auto + qed + + thus ?case using config_typeI elim by (meson order_refl) +next + case (reduce_let_lenI \ \ x v s) + hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_len (V_lit (L_bitvec v))) s \ \ \ \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using check_s_elims config_type_elims by metis + + then obtain t where t: "\; \; {||}; GNil; \ \ AE_len (V_lit (L_bitvec v)) \ t" using check_s_elims by meson + + moreover then obtain z::x where "t = \ z : B_int | CE_val (V_var z) == CE_len ([V_lit (L_bitvec v)]\<^sup>c\<^sup>e) \" using infer_e_elims by meson + + moreover obtain z'::x where "atom z' \ v" using obtain_fresh by metis + moreover have "\; \; {||}; GNil; \ \ AE_val (V_lit (L_num (int (length v)))) \ \ z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \" + using infer_e_valI infer_v_litI infer_l.intros(3) t check_s_wf elim + by (metis infer_l_form2 type_for_lit.simps(3)) + + moreover have "\ ; {||} ; GNil \ \ z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \ \ + \ z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \" using subtype_len check_s_wf elim by auto + + ultimately have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_lit (L_num (int (length v))))) s \ \" using subtype_let by (meson elim) + thus ?case using config_typeI elim by (meson order_refl) +next + case (reduce_let_splitI n v v1 v2 \ \ x s) + hence elim: "\; \; {||}; GNil; \ \ AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \ \ \ + \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by metis + + obtain z::x where z: "atom z \ (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n)), GNil, CE_val (V_lit (L_bitvec (v1 @ v2))), +([ L_bitvec v1 ]\<^sup>v, [ L_bitvec v2 ]\<^sup>v), \, {||}::bv fset)" + using obtain_fresh by metis + + have *:"\ ; {||} \\<^sub>w\<^sub>f GNil" using check_s_wf elim by auto + + have "\; \; {||}; GNil; \ \ AS_let x (AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) s \ \" proof(rule subtype_let) + + show \ \; \; {||}; GNil; \ \ AS_let x (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) s \ \\ using elim by auto + show \\; \; {||}; GNil; \ \ (AE_split (V_lit (L_bitvec v)) (V_lit (L_num n))) \ \ z : B_pair B_bitvec B_bitvec + | ((CE_val (V_lit (L_bitvec v))) == (CE_concat (CE_fst (CE_val (V_var z))) (CE_snd (CE_val (V_var z))))) + AND (((CE_len (CE_fst (CE_val (V_var z))))) == (CE_val (V_lit (L_num n)))) \ \ + (is "\; \; {||}; GNil; \ \ ?e1 \ ?t1") + proof + show \ \ ; {||} ; GNil \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto + show \ \ \\<^sub>w\<^sub>f \ \ using check_s_wf elim by auto + show \ \ ; {||} ; GNil \ V_lit (L_bitvec v) \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec v)) \\ + using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ fresh_GNil by auto + show "\ ; {||} ; GNil \ ([ L_num + n ]\<^sup>v) \ \ z : B_int | (([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e) == ([ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e)) AND [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec + v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \" using split_n reduce_let_splitI check_v_num_leq * wfX_wfY by metis + show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto + show \atom z \ GNil\ using z fresh_Pair by auto + show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto + show \atom z \ GNil\ using z fresh_Pair by auto + show \atom z \ AE_split [ L_bitvec v ]\<^sup>v [ L_num n ]\<^sup>v\ using z fresh_Pair by auto + show \atom z \ GNil\ using z fresh_Pair by auto + qed + + show \\; \; {||}; GNil; \ \ AE_val (V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2))) \ \ z : B_pair B_bitvec B_bitvec | CE_val (V_var z) == CE_val ((V_pair (V_lit (L_bitvec v1)) (V_lit (L_bitvec v2)))) \ \ + (is "\; \; {||}; GNil; \ \ ?e2 \ ?t2") + apply(rule infer_e_valI) + using check_s_wf elim apply metis + using check_s_wf elim apply metis + apply(rule infer_v_pair2I) + using z fresh_prodN apply metis + using z fresh_GNil fresh_prodN apply metis + using infer_v_litI infer_l.intros \\ ; {||} \\<^sub>w\<^sub>f GNil\ b_of.simps apply blast+ + using b_of.simps apply simp+ + done + show \\ ; {||} ; GNil \ ?t2 \ ?t1\ using subtype_split check_s_wf elim reduce_let_splitI by auto + qed + + thus ?case using config_typeI elim by (meson order_refl) +next + case (reduce_assert1I \ \ c v) + + hence elim: "\; \; {||}; GNil; \ \ AS_assert c [v]\<^sup>s \ \ \ + \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims reduce_assert1I by metis + hence *:"\; \; {||}; GNil; \ \ AS_assert c [v]\<^sup>s \ \" by auto + + have "\; \; {||}; GNil; \ \ [v]\<^sup>s \ \" using check_assert_s * by metis + thus ?case using elim config_typeI by blast +next + case (reduce_assert2I \ \ s \' s' c) + + hence elim: "\; \; {||}; GNil; \ \ AS_assert c s \ \ \ + \ \ \ \ \ \ (\fd\set \. check_fundef \ \ fd)" + using config_type_elims by metis + hence *:"\; \; {||}; GNil; \ \ AS_assert c s \ \" by auto + + have cv: "\; \; {||}; GNil; \ \ s \ \ \ \ ; {||} ; GNil \ c " using check_assert_s * by metis + + hence "\; \; \ \ \\, s\ \ \" using elim config_typeI by simp + then obtain \' where D: "\; \; \' \ \ \' , s' \ \ \ \ \ \ \'" using reduce_assert2I by metis + hence **:"\; \; {||}; GNil; \' \ s' \ \ \ \ \ \' \ \'" using config_type_elims by metis + + obtain x::x where x:"atom x \ (\, \, ({||}::bv fset), GNil, \', c, \, s')" using obtain_fresh by metis + + have *:"\; \; {||}; GNil; \' \ AS_assert c s' \ \" proof + show "atom x \ (\, \, {||}, GNil, \', c, \, s')" using x by auto + have "\ ; {||} ; GNil \\<^sub>w\<^sub>f c" using * check_s_wf by auto + hence wfg:"\ ; {||} \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ GNil" using wfC_wfG wfB_boolI check_s_wf * fresh_GNil by auto + moreover have cs: "\; \; {||}; GNil; \' \ s' \ \" using ** by auto + ultimately show "\ ; \ ; {||} ; (x, B_bool, c) #\<^sub>\ GNil ; \' \ s' \ \" using check_s_g_weakening(1)[OF cs _ wfg] toSet.simps by simp + show "\ ; {||} ; GNil \ c " using cv by auto + show "\ ; {||} ; GNil \\<^sub>w\<^sub>f \' " using check_s_wf ** by auto + qed + + thus ?case using elim config_typeI D ** by metis +qed + +lemma preservation_many: + assumes " \ \ \\, s\ \\<^sup>* \ \' , s' \" + shows "\; \; \ \ \\, s\ \ \ \ \\'. \; \; \' \ \ \' , s' \ \ \ \ \ \ \'" + using assms proof(induct arbitrary: \ rule: reduce_stmt_many.induct) + case (reduce_stmt_many_oneI \ \ s \' s') + then show ?case using preservation by simp +next + case (reduce_stmt_many_manyI \ \ s \' s' \'' s'') + then show ?case using preservation subset_trans by metis +qed + +section \Progress\ +text \We prove that a well typed program is either a value or we can make a step\ + +lemma check_let_op_infer: + assumes "\; \; {||}; \; \ \ LET x = (AE_op opp v1 v2) IN s \ \" and "supp ( LET x = (AE_op opp v1 v2) IN s) \ atom`fst`setD \" + shows "\ z b c. \; \; {||}; \; \ \ (AE_op opp v1 v2) \ \z:b|c\" +proof - + have xx: "\; \; {||}; \; \ \ LET x = (AE_op opp v1 v2) IN s \ \" using assms by simp + then show ?thesis using check_s_elims(2)[OF xx] by meson +qed + +lemma infer_pair: + assumes "\ ; B; \ \ v \ \ z : B_pair b1 b2 | c \" and "supp v = {}" + obtains v1 and v2 where "v = V_pair v1 v2" + using assms proof(nominal_induct v rule: v.strong_induct) + case (V_lit x) + then show ?case by auto +next + case (V_var x) + then show ?case using v.supp supp_at_base by auto +next + case (V_pair x1a x2a) + then show ?case by auto +next + case (V_cons x1a x2a x3) + then show ?case by auto +next + case (V_consp x1a x2a x3 x4) + then show ?case by auto +qed + +lemma progress_fst: + assumes "\; \; {||}; \; \ \ LET x = (AE_fst v) IN s \ \" and "\ \ \ \ \" and + "supp (LET x = (AE_fst v) IN s) \ atom`fst`setD \" + shows "\\' s'. \ \ \ \ , LET x = (AE_fst v) IN s \ \ \\', s'\" +proof - + have *:"supp v = {}" using assms s_branch_s_branch_list.supp by auto + obtain z and b and c where "\; \; {||}; \; \ \ (AE_fst v ) \ \ z : b | c \" + using check_s_elims(2) using assms by meson + moreover obtain z' and b' and c' where "\ ; {||} ; \ \ v \ \ z' : B_pair b b' | c' \" + using infer_e_elims(8) using calculation by auto + moreover then obtain v1 and v2 where "V_pair v1 v2 = v" + using * infer_pair by metis + ultimately show ?thesis using reduce_let_fstI assms by metis +qed + +lemma progress_let: + assumes "\; \; {||}; \; \ \ LET x = e IN s \ \" and "\ \ \ \ \" and + "supp (LET x = e IN s) \ atom ` fst ` setD \" and "sble \ \" + shows "\\' s'. \ \ \ \ , LET x = e IN s\ \ \ \' , s'\" +proof - + obtain z b c where *: "\; \; {||}; \; \ \ e \ \ z : b | c \ " using check_s_elims(2)[OF assms(1)] by metis + have **: "supp e \ atom ` fst ` setD \" using assms s_branch_s_branch_list.supp by auto + from * ** assms show ?thesis proof(nominal_induct "\ z : b | c \" rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \ \ v) + then show ?case using reduce_stmt_elims reduce_let_valI by metis + next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence vf: "supp v1 = {} \ supp v2 = {}" by force + then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \ v2 = (V_lit (L_num n2))" using infer_int infer_e_plusI by metis + then show ?case using reduce_let_plusI * by metis + next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence vf: "supp v1 = {} \ supp v2 = {}" by force + then obtain n1 and n2 where *: "v1 = V_lit (L_num n1) \ v2 = (V_lit (L_num n2))" using infer_int infer_e_leqI by metis + then show ?case using reduce_let_leqI * by metis + next + case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) + hence vf: "supp v1 = {} \ supp v2 = {}" by force + then obtain n1 and n2 where *: "v1 = V_lit n1 \ v2 = (V_lit n2)" using infer_lit infer_e_eqI by metis + then show ?case using reduce_let_eqI by blast + next + case (infer_e_appI \ \ \ \ \ f x b c \' s' v) + then show ?case using reduce_let_appI by metis + next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v) + then show ?case using reduce_let_appPI by metis + next + case (infer_e_fstI \ \ \ \ \ v z' b2 c z) + hence "supp v = {}" by force + then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_fstI infer_pair by metis + then show ?case using reduce_let_fstI * by metis + next + case (infer_e_sndI \ \ \ \ \ v z' b1 c z) + hence "supp v = {}" by force + then obtain v1 and v2 where "v = V_pair v1 v2" using infer_e_sndI infer_pair by metis + then show ?case using reduce_let_sndI * by metis + next + case (infer_e_lenI \ \ \ \ \ v z' c za) + hence "supp v = {}" by force + then obtain bvec where "v = V_lit (L_bitvec bvec)" using infer_e_lenI infer_bitvec by metis + then show ?case using reduce_let_lenI * by metis + next + case (infer_e_mvarI \ \ \ \ \ u) + hence "(u, \ z : b | c \) \ setD \" using infer_e_elims(10) by meson + then obtain v where "(u,v) \ set \" using infer_e_mvarI delta_sim_delta_lookup by meson + then show ?case using reduce_let_mvar by metis + next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence vf: "supp v1 = {} \ supp v2 = {}" by force + then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \ v2 = (V_lit (L_bitvec n2))" using infer_bitvec infer_e_concatI by metis + then show ?case using reduce_let_concatI * by metis + next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + hence vf: "supp v1 = {} \ supp v2 = {}" by force + then obtain n1 and n2 where *: "v1 = V_lit (L_bitvec n1) \ v2 = (V_lit (L_num n2))" using infer_bitvec infer_e_splitI check_int by metis + + have "0 \ n2 \ n2 \ int (length n1)" using check_v_range[OF _ * ] infer_e_splitI by simp + then obtain bv1 and bv2 where "split n2 n1 (bv1 , bv2)" using obtain_split by metis + then show ?case using reduce_let_splitI * by metis + qed +qed + +lemma check_css_lookup_branch_exist: + fixes s::s and cs::branch_s and css::branch_list and v::v + shows + "\; \; B; G; \ \ s \ \ \ True" and + "check_branch_s \ \ {||} GNil \ tid dc const v cs \ \ True" and + "\; \; \; \; \; tid; dclist; v \ css \ \ \ (dc, t) \ set dclist \ + \x' s'. Some (AS_branch dc x' s') = lookup_branch dc css" +proof(nominal_induct \ and \ and \ rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_branch_list_consI \ \ \ \ \ tid cons const v cs \ dclist css) + then show ?case using lookup_branch.simps check_branch_list_finalI by force +next + case (check_branch_list_finalI \ \ \ \ \ tid cons const v cs \) + then show ?case using lookup_branch.simps check_branch_list_finalI by force +qed(auto+) + +lemma progress_aux: + shows "\; \; \; \; \ \ s \ \ \ \ = {||} \ sble \ \ \ supp s \ atom ` fst ` setD \ \ \ \ \ \ \ \ + (\v. s = [v]\<^sup>s) \ (\\' s'. \ \ \\, s\ \ \\', s'\)" and + "\; \; {||}; \; \; tid; dc; const; v2 \ cs \ \ \ supp cs = {} \ True " + "\; \; {||}; \; \; tid; dclist; v2 \ css \ \ \ supp css = {} \ True " +proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ v \' \) + then show ?case by auto +next + case (check_letI x \ \ \ \ \ e \ z s b c) + hence "\; \; {||}; \; \ \ AS_let x e s \ \" using Typing.check_letI by meson + then show ?case using progress_let check_letI by metis +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + then show ?case by auto +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case by auto +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case by auto +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + have "supp v = {}" using check_ifI s_branch_s_branch_list.supp by auto + hence "v = V_lit L_true \ v = V_lit L_false" using check_bool_options check_ifI by auto + then show ?case using reduce_if_falseI reduce_if_trueI check_ifI by meson +next + case (check_let2I x \ \ \ G \ t s1 \ s2 ) + then consider " (\v. s1 = AS_val v)" | "(\\' a. \ \ \\, s1\ \ \\', a\)" by auto + then show ?case proof(cases) + case 1 + then show ?thesis using reduce_let2_valI by fast + next + case 2 + then show ?thesis using reduce_let2I check_let2I by meson + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + + obtain uu::u where uf: "atom uu \ (u,\,s) " using obtain_fresh by blast + obtain sa where " (uu \ u ) \ s = sa" by presburger + moreover have "atom uu \ s" using uf fresh_prod3 by auto + ultimately have "AS_var uu \' v sa = AS_var u \' v s" using s_branch_s_branch_list.eq_iff(7) Abs1_eq_iff(3)[of uu sa u s] by auto + + moreover have "atom uu \ \" using uf fresh_prod3 by auto + ultimately have "\ \ \\, AS_var u \' v s\ \ \(uu, v) # \, sa\" + using reduce_varI uf by metis + then show ?case by auto +next + case (check_assignI \ u \ P G v z \') + then show ?case using reduce_assignI by blast +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + obtain x::x where "atom x \ (s1,s2)" using obtain_fresh by metis + moreover obtain z::x where "atom z \ x" using obtain_fresh by metis + ultimately show ?case using reduce_whileI by fast +next + case (check_seqI P \ \ G \ s1 z s2 \) + thus ?case proof(cases "\v. s1 = AS_val v") + case True + then obtain v where v: "s1 = AS_val v" by blast + hence "supp v = {}" using check_seqI by auto + have "\z1 c1. P; \; G \ v \ (\ z1 : B_unit | c1 \)" proof - + obtain t where t:"P; \; G \ v \ t \ P; \ ; G \ t \ (\ z : B_unit | TRUE \)" + using v check_seqI(1) check_s_elims(1) by blast + obtain z1 and b1 and c1 where teq: "t = (\ z1 : b1 | c1 \)" using obtain_fresh_z by meson + hence "b1 = B_unit" using subtype_eq_base t by meson + thus ?thesis using t teq by fast + qed + then obtain z1 and c1 where "P ; \ ; G \ v \ (\ z1 : B_unit | c1 \)" by auto + hence "v = V_lit L_unit" using infer_v_unit_form \supp v = {}\ by simp + hence "s1 = AS_val (V_lit L_unit)" using v by auto + then show ?thesis using check_seqI reduce_seq1I by meson + next + case False + then show ?thesis using check_seqI reduce_seq2I + by (metis Un_subset_iff s_branch_s_branch_list.supp(9)) + qed + +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + hence "supp v = {}" by auto + + then obtain v' and dc and t::\ where v: "v = V_cons tid dc v' \ (dc, t) \ set dclist" + using check_v_tid_form check_caseI by metis + obtain z and b and c where teq: "t = (\ z : b | c \)" using obtain_fresh_z by meson + + moreover then obtain x' s' where "Some (AS_branch dc x' s') = lookup_branch dc cs" using v teq check_caseI check_css_lookup_branch_exist by metis + ultimately show ?case using reduce_caseI v check_caseI dc_of.cases by metis +next + case (check_assertI x \ \ \ \ \ c \ s) + hence sps: "supp s \ atom ` fst ` setD \" by auto + have "atom x \ c " using check_assertI by auto + have "atom x \ \" using check_assertI check_s_wf wfG_elims by metis + have "sble \ ((x, B_bool, c) #\<^sub>\ \)" proof - + obtain i' where i':" i' \ \ \ \; \ \ i'" using check_assertI sble_def by metis + obtain i::valuation where i:"i = i' ( x \ SBool True)" by auto + + have "i \ (x, B_bool, c) #\<^sub>\ \" proof - + have "i' \ c" using valid.simps i' check_assertI by metis + hence "i \ c" using is_satis_weakening_x i \atom x \ c\ by auto + moreover have "i \ \" using is_satis_g_weakening_x i' i check_assertI \atom x \ \\ by metis + ultimately show ?thesis using is_satis_g.simps i by auto + qed + moreover have "\ ; ((x, B_bool, c) #\<^sub>\ \) \ i" proof(rule wfI_cons) + show \ i' \ \ \ using i' by auto + show \ \ ; \ \ i'\ using i' by auto + show \i = i'(x \ SBool True)\ using i by auto + show \ \ \ SBool True: B_bool\ using wfRCV_BBoolI by auto + show \atom x \ \\ using check_assertI check_s_wf wfG_elims by auto + qed + ultimately show ?thesis using sble_def by auto + qed + then consider "(\v. s = [v]\<^sup>s)" | "(\\' a. \ \ \\, s\ \ \ \', a\)" using check_assertI sps by metis + hence "(\\' a. \ \ \\, ASSERT c IN s\ \ \\', a\)" proof(cases) + case 1 + then show ?thesis using reduce_assert1I by metis + next + case 2 + then show ?thesis using reduce_assert2I by metis + qed + thus ?case by auto +qed + +lemma progress: + assumes "\; \; \ \ \\, s\ \ \" + shows "(\v. s = [v]\<^sup>s) \ (\\' s'. \ \ \\, s\ \ \\', s'\)" +proof - + have "\; \; {||}; GNil; \ \ s \ \" and "\ \ \ \ \" + using config_type_elims[OF assms(1)] by auto+ + moreover hence "supp s \ atom ` fst ` setD \" using check_s_wf wfS_supp by fastforce + moreover have "sble \ GNil" using sble_def wfI_def is_satis_g.simps by simp + ultimately show ?thesis using progress_aux by blast +qed + +section \Safety\ + +lemma safety_stmt: + assumes "\ \ \\, s\ \\<^sup>* \\', s'\" and "\; \; \ \ \\, s\ \ \" + shows "(\v. s' = [v]\<^sup>s) \ (\\'' s''. \ \ \\', s'\ \ \\'', s''\)" + using preservation_many progress assms by meson + +lemma safety: + assumes "\ \PROG \ \ \ s\ \ \" and "\ \ \\_of \, s\ \\<^sup>* \\', s'\" + shows "(\v. s' = [v]\<^sup>s) \ (\\'' s''. \ \ \\', s'\ \ \\'', s''\)" + using assms config_type_prog_elims safety_stmt by metis + +end \ No newline at end of file diff --git a/thys/MiniSail/SubstMethods.thy b/thys/MiniSail/SubstMethods.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/SubstMethods.thy @@ -0,0 +1,71 @@ +theory SubstMethods + (* Its seems that it's best to load the Eisbach tools last *) + imports IVSubst WellformedL "HOL-Eisbach.Eisbach_Tools" +begin + +text \ + See Eisbach/Examples.thy as well as Eisbach User Manual. + +Freshness for various substitution situations. It seems that if undirected and we throw all the +facts at them to try to solve in one shot, the automatic methods are *sometimes* unable +to handle the different variants, so some guidance is needed. +First we split into subgoals using fresh\_prodN and intro conjI + +The 'add', for example, will be induction premises that will contain freshness facts or freshness conditions from +prior obtains + +Use different arguments for different things or just lump into one bucket\ + +method fresh_subst_mth_aux uses add = ( + (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: fresh_subst_gv_if[of "atom z" \ v x] add\) + | (match conclusion in "atom z \ (v'::v)[x::=v]\<^sub>v\<^sub>v" for z x v v' \ \auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_v_def add\ ) + | (match conclusion in "atom z \ (ce::ce)[x::=v]\<^sub>c\<^sub>e\<^sub>v" for z x v ce \ \auto simp add: fresh_subst_v_if subst_v_ce_def add\ ) + | (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: fresh_subst_v_if fresh_subst_dv_if add\ ) + | (match conclusion in "atom z \ \'[x::=v]\<^sub>\\<^sub>v @ \" for z x v \' \ \ \metis add \ ) + | (match conclusion in "atom z \ (\::\)[x::=v]\<^sub>\\<^sub>v" for z x v \ \ \auto simp add: v.fresh fresh_subst_v_if pure_fresh subst_v_\_def add\ ) + | (match conclusion in "atom z \ ({||} :: bv fset)" for z \ \auto simp add: fresh_empty_fset\) + (* tbc delta and types *) + | (auto simp add: add x_fresh_b pure_fresh) (* Cases where there is no subst and so can most likely get what we want from induction premises *) + ) + +method fresh_mth uses add = ( + (unfold fresh_prodN, intro conjI)?, + (fresh_subst_mth_aux add: add)+) + + +notepad +begin + fix \::\ and z::x and x::x and v::v and \::\ and v'::v and w::x and tyid::string and dc::string and b::b and ce::ce and bv::bv + + assume as:"atom z \ (\,v',\, v,w,ce) \ atom bv \ (\,v',\, v,w,ce,b) " + + have "atom z \ \[x::=v]\<^sub>\\<^sub>v" + by (fresh_mth add: as) + + hence "atom z \ v'[x::=v]\<^sub>v\<^sub>v" + by (fresh_mth add: as) + + hence "atom z \ \" + by (fresh_mth add: as) + + hence "atom z \ \" + by (fresh_mth add: as) + + hence "atom z \ (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v" + using as by auto + + hence "atom bv \ (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v" + using as by auto + + have "atom z \ (\,\[x::=v]\<^sub>\\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v, (CE_val v == ce)[x::=v]\<^sub>c\<^sub>v) " + by (fresh_mth add: as) + + have "atom bv \ (\,\[x::=v]\<^sub>\\<^sub>v,v'[x::=v]\<^sub>v\<^sub>v,w, V_pair v v, V_consp tyid dc b v) " + by (fresh_mth add: as) + +end + + + + +end \ No newline at end of file diff --git a/thys/MiniSail/Syntax.thy b/thys/MiniSail/Syntax.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Syntax.thy @@ -0,0 +1,1174 @@ +(*<*) +theory Syntax + imports "Nominal2.Nominal2" "Nominal-Utils" +begin + (*>*) + +chapter \Syntax\ + +text \Syntax of MiniSail programs and the contexts we use in judgements.\ + +section \Program Syntax\ + +subsection \AST Datatypes\ + +type_synonym num_nat = nat + +atom_decl x (* Immutable variables*) +atom_decl u (* Mutable variables *) +atom_decl bv (* Basic-type variables *) + +type_synonym f = string (* Function name *) +type_synonym dc = string (* Data constructor *) +type_synonym tyid = string + +text \Basic types. Types without refinement constraints\ +nominal_datatype "b" = + B_int | B_bool | B_id tyid +| B_pair b b ("[ _ , _ ]\<^sup>b") +| B_unit | B_bitvec | B_var bv +| B_app tyid b + +nominal_datatype bit = BitOne | BitZero + +text \ Literals \ +nominal_datatype "l" = + L_num "int" | L_true | L_false | L_unit | L_bitvec "bit list" + +text \ Values. We include a type identifier, tyid, in the literal for constructors +to make typing and well-formedness checking easier \ + +nominal_datatype "v" = + V_lit "l" ( "[ _ ]\<^sup>v") + | V_var "x" ( "[ _ ]\<^sup>v") + | V_pair "v" "v" ( "[ _ , _ ]\<^sup>v") + | V_cons tyid dc "v" + | V_consp tyid dc b "v" + +text \ Binary Operations \ +nominal_datatype "opp" = Plus ( "plus") | LEq ("leq") | Eq ("eq") + +text \ Expressions \ +nominal_datatype "e" = + AE_val "v" ( "[ _ ]\<^sup>e" ) + | AE_app "f" "v" ( "[ _ ( _ ) ]\<^sup>e" ) + | AE_appP "f" "b" "v" ( "[_ [ _ ] ( _ )]\<^sup>e" ) + | AE_op "opp" "v" "v" ( "[ _ _ _ ]\<^sup>e" ) + | AE_concat v v ( "[ _ @@ _ ]\<^sup>e" ) + | AE_fst "v" ( "[#1_ ]\<^sup>e" ) + | AE_snd "v" ( "[#2_ ]\<^sup>e" ) + | AE_mvar "u" ( "[ _ ]\<^sup>e" ) + | AE_len "v" ( "[| _ |]\<^sup>e" ) + | AE_split "v" "v" ( "[ _ / _ ]\<^sup>e" ) + +text \ Expressions for constraints\ +nominal_datatype "ce" = + CE_val "v" ( "[ _ ]\<^sup>c\<^sup>e" ) + | CE_op "opp" "ce" "ce" ( "[ _ _ _ ]\<^sup>c\<^sup>e" ) + | CE_concat ce ce ( "[ _ @@ _ ]\<^sup>c\<^sup>e" ) + | CE_fst "ce" ( "[#1_]\<^sup>c\<^sup>e" ) + | CE_snd "ce" ( "[#2_]\<^sup>c\<^sup>e" ) + | CE_len "ce" ( "[| _ |]\<^sup>c\<^sup>e" ) + +text \ Constraints \ +nominal_datatype "c" = + C_true ( "TRUE" [] 50 ) + | C_false ( "FALSE" [] 50 ) + | C_conj "c" "c" ("_ AND _ " [50, 50] 50) + | C_disj "c" "c" ("_ OR _ " [50,50] 50) + | C_not "c" ( "\ _ " [] 50 ) + | C_imp "c" "c" ("_ IMP _ " [50, 50] 50) + | C_eq "ce" "ce" ("_ == _ " [50, 50] 50) + +text \ Refined types \ +nominal_datatype "\" = + T_refined_type x::x b c::c binds x in c ("\ _ : _ | _ \" [50, 50] 1000) + +text \ Statements \ + +nominal_datatype + s = + AS_val v ( "[_]\<^sup>s") + | AS_let x::x e s::s binds x in s ( "(LET _ = _ IN _)") + | AS_let2 x::x \ s s::s binds x in s ( "(LET _ : _ = _ IN _)") + | AS_if v s s ( "(IF _ THEN _ ELSE _)" [0, 61, 0] 61) + | AS_var u::u \ v s::s binds u in s ( "(VAR _ : _ = _ IN _)") + | AS_assign u v ( "(_ ::= _)") + | AS_match v branch_list ( "(MATCH _ WITH { _ })") + | AS_while s s ( "(WHILE _ DO { _ } )" [0, 0] 61) + | AS_seq s s ( "( _ ;; _ )" [1000, 61] 61) + | AS_assert c s ( "(ASSERT _ IN _ )" ) + and branch_s = + AS_branch dc x::x s::s binds x in s ( "( _ _ \ _ )") + and branch_list = + AS_final branch_s ( "{ _ }" ) + | AS_cons branch_s branch_list ( "( _ | _ )") + +text \ Function and union type definitions \ + +nominal_datatype "fun_typ" = + AF_fun_typ x::x "b" c::c \::\ s::s binds x in c \ s + +nominal_datatype "fun_typ_q" = + AF_fun_typ_some bv::bv ft::fun_typ binds bv in ft + | AF_fun_typ_none fun_typ + +nominal_datatype "fun_def" = AF_fundef f fun_typ_q + +nominal_datatype "type_def" = + AF_typedef string "(string * \) list" + | AF_typedef_poly string bv::bv dclist::"(string * \) list" binds bv in dclist + +lemma check_typedef_poly: + "AF_typedef_poly ''option'' bv [ (''None'', \ zz : B_unit | TRUE \), (''Some'', \ zz : B_var bv | TRUE \) ] = + AF_typedef_poly ''option'' bv2 [ (''None'', \ zz : B_unit | TRUE \), (''Some'', \ zz : B_var bv2 | TRUE \) ]" + by auto + +nominal_datatype "var_def" = AV_def u \ v + +text \ Programs \ +nominal_datatype "p" = + AP_prog "type_def list" "fun_def list" "var_def list" "s" ("PROG _ _ _ _") + +declare l.supp [simp] v.supp [simp] e.supp [simp] s_branch_s_branch_list.supp [simp] \.supp [simp] c.supp [simp] b.supp[simp] + +subsection \Lemmas\ + +text \These lemmas deal primarily with freshness and alpha-equivalence\ + +subsubsection \Atoms\ + +lemma x_not_in_u_atoms[simp]: + fixes u::u and x::x and us::"u set" + shows "atom x \ atom`us" + by (simp add: image_iff) + +lemma x_fresh_u[simp]: + fixes u::u and x::x + shows "atom x \ u" + by auto + +lemma x_not_in_b_set[simp]: + fixes x::x and bs::"bv fset" + shows "atom x \ supp bs" + by(induct bs,auto, simp add: supp_finsert supp_at_base) + +lemma x_fresh_b[simp]: + fixes x::x and b::b + shows "atom x \ b" + apply (induct b rule: b.induct, auto simp: pure_supp) + using pure_supp fresh_def by blast+ + +lemma x_fresh_bv[simp]: + fixes x::x and bv::bv + shows "atom x \ bv" + using fresh_def supp_at_base by auto + +lemma u_not_in_x_atoms[simp]: + fixes u::u and x::x and xs::"x set" + shows "atom u \ atom`xs" + by (simp add: image_iff) + +lemma bv_not_in_x_atoms[simp]: + fixes bv::bv and x::x and xs::"x set" + shows "atom bv \ atom`xs" + by (simp add: image_iff) + +lemma u_not_in_b_atoms[simp]: + fixes b :: b and u::u + shows "atom u \ supp b" + by (induct b rule: b.induct,auto simp: pure_supp supp_at_base) + +lemma u_not_in_b_set[simp]: + fixes u::u and bs::"bv fset" + shows "atom u \ supp bs" + by(induct bs, auto simp add: supp_at_base supp_finsert) + +lemma u_fresh_b[simp]: + fixes x::u and b::b + shows "atom x \ b" + by(induct b rule: b.induct, auto simp: pure_fresh ) + +lemma supp_b_v_disjoint: + fixes x::x and bv::bv + shows "supp (V_var x) \ supp (B_var bv) = {}" + by (simp add: supp_at_base) + +lemma supp_b_u_disjoint[simp]: + fixes b::b and u::u + shows "supp u \ supp b = {}" + by(nominal_induct b rule:b.strong_induct,(auto simp add: pure_supp b.supp supp_at_base)+) + +lemma u_fresh_bv[simp]: + fixes u::u and b::bv + shows "atom u \ b" + using fresh_at_base by simp + +subsubsection \Basic Types\ + +nominal_function b_of :: "\ \ b" where + "b_of \ z : b | c \ = b" + apply(auto,simp add: eqvt_def b_of_graph_aux_def ) + by (meson \.exhaust) +nominal_termination (eqvt) by lexicographic_order + +lemma supp_b_empty[simp]: + fixes b :: b and x::x + shows "atom x \ supp b" + by (induct b rule: b.induct, auto simp: pure_supp supp_at_base x_not_in_b_set) + +lemma flip_b_id[simp]: + fixes x::x and b::b + shows "(x \ x') \ b = b" + by(rule flip_fresh_fresh, auto simp add: fresh_def) + +lemma flip_x_b_cancel[simp]: + fixes x::x and y::x and b::b and bv::bv + shows "(x \ y) \ b = b" and "(x \ y) \ bv = bv" + using flip_b_id apply simp + by (metis b.eq_iff(7) b.perm_simps(7) flip_b_id) + +lemma flip_bv_x_cancel[simp]: + fixes bv::bv and z::bv and x::x + shows "(bv \ z) \ x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto + +lemma flip_bv_u_cancel[simp]: + fixes bv::bv and z::bv and x::u + shows "(bv \ z) \ x = x" using flip_fresh_fresh[of bv x z] fresh_at_base by auto + +subsubsection \Literals\ + +lemma supp_bitvec_empty: + fixes bv::"bit list" + shows "supp bv = {}" +proof(induct bv) + case Nil + then show ?case using supp_Nil by auto +next + case (Cons a bv) + then show ?case using supp_Cons bit.supp + by (metis (mono_tags, hide_lams) bit.strong_exhaust l.supp(5) sup_bot.right_neutral) +qed + +lemma bitvec_pure[simp]: + fixes bv::"bit list" and x::x + shows "atom x \ bv" using fresh_def supp_bitvec_empty by auto + +lemma supp_l_empty[simp]: + fixes l:: l + shows "supp (V_lit l) = {}" + by(nominal_induct l rule: l.strong_induct, + auto simp add: l.strong_exhaust pure_supp v.fv_defs supp_bitvec_empty) + +lemma type_l_nosupp[simp]: + fixes x::x and l::l + shows "atom x \ supp (\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \)" + using supp_at_base supp_l_empty ce.supp(1) c.supp \.supp by force + +lemma flip_bitvec0: + fixes x::"bit list" + assumes "atom c \ (z, x, z')" + shows "(z \ c) \ x = (z' \ c) \ x" +proof - + have "atom z \ x" and "atom z' \ x" + using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+ + moreover have "atom c \ x" using supp_bitvec_empty fresh_def by auto + ultimately show ?thesis using assms flip_fresh_fresh by metis +qed + +lemma flip_bitvec: + assumes "atom c \ (z, L_bitvec x, z')" + shows "(z \ c) \ x = (z' \ c) \ x" +proof - + have "atom z \ x" and "atom z' \ x" + using flip_fresh_fresh assms supp_bitvec_empty fresh_def by blast+ + moreover have "atom c \ x" using supp_bitvec_empty fresh_def by auto + ultimately show ?thesis using assms flip_fresh_fresh by metis +qed + +lemma type_l_eq: + shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \ = (\ z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == [V_lit l]\<^sup>c\<^sup>e \)" + by(auto,nominal_induct l rule: l.strong_induct,auto, metis permute_pure, auto simp add: flip_bitvec) + +lemma flip_l_eq: + fixes x::l + shows "(z \ c) \ x = (z' \ c) \ x" +proof - + have "atom z \ x" and "atom c \ x" and "atom z' \ x" + using flip_fresh_fresh fresh_def supp_l_empty by fastforce+ + thus ?thesis using flip_fresh_fresh by metis +qed + +lemma flip_l_eq1: + fixes x::l + assumes "(z \ c) \ x = (z' \ c) \ x'" + shows "x' = x" +proof - + have "atom z \ x" and "atom c \ x'" and "atom c \ x" and "atom z' \ x'" + using flip_fresh_fresh fresh_def supp_l_empty by fastforce+ + thus ?thesis using flip_fresh_fresh assms by metis +qed + +subsubsection \Types\ + +lemma flip_base_eq: + fixes b::b and x::x and y::x + shows "(x \ y) \ b = b" + using b.fresh by (simp add: flip_fresh_fresh fresh_def) + +text \ Obtain an alpha-equivalent type where the bound variable is fresh in some term t \ +lemma has_fresh_z0: + fixes t::"'b::fs" + shows "\z. atom z \ (c',t) \ (\z' : b | c' \) = (\ z : b | (z \ z' ) \ c' \)" +proof - + obtain z::x where fr: "atom z \ (c',t)" using obtain_fresh by blast + moreover hence "(\ z' : b | c' \) = (\ z : b | (z \ z') \ c' \)" + using \.eq_iff Abs1_eq_iff + by (metis flip_commute flip_fresh_fresh fresh_PairD(1)) + ultimately show ?thesis by fastforce +qed + +lemma has_fresh_z: + fixes t::"'b::fs" + shows "\z b c. atom z \ t \ \ = \ z : b | c \" +proof - + obtain z' and b and c' where teq: "\ = (\ z' : b | c' \)" using \.exhaust by blast + obtain z::x where fr: "atom z \ (t,c')" using obtain_fresh by blast + hence "(\ z' : b | c' \) = (\ z : b | (z \ z') \ c' \)" using \.eq_iff Abs1_eq_iff + flip_commute flip_fresh_fresh fresh_PairD(1) by (metis fresh_PairD(2)) + hence "atom z \ t \ \ = (\ z : b | (z \ z') \ c' \)" using fr teq by force + thus ?thesis using teq fr by fast +qed + +lemma obtain_fresh_z: + fixes t::"'b::fs" + obtains z and b and c where "atom z \ t \ \ = \ z : b | c \" + using has_fresh_z by blast + +lemma has_fresh_z2: + fixes t::"'b::fs" + shows "\z c. atom z \ t \ \ = \ z : b_of \ | c \" +proof - + obtain z and b and c where "atom z \ t \ \ = \ z : b | c \" using obtain_fresh_z by metis + moreover then have "b_of \ = b" using \.eq_iff by simp + ultimately show ?thesis using obtain_fresh_z \.eq_iff by auto +qed + +lemma obtain_fresh_z2: + fixes t::"'b::fs" + obtains z and c where "atom z \ t \ \ = \ z : b_of \ | c \" + using has_fresh_z2 by blast + +subsubsection \Values\ + +lemma u_notin_supp_v[simp]: + fixes u::u and v::v + shows "atom u \ supp v" +proof(nominal_induct v rule: v.strong_induct) + case (V_lit l) + then show ?case using supp_l_empty by auto +next + case (V_var x) + then show ?case + by (simp add: supp_at_base) +next + case (V_pair v1 v2) + then show ?case by auto +next + case (V_cons tyid list v) + then show ?case using pure_supp by auto +next + case (V_consp tyid list b v) + then show ?case using pure_supp by auto +qed + +lemma u_fresh_xv[simp]: + fixes u::u and x::x and v::v + shows "atom u \ (x,v)" +proof - + have "atom u \ x" using fresh_def by fastforce + moreover have "atom u \ v" using fresh_def u_notin_supp_v by metis + ultimately show ?thesis using fresh_prod2 by auto +qed + +text \ Part of an effort to make the proofs across inductive cases more uniform by distilling +the non-uniform parts into lemmas like this\ +lemma v_flip_eq: + fixes v::v and va::v and x::x and c::x + assumes "atom c \ (v, va)" and "atom c \ (x, xa, v, va)" and "(x \ c) \ v = (xa \ c) \ va" + shows "((v = V_lit l \ (\l'. va = V_lit l' \ (x \ c) \ l = (xa \ c) \ l'))) \ + ((v = V_var y \ (\y'. va = V_var y' \ (x \ c) \ y = (xa \ c) \ y'))) \ + ((v = V_pair vone vtwo \ (\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2'))) \ + ((v = V_cons tyid dc vone \ (\v1'. va = V_cons tyid dc v1'\ (x \ c) \ vone = (xa \ c) \ v1'))) \ + ((v = V_consp tyid dc b vone \ (\v1'. va = V_consp tyid dc b v1'\ (x \ c) \ vone = (xa \ c) \ v1')))" + using assms proof(nominal_induct v rule:v.strong_induct) + case (V_lit l) + then show ?case using assms v.perm_simps + empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh + by (metis permute_swap_cancel2 v.distinct) +next + case (V_var x) + then show ?case using assms v.perm_simps + empty_iff flip_def fresh_def fresh_permute_iff supp_l_empty swap_fresh_fresh v.fresh + by (metis permute_swap_cancel2 v.distinct) +next + case (V_pair v1 v2) + have " (V_pair v1 v2 = V_pair vone vtwo \ (\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2'))" proof + assume "V_pair v1 v2= V_pair vone vtwo" + thus "(\v1' v2'. va = V_pair v1' v2' \ (x \ c) \ vone = (xa \ c) \ v1' \ (x \ c) \ vtwo = (xa \ c) \ v2')" + using V_pair assms + by (metis (no_types, hide_lams) flip_def permute_swap_cancel v.perm_simps(3)) + qed + thus ?case using V_pair by auto +next + case (V_cons tyid dc v1) + have " (V_cons tyid dc v1 = V_cons tyid dc vone \ (\ v1'. va = V_cons tyid dc v1' \ (x \ c) \ vone = (xa \ c) \ v1'))" proof + assume as: "V_cons tyid dc v1 = V_cons tyid dc vone" + hence "(x \ c) \ (V_cons tyid dc vone) = V_cons tyid dc ((x \ c) \ vone)" proof - + have "(x \ c) \ dc = dc" using pure_permute_id by metis + moreover have "(x \ c) \ tyid = tyid" using pure_permute_id by metis + ultimately show ?thesis using v.perm_simps(4) by simp + qed + then obtain v1' where "(xa \ c) \ va = V_cons tyid dc v1' \ (x \ c) \ vone = v1'" using assms V_cons + using as by fastforce + hence " va = V_cons tyid dc ((xa \ c) \ v1') \ (x \ c) \ vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh + by (metis pure_fresh v.perm_simps(4)) + + thus "(\v1'. va = V_cons tyid dc v1' \ (x \ c) \ vone = (xa \ c) \ v1')" + using V_cons assms by simp + qed + thus ?case using V_cons by auto +next + case (V_consp tyid dc b v1) + have " (V_consp tyid dc b v1 = V_consp tyid dc b vone \ (\ v1'. va = V_consp tyid dc b v1' \ (x \ c) \ vone = (xa \ c) \ v1'))" proof + assume as: "V_consp tyid dc b v1 = V_consp tyid dc b vone" + hence "(x \ c) \ (V_consp tyid dc b vone) = V_consp tyid dc b ((x \ c) \ vone)" proof - + have "(x \ c) \ dc = dc" using pure_permute_id by metis + moreover have "(x \ c) \ tyid = tyid" using pure_permute_id by metis + ultimately show ?thesis using v.perm_simps(4) by simp + qed + then obtain v1' where "(xa \ c) \ va = V_consp tyid dc b v1' \ (x \ c) \ vone = v1'" using assms V_consp + using as by fastforce + hence " va = V_consp tyid dc b ((xa \ c) \ v1') \ (x \ c) \ vone = v1'" using permute_flip_cancel empty_iff flip_def fresh_def supp_b_empty swap_fresh_fresh + pure_fresh v.perm_simps + by (metis (mono_tags, hide_lams)) + thus "(\v1'. va = V_consp tyid dc b v1' \ (x \ c) \ vone = (xa \ c) \ v1')" + using V_consp assms by simp + qed + thus ?case using V_consp by auto +qed + +lemma flip_eq: + fixes x::x and xa::x and s::"'a::fs" and sa::"'a::fs" + assumes "(\c. atom c \ (s, sa) \ atom c \ (x, xa, s, sa) \ (x \ c) \ s = (xa \ c) \ sa)" and "x \ xa" + shows "(x \ xa) \ s = sa" +proof - + have " ([[atom x]]lst. s = [[atom xa]]lst. sa)" using assms Abs1_eq_iff_all by simp + hence "(xa = x \ sa = s \ xa \ x \ sa = (xa \ x) \ s \ atom xa \ s)" using assms Abs1_eq_iff[of xa sa x s] by simp + thus ?thesis using assms + by (metis flip_commute) +qed + +lemma swap_v_supp: + fixes v::v and d::x and z::x + assumes "atom d \ v" + shows "supp ((z \ d ) \ v) \ supp v - { atom z } \ { atom d}" + using assms +proof(nominal_induct v rule:v.strong_induct) + case (V_lit l) + then show ?case using l.supp by (metis supp_l_empty empty_subsetI l.strong_exhaust pure_supp supp_eqvt v.supp) +next + case (V_var x) + hence "d \ x" using fresh_def by fastforce + thus ?case apply(cases "z = x") using supp_at_base V_var \d\x\ by fastforce+ +next + case (V_cons tyid dc v) + show ?case using v.supp(4) pure_supp + using V_cons.hyps V_cons.prems fresh_def by auto +next + case (V_consp tyid dc b v) + show ?case using v.supp(4) pure_supp + using V_consp.hyps V_consp.prems fresh_def by auto +qed(force+) + +subsubsection \Expressions\ + +lemma swap_e_supp: + fixes e::e and d::x and z::x + assumes "atom d \ e" + shows "supp ((z \ d ) \ e) \ supp e - { atom z } \ { atom d}" + using assms +proof(nominal_induct e rule:e.strong_induct) + case (AE_val v) + then show ?case using swap_v_supp by simp +next + case (AE_app f v) + then show ?case using swap_v_supp by (simp add: pure_supp) +next + case (AE_appP b f v) + hence df: "atom d \ v" using fresh_def e.supp by force + have "supp ((z \ d ) \ (AE_appP b f v)) = supp (AE_appP b f ((z \ d ) \ v))" using e.supp + by (metis b.eq_iff(3) b.perm_simps(3) e.perm_simps(3) flip_b_id) + also have "... = supp b \ supp f \ supp ((z \ d ) \ v)" using e.supp by auto + also have "... \ supp b \ supp f \ supp v - { atom z } \ { atom d}" using swap_v_supp[OF df] pure_supp by auto + finally show ?case using e.supp by auto +next + case (AE_op opp v1 v2) + hence df: "atom d \ v1 \ atom d \ v2" using fresh_def e.supp by force + have "((z \ d ) \ (AE_op opp v1 v2)) = AE_op opp ((z \ d ) \ v1) ((z \ d ) \ v2)" using + e.perm_simps flip_commute opp.perm_simps AE_op opp.strong_exhaust pure_supp + by (metis (full_types)) + + hence "supp ((z \ d) \ AE_op opp v1 v2) = supp (AE_op opp ((z \ d) \v1) ((z \ d) \v2))" by simp + also have "... = supp ((z \ d) \v1) \ supp ((z \ d) \v2)" using e.supp + by (metis (mono_tags, hide_lams) opp.strong_exhaust opp.supp sup_bot.left_neutral) + also have "... \ (supp v1 - { atom z } \ { atom d}) \ (supp v2 - { atom z } \ { atom d})" using swap_v_supp AE_op df by blast + finally show ?case using e.supp opp.supp by blast +next + case (AE_fst v) + then show ?case using swap_v_supp by auto +next + case (AE_snd v) + then show ?case using swap_v_supp by auto +next + case (AE_mvar u) + then show ?case using + Diff_empty Diff_insert0 Un_upper1 atom_x_sort flip_def flip_fresh_fresh fresh_def set_eq_subset supp_eqvt swap_set_in_eq + by (metis sort_of_atom_eq) +next + case (AE_len v) + then show ?case using swap_v_supp by auto +next + case (AE_concat v1 v2) + then show ?case using swap_v_supp by auto +next + case (AE_split v1 v2) + then show ?case using swap_v_supp by auto +qed + +lemma swap_ce_supp: + fixes e::ce and d::x and z::x + assumes "atom d \ e" + shows "supp ((z \ d ) \ e) \ supp e - { atom z } \ { atom d}" + using assms +proof(nominal_induct e rule:ce.strong_induct) + case (CE_val v) + then show ?case using swap_v_supp ce.fresh ce.supp by simp +next + case (CE_op opp v1 v2) + hence df: "atom d \ v1 \ atom d \ v2" using fresh_def e.supp by force + have "((z \ d ) \ (CE_op opp v1 v2)) = CE_op opp ((z \ d ) \ v1) ((z \ d ) \ v2)" using + ce.perm_simps flip_commute opp.perm_simps CE_op opp.strong_exhaust x_fresh_b pure_supp + by (metis (full_types)) + + hence "supp ((z \ d) \ CE_op opp v1 v2) = supp (CE_op opp ((z \ d) \v1) ((z \ d) \v2))" by simp + also have "... = supp ((z \ d) \v1) \ supp ((z \ d) \v2)" using ce.supp + by (metis (mono_tags, hide_lams) opp.strong_exhaust opp.supp sup_bot.left_neutral) + also have "... \ (supp v1 - { atom z } \ { atom d}) \ (supp v2 - { atom z } \ { atom d})" using swap_v_supp CE_op df by blast + finally show ?case using ce.supp opp.supp by blast +next + case (CE_fst v) + then show ?case using ce.supp ce.fresh swap_v_supp by auto +next + case (CE_snd v) + then show ?case using ce.supp ce.fresh swap_v_supp by auto +next + case (CE_len v) + then show ?case using ce.supp ce.fresh swap_v_supp by auto +next + case (CE_concat v1 v2) + then show ?case using ce.supp ce.fresh swap_v_supp ce.perm_simps + proof - + have "\x v xa. \ atom (x::x) \ (v::v) \ supp ((xa \ x) \ v) \ supp v - {atom xa} \ {atom x}" + by (meson swap_v_supp) (* 0.0 ms *) + then show ?thesis + using CE_concat ce.supp by auto + qed +qed + +lemma swap_c_supp: + fixes c::c and d::x and z::x + assumes "atom d \ c" + shows "supp ((z \ d ) \ c) \ supp c - { atom z } \ { atom d}" + using assms +proof(nominal_induct c rule:c.strong_induct) + case (C_eq e1 e2) + then show ?case using swap_ce_supp by auto +qed(auto+) + +lemma type_e_eq: + assumes "atom z \ e" and "atom z' \ e" + shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \ = (\ z' : b | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \)" + by (auto,metis (full_types) assms(1) assms(2) flip_fresh_fresh fresh_PairD(1) fresh_PairD(2)) + +lemma type_e_eq2: + assumes "atom z \ e" and "atom z' \ e" and "b=b'" + shows "\ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == e \ = (\ z' : b' | [[z']\<^sup>v]\<^sup>c\<^sup>e == e \)" + using assms type_e_eq by fast + +lemma e_flip_eq: + fixes e::e and ea::e + assumes "atom c \ (e, ea)" and "atom c \ (x, xa, e, ea)" and "(x \ c) \ e = (xa \ c) \ ea" + shows "(e = AE_val w \ (\w'. ea = AE_val w' \ (x \ c) \ w = (xa \ c) \ w')) \ + (e = AE_op opp v1 v2 \ (\v1' v2'. ea = AS_op opp v1' v2' \ (x \ c) \ v1 = (xa \ c) \ v1') \ (x \ c) \ v2 = (xa \ c) \ v2') \ + (e = AE_fst v \ (\v'. ea = AE_fst v' \ (x \ c) \ v = (xa \ c) \ v')) \ + (e = AE_snd v \ (\v'. ea = AE_snd v' \ (x \ c) \ v = (xa \ c) \ v')) \ + (e = AE_len v \ (\v'. ea = AE_len v' \ (x \ c) \ v = (xa \ c) \ v')) \ + (e = AE_concat v1 v2 \ (\v1' v2'. ea = AS_concat v1' v2' \ (x \ c) \ v1 = (xa \ c) \ v1') \ (x \ c) \ v2 = (xa \ c) \ v2') \ + (e = AE_app f v \ (\v'. ea = AE_app f v' \ (x \ c) \ v = (xa \ c) \ v'))" + by (metis assms e.perm_simps permute_flip_cancel2) + +lemma fresh_opp_all: + fixes opp::opp + shows "z \ opp" + using e.fresh opp.exhaust opp.fresh by metis + +lemma fresh_e_opp_all: + shows "(z \ v1 \ z \ v2) = z \ AE_op opp v1 v2" + using e.fresh opp.exhaust opp.fresh fresh_opp_all by simp + +lemma fresh_e_opp: + fixes z::x + assumes "atom z \ v1 \ atom z \ v2" + shows "atom z \ AE_op opp v1 v2" + using e.fresh opp.exhaust opp.fresh opp.supp by (metis assms) + +subsubsection \Statements\ + +lemma branch_s_flip_eq: + fixes v::v and va::v + assumes "atom c \ (v, va)" and "atom c \ (x, xa, v, va)" and "(x \ c) \ s = (xa \ c) \ sa" + shows "(s = AS_val w \ (\w'. sa = AS_val w' \ (x \ c) \ w = (xa \ c) \ w')) \ + (s = AS_seq s1 s2 \ (\s1' s2'. sa = AS_seq s1' s2' \ (x \ c) \ s1 = (xa \ c) \ s1') \ (x \ c) \ s2 = (xa \ c) \ s2') \ + (s = AS_if v s1 s2 \ (\v' s1' s2'. sa = AS_if seq s1' s2' \ (x \ c) \ s1 = (xa \ c) \ s1') \ (x \ c) \ s2 = (xa \ c) \ s2' \ (x \ c) \ c = (xa \ c) \ v')" + by (metis assms s_branch_s_branch_list.perm_simps permute_flip_cancel2) + +section \Context Syntax\ + +subsection \Datatypes\ + +text \Type and function/type definition contexts\ +type_synonym \ = "fun_def list" +type_synonym \ = "type_def list" +type_synonym \ = "bv fset" + +datatype \ = + GNil + | GCons "x*b*c" \ (infixr "#\<^sub>\" 65) + +datatype \ = + DNil ("[]\<^sub>\") + | DCons "u*\" \ (infixr "#\<^sub>\" 65) + +subsection \Functions and Lemmas\ + +lemma \_induct [case_names GNil GCons] : "P GNil \ (\x b c \'. P \' \ P ((x,b,c) #\<^sub>\ \')) \ P \" +proof(induct \ rule:\.induct) + case GNil + then show ?case by auto +next + case (GCons x1 x2) + then obtain x and b and c where "x1=(x,b,c)" using prod_cases3 by blast + then show ?case using GCons by presburger +qed + +instantiation \ :: pt +begin + +primrec permute_\ + where + DNil_eqvt: "p \ DNil = DNil" + | DCons_eqvt: "p \ (x #\<^sub>\ xs) = p \ x #\<^sub>\ p \ (xs::\)" + +instance by standard (induct_tac [!] x, simp_all) +end + +lemmas [eqvt] = permute_\.simps + +lemma \_induct [case_names DNil DCons] : "P DNil \ (\u t \'. P \' \ P ((u,t) #\<^sub>\ \')) \ P \" +proof(induct \ rule: \.induct) + case DNil + then show ?case by auto +next + case (DCons x1 x2) + then obtain u and t where "x1=(u,t)" by fastforce + then show ?case using DCons by presburger +qed + +lemma \_induct [case_names PNil PConsNone PConsSome] : "P [] \ (\f x b c \ s' \'. P \' \ P ((AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s'))) # \')) \ + (\f bv x b c \ s' \'. P \' \ P ((AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s'))) # \')) \ P \" +proof(induct \ rule: list.induct) + case Nil + then show ?case by auto +next + case (Cons x1 x2) + then obtain f and t where ft: "x1 = (AF_fundef f t)" + by (meson fun_def.exhaust) + then show ?case proof(nominal_induct t rule:fun_typ_q.strong_induct) + case (AF_fun_typ_some bv ft) + then show ?case using Cons ft + by (metis fun_typ.exhaust) + next + case (AF_fun_typ_none ft) + then show ?case using Cons ft + by (metis fun_typ.exhaust) + qed +qed + +lemma \_induct [case_names TNil AF_typedef AF_typedef_poly] : "P [] \ (\tid dclist \'. P \' \ P ((AF_typedef tid dclist) # \')) \ + (\tid bv dclist \'. P \' \ P ((AF_typedef_poly tid bv dclist ) # \')) \ P \" +proof(induct \ rule: list.induct) + case Nil + then show ?case by auto +next + case (Cons td T) + show ?case by(cases td rule: type_def.exhaust, (simp add: Cons)+) +qed + +instantiation \ :: pt +begin + +primrec permute_\ + where + GNil_eqvt: "p \ GNil = GNil" + | GCons_eqvt: "p \ (x #\<^sub>\ xs) = p \ x #\<^sub>\ p \ (xs::\)" + +instance by standard (induct_tac [!] x, simp_all) +end + +lemmas [eqvt] = permute_\.simps + +lemma G_cons_eqvt[simp]: + fixes \::\ + shows "p \ ((x,b,c) #\<^sub>\ \) = ((p \ x, p \ b , p \ c) #\<^sub>\ (p \ \ ))" (is "?A = ?B" ) + using Cons_eqvt triple_eqvt supp_b_empty by simp + +lemma G_cons_flip[simp]: + fixes x::x and \::\ + shows "(x\x') \ ((x'',b,c) #\<^sub>\ \) = (((x\x') \ x'', b , (x\x') \ c) #\<^sub>\ ((x\x') \ \ ))" + using Cons_eqvt triple_eqvt supp_b_empty by auto + +lemma G_cons_flip_fresh[simp]: + fixes x::x and \::\ + assumes "atom x \ (c,\)" and "atom x' \ (c,\)" + shows "(x\x') \ ((x',b,c) #\<^sub>\ \) = ((x, b , c) #\<^sub>\ \ )" + using G_cons_flip flip_fresh_fresh assms by force + +lemma G_cons_flip_fresh2[simp]: + fixes x::x and \::\ + assumes "atom x \ (c,\)" and "atom x' \ (c,\)" + shows "(x\x') \ ((x,b,c) #\<^sub>\ \) = ((x', b , c) #\<^sub>\ \ )" + using G_cons_flip flip_fresh_fresh assms by force + +lemma G_cons_flip_fresh3[simp]: + fixes x::x and \::\ + assumes "atom x \ \" and "atom x' \ \" + shows "(x\x') \ ((x',b,c) #\<^sub>\ \) = ((x, b , (x\x') \ c) #\<^sub>\ \ )" + using G_cons_flip flip_fresh_fresh assms by force + +lemma neq_GNil_conv: "(xs \ GNil) = (\y ys. xs = y #\<^sub>\ ys)" + by (induct xs) auto + +nominal_function toList :: "\ \ (x*b*c) list" where + "toList GNil = []" +| "toList (GCons xbc G) = xbc#(toList G)" + apply (auto, simp add: eqvt_def toList_graph_aux_def ) + using neq_GNil_conv surj_pair by metis +nominal_termination (eqvt) + by lexicographic_order + +nominal_function toSet :: "\ \ (x*b*c) set" where + "toSet GNil = {}" +| "toSet (GCons xbc G) = {xbc} \ (toSet G)" + apply (auto,simp add: eqvt_def toSet_graph_aux_def ) + using neq_GNil_conv surj_pair by metis +nominal_termination (eqvt) + by lexicographic_order + +nominal_function append_g :: "\ \ \ \ \" (infixr "@" 65) where + "append_g GNil g = g" +| "append_g (xbc #\<^sub>\ g1) g2 = (xbc #\<^sub>\ (g1@g2))" + apply (auto,simp add: eqvt_def append_g_graph_aux_def ) + using neq_GNil_conv surj_pair by metis +nominal_termination (eqvt) by lexicographic_order + +nominal_function dom :: "\ \ x set" where + "dom \ = (fst` (toSet \))" + apply auto + unfolding eqvt_def dom_graph_aux_def lfp_eqvt toSet.eqvt by simp +nominal_termination (eqvt) by lexicographic_order + +text \ Use of this is sometimes mixed in with use of freshness and support for the context however it makes it clear +that for immutable variables, the context is `self-supporting'\ + +nominal_function atom_dom :: "\ \ atom set" where + "atom_dom \ = atom`(dom \)" + apply auto + unfolding eqvt_def atom_dom_graph_aux_def lfp_eqvt toSet.eqvt by simp +nominal_termination (eqvt) by lexicographic_order + +subsection \Immutable Variable Context Lemmas\ + +lemma append_GNil[simp]: + "GNil @ G = G" + by simp + +lemma append_g_toSetU [simp]: "toSet (G1@G2) = toSet G1 \ toSet G2" + by(induct G1, auto+) + +lemma supp_GNil: + shows "supp GNil = {}" + by (simp add: supp_def) + +lemma supp_GCons: + fixes xs::\ + shows "supp (x #\<^sub>\ xs) = supp x \ supp xs" + by (simp add: supp_def Collect_imp_eq Collect_neg_eq) + +lemma atom_dom_eq[simp]: + fixes G::\ + shows "atom_dom ((x, b, c) #\<^sub>\ G) = atom_dom ((x, b, c') #\<^sub>\ G)" + using atom_dom.simps toSet.simps by simp + +lemma dom_append[simp]: + "atom_dom (\@\') = atom_dom \ \ atom_dom \'" + using image_Un append_g_toSetU atom_dom.simps dom.simps by metis + +lemma dom_cons[simp]: + "atom_dom ((x,b,c) #\<^sub>\ G) = { atom x } \ atom_dom G" + using image_Un append_g_toSetU atom_dom.simps by auto + +lemma fresh_GNil[ms_fresh]: + shows "a \ GNil" + by (simp add: fresh_def supp_GNil) + +lemma fresh_GCons[ms_fresh]: + fixes xs::\ + shows "a \ (x #\<^sub>\ xs) \ a \ x \ a \ xs" + by (simp add: fresh_def supp_GCons) + +lemma dom_supp_g[simp]: + "atom_dom G \ supp G" + apply(induct G rule: \_induct,simp) + using supp_at_base supp_Pair atom_dom.simps supp_GCons by fastforce + +lemma fresh_append_g[ms_fresh]: + fixes xs::\ + shows "a \ (xs @ ys) \ a \ xs \ a \ ys" + by (induct xs) (simp_all add: fresh_GNil fresh_GCons) + +lemma append_g_assoc: + fixes xs::\ + shows "(xs @ ys) @ zs = xs @ (ys @ zs)" + by (induct xs) simp_all + +lemma append_g_inside: + fixes xs::\ + shows "xs @ (x #\<^sub>\ ys) = (xs @ (x #\<^sub>\ GNil)) @ ys" + by(induct xs,auto+) + +lemma finite_\: + "finite (toSet \)" + by(induct \ rule: \_induct,auto) + +lemma supp_\: + "supp \ = supp (toSet \)" +proof(induct \ rule: \_induct) + case GNil + then show ?case using supp_GNil toSet.simps + by (simp add: supp_set_empty) +next + case (GCons x b c \') + then show ?case using supp_GCons toSet.simps finite_\ supp_of_finite_union + using supp_of_finite_insert by fastforce +qed + +lemma supp_of_subset: + fixes G::"('a::fs set)" + assumes "finite G" and "finite G'" and "G \ G'" + shows "supp G \ supp G'" + using supp_of_finite_sets assms by (metis subset_Un_eq supp_of_finite_union) + +lemma supp_weakening: + assumes "toSet G \ toSet G'" + shows "supp G \ supp G'" + using supp_\ finite_\ by (simp add: supp_of_subset assms) + +lemma fresh_weakening[ms_fresh]: + assumes "toSet G \ toSet G'" and "x \ G'" + shows "x \ G" +proof(rule ccontr) + assume "\ x \ G" + hence "x \ supp G" using fresh_def by auto + hence "x \ supp G'" using supp_weakening assms by auto + thus False using fresh_def assms by auto +qed + +instance \ :: fs + by (standard, induct_tac x, simp_all add: supp_GNil supp_GCons finite_supp) + +lemma fresh_gamma_elem: + fixes \::\ + assumes "a \ \" + and "e \ toSet \" + shows "a \ e" + using assms by(induct \,auto simp add: fresh_GCons) + +lemma fresh_gamma_append: + fixes xs::\ + shows "a \ (xs @ ys) \ a \ xs \ a \ ys" + by (induct xs, simp_all add: fresh_GNil fresh_GCons) + +lemma supp_triple[simp]: + shows "supp (x, y, z) = supp x \ supp y \ supp z" +proof - + have "supp (x,y,z) = supp (x,(y,z))" by auto + hence "supp (x,y,z) = supp x \ (supp y \ supp z)" using supp_Pair by metis + thus ?thesis by auto +qed + +lemma supp_append_g: + fixes xs::\ + shows "supp (xs @ ys) = supp xs \ supp ys" + by(induct xs,auto simp add: supp_GNil supp_GCons ) + +lemma fresh_in_g[simp]: + fixes \::\ and x'::x + shows "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \ = (atom x' \ supp \' \ supp x \ supp b0 \ supp c0 \ supp \)" +proof - + have "atom x' \ \' @ (x, b0, c0) #\<^sub>\ \ = (atom x' \ supp (\' @((x,b0,c0) #\<^sub>\ \)))" + using fresh_def by auto + also have "... = (atom x' \ supp \' \ supp ((x,b0,c0) #\<^sub>\ \))" using supp_append_g by fast + also have "... = (atom x' \ supp \' \ supp x \ supp b0 \ supp c0 \ supp \)" using supp_GCons supp_append_g supp_triple by auto + finally show ?thesis by fast +qed + +lemma fresh_suffix[ms_fresh]: + fixes \::\ + assumes "atom x \ \'@\" + shows "atom x \ \" + using assms by(induct \' rule: \_induct, auto simp add: append_g.simps fresh_GCons) + +lemma not_GCons_self [simp]: + fixes xs::\ + shows "xs \ x #\<^sub>\ xs" + by (induct xs) auto + +lemma not_GCons_self2 [simp]: + fixes xs::\ + shows "x #\<^sub>\ xs \ xs" + by (rule not_GCons_self [symmetric]) + +lemma fresh_restrict: + fixes y::x and \::\ + assumes "atom y \ (\' @ (x, b, c) #\<^sub>\ \)" + shows "atom y \ (\'@\)" + using assms by(induct \' rule: \_induct, auto simp add:fresh_GCons fresh_GNil ) + +lemma fresh_dom_free: + assumes "atom x \ \" + shows "(x,b,c) \ toSet \" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + hence "x\x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by blast + moreover have "atom x \ \'" using fresh_GCons GCons by auto + ultimately show ?case using toSet.simps GCons by auto +qed + +lemma \_set_intros: "x \ toSet ( x #\<^sub>\ xs)" and "y \ toSet xs \ y \ toSet (x #\<^sub>\ xs)" + by simp+ + +lemma fresh_dom_free2: + assumes "atom x \ atom_dom \" + shows "(x,b,c) \ toSet \" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + hence "x\x'" using fresh_def fresh_GCons fresh_Pair supp_at_base by auto + moreover have "atom x \ atom_dom \'" using fresh_GCons GCons by auto + ultimately show ?case using toSet.simps GCons by auto +qed + +subsection \Mutable Variable Context Lemmas\ + +lemma supp_DNil: + shows "supp DNil = {}" + by (simp add: supp_def) + +lemma supp_DCons: + fixes xs::\ + shows "supp (x #\<^sub>\ xs) = supp x \ supp xs" + by (simp add: supp_def Collect_imp_eq Collect_neg_eq) + +lemma fresh_DNil[ms_fresh]: + shows "a \ DNil" + by (simp add: fresh_def supp_DNil) + +lemma fresh_DCons[ms_fresh]: + fixes xs::\ + shows "a \ (x #\<^sub>\ xs) \ a \ x \ a \ xs" + by (simp add: fresh_def supp_DCons) + +instance \ :: fs + by (standard, induct_tac x, simp_all add: supp_DNil supp_DCons finite_supp) + +subsection \Lookup Functions\ + +nominal_function lookup :: "\ \ x \ (b*c) option" where + "lookup GNil x = None" +| "lookup ((x,b,c)#\<^sub>\G) y = (if x=y then Some (b,c) else lookup G y)" + by (auto,simp add: eqvt_def lookup_graph_aux_def, metis neq_GNil_conv surj_pair) +nominal_termination (eqvt) by lexicographic_order + +nominal_function replace_in_g :: "\ \ x \ c \ \" ("_[_\_]" [1000,0,0] 200) where + "replace_in_g GNil _ _ = GNil" +| "replace_in_g ((x,b,c)#\<^sub>\G) x' c' = (if x=x' then ((x,b,c')#\<^sub>\G) else (x,b,c)#\<^sub>\(replace_in_g G x' c'))" + apply(auto,simp add: eqvt_def replace_in_g_graph_aux_def) + using surj_pair \.exhaust by metis +nominal_termination (eqvt) by lexicographic_order + +text \ Functions for looking up data-constructors in the Pi context \ + +nominal_function lookup_fun :: "\ \ f \ fun_def option" where + "lookup_fun [] g = None" +| "lookup_fun ((AF_fundef f ft)#\) g = (if (f=g) then Some (AF_fundef f ft) else lookup_fun \ g)" + apply(auto,simp add: eqvt_def lookup_fun_graph_aux_def ) + by (metis fun_def.exhaust neq_Nil_conv) +nominal_termination (eqvt) by lexicographic_order + +nominal_function lookup_td :: "\ \ string \ type_def option" where + "lookup_td [] g = None" +| "lookup_td ((AF_typedef s lst ) # (\::\)) g = (if (s = g) then Some (AF_typedef s lst ) else lookup_td \ g)" +| "lookup_td ((AF_typedef_poly s bv lst ) # (\::\)) g = (if (s = g) then Some (AF_typedef_poly s bv lst ) else lookup_td \ g)" + apply(auto,simp add: eqvt_def lookup_td_graph_aux_def ) + by (metis type_def.exhaust neq_Nil_conv) +nominal_termination (eqvt) by lexicographic_order + +nominal_function name_of_type ::"type_def \ f " where + "name_of_type (AF_typedef f _ ) = f" +| "name_of_type (AF_typedef_poly f _ _) = f" + apply(auto,simp add: eqvt_def name_of_type_graph_aux_def ) + using type_def.exhaust by blast +nominal_termination (eqvt) by lexicographic_order + +nominal_function name_of_fun ::"fun_def \ f " where + "name_of_fun (AF_fundef f ft) = f" + apply(auto,simp add: eqvt_def name_of_fun_graph_aux_def ) + using fun_def.exhaust by blast +nominal_termination (eqvt) by lexicographic_order + +nominal_function remove2 :: "'a::pt \ 'a list \ 'a list" where + "remove2 x [] = []" | + "remove2 x (y # xs) = (if x = y then xs else y # remove2 x xs)" + by (simp add: eqvt_def remove2_graph_aux_def,auto+,meson list.exhaust) +nominal_termination (eqvt) by lexicographic_order + +nominal_function base_for_lit :: "l \ b" where + "base_for_lit (L_true) = B_bool " +| "base_for_lit (L_false) = B_bool " +| "base_for_lit (L_num n) = B_int " +| "base_for_lit (L_unit) = B_unit " +| "base_for_lit (L_bitvec v) = B_bitvec" + apply (auto simp: eqvt_def base_for_lit_graph_aux_def ) + using l.strong_exhaust by blast +nominal_termination (eqvt) by lexicographic_order + +lemma neq_DNil_conv: "(xs \ DNil) = (\y ys. xs = y #\<^sub>\ ys)" + by (induct xs) auto + +nominal_function setD :: "\ \ (u*\) set" where + "setD DNil = {}" +| "setD (DCons xbc G) = {xbc} \ (setD G)" + apply (auto,simp add: eqvt_def setD_graph_aux_def ) + using neq_DNil_conv surj_pair by metis +nominal_termination (eqvt) by lexicographic_order + +lemma eqvt_triple: + fixes y::"'a::at" and ya::"'a::at" and xa::"'c::at" and va::"'d::fs" and s::s and sa::s and f::"s*'c*'d \ s" + assumes "atom y \ (xa, va)" and "atom ya \ (xa, va)" and + "\c. atom c \ (s, sa) \ atom c \ (y, ya, s, sa) \ (y \ c) \ s = (ya \ c) \ sa" + and "eqvt_at f (s,xa,va)" and "eqvt_at f (sa,xa,va)" and + "atom c \ (s, va, xa, sa)" and "atom c \ (y, ya, f (s, xa, va), f (sa, xa, va))" + shows "(y \ c) \ f (s, xa, va) = (ya \ c) \ f (sa, xa, va)" +proof - + have " (y \ c) \ f (s, xa, va) = f ( (y \ c) \ (s,xa,va))" using assms eqvt_at_def by metis + also have "... = f ( (y \ c) \ s, (y \ c) \ xa ,(y \ c) \ va)" by auto + also have "... = f ( (ya \ c) \ sa, (ya \ c) \ xa ,(ya \ c) \ va)" proof - + have " (y \ c) \ s = (ya \ c) \ sa" using assms Abs1_eq_iff_all by auto + moreover have "((y \ c) \ xa) = ((ya \ c) \ xa)" using assms flip_fresh_fresh fresh_prodN by metis + moreover have "((y \ c) \ va) = ((ya \ c) \ va)" using assms flip_fresh_fresh fresh_prodN by metis + ultimately show ?thesis by auto + qed + also have "... = f ( (ya \ c) \ (sa,xa,va))" by auto + finally show ?thesis using assms eqvt_at_def by metis +qed + +section \Functions for bit list/vectors\ + +inductive split :: "int \ bit list \ bit list * bit list \ bool" where + "split 0 xs ([], xs)" +| "split m xs (ys,zs) \ split (m+1) (x#xs) ((x # ys), zs)" +equivariance split +nominal_inductive split . + +lemma split_concat: + assumes "split n v (v1,v2)" + shows "v = append v1 v2" + using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts) + case 1 + then show ?case by auto +next + case (2 m xs ys zs x) + then show ?case by auto +qed + +lemma split_n: + assumes "split n v (v1,v2)" + shows "0 \ n \ n \ int (length v)" + using assms proof(induct rule: split.inducts) + case (1 xs) + then show ?case by auto +next + case (2 m xs ys zs x) + then show ?case by auto +qed + +lemma split_length: + assumes "split n v (v1,v2)" + shows "n = int (length v1)" + using assms proof(induct "(v1,v2)" arbitrary: v1 v2 rule: split.inducts) + case (1 xs) + then show ?case by auto +next + case (2 m xs ys zs x) + then show ?case by auto +qed + +lemma obtain_split: + assumes "0 \ n" and "n \ int (length bv)" + shows "\ bv1 bv2. split n bv (bv1 , bv2)" + using assms proof(induct bv arbitrary: n) + case Nil + then show ?case using split.intros by auto +next + case (Cons b bv) + show ?case proof(cases "n = 0") + case True + then show ?thesis using split.intros by auto + next + case False + then obtain m where m:"n=m+1" using Cons + by (metis add.commute add_minus_cancel) + moreover have "0 \ m" using False m Cons by linarith + then obtain bv1 and bv2 where "split m bv (bv1 , bv2)" using Cons m by force + hence "split n (b # bv) ((b#bv1), bv2)" using m split.intros by auto + then show ?thesis by auto + qed +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/SyntaxL.thy b/thys/MiniSail/SyntaxL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/SyntaxL.thy @@ -0,0 +1,383 @@ +(*<*) +theory SyntaxL + imports Syntax IVSubst +begin + (*>*) + +chapter \Syntax Lemmas\ + +section \Support, lookup and contexts\ + +lemma supp_v_tau [simp]: + assumes "atom z \ v" + shows "supp (\ z : b | CE_val (V_var z) == CE_val v \) = supp v \ supp b" + using assms \.supp c.supp ce.supp + by (simp add: fresh_def supp_at_base) + +lemma supp_v_var_tau [simp]: + assumes "z \ x" + shows "supp (\ z : b | CE_val (V_var z) == CE_val (V_var x) \) = { atom x } \ supp b" + using supp_v_tau assms + using supp_at_base by fastforce + +text \ Sometimes we need to work with a version of a binder where the variable is fresh +in something else, such as a bigger context. I think these could be generated automatically \ + +lemma obtain_fresh_fun_def: + fixes t::"'b::fs" + shows "\y::x. atom y \ (s,c,\,t) \ (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \ x) \ c ) ((y \ x) \ \) ((y \ x) \ s))))" +proof - + obtain y::x where y: "atom y \ (s,c,\,t)" using obtain_fresh by blast + moreover have "AF_fundef f (AF_fun_typ_none (AF_fun_typ y b ((y \ x) \ c ) ((y \ x) \ \) ((y \ x) \ s))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)))" + proof(cases "x=y") + case True + then show ?thesis using fun_def.eq_iff Abs1_eq_iff(3) flip_commute flip_fresh_fresh fresh_PairD by auto + next + case False + + have "(AF_fun_typ y b ((y \ x) \ c) ((y \ x) \ \) ((y \ x) \ s)) =(AF_fun_typ x b c \ s)" proof(subst fun_typ.eq_iff, subst Abs1_eq_iff(3)) + show \(y = x \ (((y \ x) \ c, (y \ x) \ \), (y \ x) \ s) = ((c, \), s) \ + y \ x \ (((y \ x) \ c, (y \ x) \ \), (y \ x) \ s) = (y \ x) \ ((c, \), s) \ atom y \ ((c, \), s)) \ + b = b\ using False flip_commute flip_fresh_fresh fresh_PairD y by auto + qed + thus ?thesis by metis + qed + ultimately show ?thesis using y fresh_Pair by metis +qed + +lemma lookup_fun_member: + assumes "Some (AF_fundef f ft) = lookup_fun \ f" + shows "AF_fundef f ft \ set \" + using assms proof (induct \) + case Nil + then show ?case by auto +next + case (Cons a \) + then show ?case using lookup_fun.simps + by (metis fun_def.exhaust insert_iff list.simps(15) option.inject) +qed + +lemma rig_dom_eq: + "dom (G[x \ c]) = dom G" +proof(induct G rule: \.induct) + case GNil + then show ?case using replace_in_g.simps by presburger +next + case (GCons xbc \') + obtain x' and b' and c' where xbc: "xbc=(x',b',c')" using prod_cases3 by blast + then show ?case using replace_in_g.simps GCons by simp +qed + +lemma lookup_in_rig_eq: + assumes "Some (b,c) = lookup \ x" + shows "Some (b,c') = lookup (\[x\c']) x" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x b c \') + then show ?case using replace_in_g.simps lookup.simps by auto +qed + +lemma lookup_in_rig_neq: + assumes "Some (b,c) = lookup \ y" and "x\y" + shows "Some (b,c) = lookup (\[x\c']) y" + using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + then show ?case using replace_in_g.simps lookup.simps by auto +qed + +lemma lookup_in_rig: + assumes "Some (b,c) = lookup \ y" + shows "\c''. Some (b,c'') = lookup (\[x\c']) y" +proof(cases "x=y") + case True + then show ?thesis using lookup_in_rig_eq using assms by blast +next + case False + then show ?thesis using lookup_in_rig_neq using assms by blast +qed + +lemma lookup_inside[simp]: + assumes "x \ fst ` toSet \'" + shows "Some (b1,c1) = lookup (\'@(x,b1,c1)#\<^sub>\\) x" + using assms by(induct \',auto) + +lemma lookup_inside2: + assumes "Some (b1,c1) = lookup (\'@((x,b0,c0)#\<^sub>\\)) y" and "x\y" + shows "Some (b1,c1) = lookup (\'@((x,b0,c0')#\<^sub>\\)) y" + using assms by(induct \' rule: \.induct,auto+) + +fun tail:: "'a list \ 'a list" where + "tail [] = []" +| "tail (x#xs) = xs" + +lemma lookup_options: + assumes "Some (b,c) = lookup (xt#\<^sub>\G) x" + shows "((x,b,c) = xt) \ (Some (b,c) = lookup G x)" + by (metis assms lookup.simps(2) option.inject surj_pair) + +lemma lookup_x: + assumes "Some (b,c) = lookup G x" + shows "x \ fst ` toSet G" + using assms + by(induct "G" rule: \.induct ,auto+) + +lemma GCons_eq_appendI: + fixes xs1::\ + shows "[| x #\<^sub>\ xs1 = ys; xs = xs1 @ zs |] ==> x #\<^sub>\ xs = ys @ zs" + by (drule sym) simp + +lemma split_G: "x : toSet xs \ \ys zs. xs = ys @ x #\<^sub>\ zs" +proof (induct xs) + case GNil thus ?case by simp +next + case GCons thus ?case using GCons_eq_appendI + by (metis Un_iff append_g.simps(1) singletonD toSet.simps(2)) +qed + +lemma lookup_not_empty: + assumes "Some \ = lookup G x" + shows "G \ GNil" + using assms by auto + +lemma lookup_in_g: + assumes "Some (b,c) = lookup \ x" + shows "(x,b,c) \ toSet \" + using assms apply(induct \, simp) + using lookup_options by fastforce + +lemma lookup_split: + fixes \::\ + assumes "Some (b,c) = lookup \ x" + shows "\G G' . \ = G'@(x,b,c)#\<^sub>\G" + by (meson assms(1) lookup_in_g split_G) + +lemma toSet_splitU[simp]: + "(x',b',c') \ toSet (\' @ (x, b, c) #\<^sub>\ \) \ (x',b',c') \ (toSet \' \ {(x, b, c)} \ toSet \)" + using append_g_toSetU toSet.simps by auto + +lemma toSet_splitP[simp]: + "(\(x', b', c')\toSet (\' @ (x, b, c) #\<^sub>\ \). P x' b' c') \ (\ (x', b', c')\toSet \'. P x' b' c') \ P x b c \ (\ (x', b', c')\toSet \. P x' b' c')" (is "?A \ ?B") + using toSet_splitU by force + +lemma lookup_restrict: + assumes "Some (b',c') = lookup (\'@(x,b,c)#\<^sub>\\) y" and "x \ y" + shows "Some (b',c') = lookup (\'@\) y" + using assms proof(induct \' rule:\_induct) + case GNil + then show ?case by auto +next + case (GCons x1 b1 c1 \') + then show ?case by auto +qed + +lemma supp_list_member: + fixes x::"'a::fs" and l::"'a list" + assumes "x \ set l" + shows "supp x \ supp l" + using assms apply(induct l, auto) + using supp_Cons by auto + +lemma GNil_append: + assumes "GNil = G1@G2" + shows "G1 = GNil \ G2 = GNil" +proof(rule ccontr) + assume "\ (G1 = GNil \ G2 = GNil)" + hence "G1@G2 \ GNil" using append_g.simps by (metis \.distinct(1) \.exhaust) + thus False using assms by auto +qed + +lemma GCons_eq_append_conv: + fixes xs::\ + shows "x#\<^sub>\xs = ys@zs = (ys = GNil \ x#\<^sub>\xs = zs \ (\ys'. x#\<^sub>\ys' = ys \ xs = ys'@zs))" + by(cases ys) auto + +lemma dclist_distinct_unique: + assumes "(dc, const) \ set dclist2" and "(cons, const1) \ set dclist2" and "dc=cons" and "distinct (List.map fst dclist2)" + shows "(const) = const1" +proof - + have "(cons, const) = (dc, const1)" + using assms by (metis (no_types, lifting) assms(3) assms(4) distinct.simps(1) distinct.simps(2) empty_iff insert_iff list.set(1) list.simps(15) list.simps(8) list.simps(9) map_of_eq_Some_iff) + thus ?thesis by auto +qed + +lemma fresh_d_fst_d: + assumes "atom u \ \" + shows "u \ fst ` set \" + using assms proof(induct \) + case Nil + then show ?case by auto +next + case (Cons ut \') + obtain u' and t' where *:"ut = (u',t') " by fastforce + hence "atom u \ ut \ atom u \ \'" using fresh_Cons Cons by auto + moreover hence "atom u \ fst ut" using * fresh_Pair[of "atom u" u' t'] Cons by auto + ultimately show ?case using Cons by auto +qed + +lemma bv_not_in_bset_supp: + fixes bv::bv + assumes "bv |\| B" + shows "atom bv \ supp B" +proof - + have *:"supp B = fset (fimage atom B)" + by (metis fimage.rep_eq finite_fset supp_finite_set_at_base supp_fset) + thus ?thesis using assms + using notin_fset by fastforce +qed + +lemma u_fresh_d: + assumes "atom u \ D" + shows "u \ fst ` setD D" + using assms proof(induct D rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' \') + then show ?case unfolding setD.simps + using fresh_DCons fresh_Pair by (simp add: fresh_Pair fresh_at_base(2)) +qed + +section \Type Definitions\ + +lemma exist_fresh_bv: + fixes tm::"'a::fs" + shows "\bva2 dclist2. AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ tm" +proof - + obtain bva2::bv where *:"atom bva2 \ (bva, dclist,tyid,tm)" using obtain_fresh by metis + moreover hence "bva2 \ bva" using fresh_at_base by auto + moreover have " dclist = (bva \ bva2) \ (bva2 \ bva) \ dclist" by simp + moreover have "atom bva \ (bva2 \ bva) \ dclist" proof - + have "atom bva2 \ dclist" using * fresh_prodN by auto + hence "atom ((bva2 \ bva) \ bva2) \ (bva2 \ bva) \ dclist" using fresh_eqvt True_eqvt + proof - + have "(bva2 \ bva) \ atom bva2 \ (bva2 \ bva) \ dclist" + by (metis True_eqvt \atom bva2 \ dclist\ fresh_eqvt) (* 62 ms *) + then show ?thesis + by simp (* 125 ms *) + qed + thus ?thesis by auto + qed + ultimately have "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 ((bva2 \ bva ) \ dclist)" + unfolding type_def.eq_iff Abs1_eq_iff by metis + thus ?thesis using * fresh_prodN by metis +qed + +lemma obtain_fresh_bv: + fixes tm::"'a::fs" + obtains bva2::bv and dclist2 where "AF_typedef_poly tyid bva dclist = AF_typedef_poly tyid bva2 dclist2 \ + atom bva2 \ tm" + using exist_fresh_bv by metis + +section \Function Definitions\ + +lemma fun_typ_flip: + fixes bv1::bv and c::bv + shows "(bv1 \ c) \ AF_fun_typ x1 b1 c1 \1 s1 = AF_fun_typ x1 ((bv1 \ c) \ b1) ((bv1 \ c) \ c1) ((bv1 \ c) \ \1) ((bv1 \ c) \ s1)" + using fun_typ.perm_simps flip_fresh_fresh supp_at_base fresh_def + flip_fresh_fresh fresh_def supp_at_base + by (simp add: flip_fresh_fresh) + +lemma fun_def_eq: + assumes "AF_fundef fa (AF_fun_typ_none (AF_fun_typ xa ba ca \a sa)) = AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))" + shows "f = fa" and "b = ba" and "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \a = [[atom x]]lst. \" and + " [[atom xa]]lst. ca = [[atom x]]lst. c" + using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis + using fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff lst_snd lst_fst using assms apply metis +proof - + have "([[atom xa]]lst. ((ca, \a), sa) = [[atom x]]lst. ((c, \), s))" using assms fun_def.eq_iff fun_typ_q.eq_iff fun_typ.eq_iff by auto + thus "[[atom xa]]lst. sa = [[atom x]]lst. s" and "[[atom xa]]lst. \a = [[atom x]]lst. \" and + " [[atom xa]]lst. ca = [[atom x]]lst. c" using lst_snd lst_fst by metis+ +qed + +lemma fun_arg_unique_aux: + assumes "AF_fun_typ x1 b1 c1 \1' s1' = AF_fun_typ x2 b2 c2 \2' s2'" + shows "\ x1 : b1 | c1 \ = \ x2 : b2 | c2\" +proof - + have " ([[atom x1]]lst. c1 = [[atom x2]]lst. c2)" using fun_def_eq assms by metis + moreover have "b1 = b2" using fun_typ.eq_iff assms by metis + ultimately show ?thesis using \.eq_iff by fast +qed + +lemma fresh_x_neq: + fixes x::x and y::x + shows "atom x \ y = (x \ y)" + using fresh_at_base fresh_def by auto + +lemma obtain_fresh_z3: + fixes tm::"'b::fs" + obtains z::x where "\ x : b | c \ = \ z : b | c[x::=V_var z]\<^sub>c\<^sub>v \ \ atom z \ tm \ atom z \ (x,c)" +proof - + obtain z::x and c'::c where z:"\ x : b | c \ = \ z : b | c' \ \ atom z \ (tm,x,c)" using obtain_fresh_z2 b_of.simps by metis + hence "c' = c[x::=V_var z]\<^sub>c\<^sub>v" proof - + have "([[atom z]]lst. c' = [[atom x]]lst. c)" using z \.eq_iff by metis + hence "c' = (z \ x) \ c" using Abs1_eq_iff[of z c' x c] fresh_x_neq fresh_prodN by fastforce + also have "... = c[x::=V_var z]\<^sub>c\<^sub>v" + using subst_v_c_def flip_subst_v[of z c x] z fresh_prod3 by metis + finally show ?thesis by auto + qed + thus ?thesis using z fresh_prodN that by metis +qed + +lemma u_fresh_v: + fixes u::u and t::v + shows "atom u \ t" + by(nominal_induct t rule:v.strong_induct,auto) + +lemma u_fresh_ce: + fixes u::u and t::ce + shows "atom u \ t" + apply(nominal_induct t rule:ce.strong_induct) + using u_fresh_v pure_fresh + apply (auto simp add: opp.fresh ce.fresh opp.fresh opp.exhaust) + unfolding ce.fresh opp.fresh opp.exhaust by (simp add: fresh_opp_all) + +lemma u_fresh_c: + fixes u::u and t::c + shows "atom u \ t" + by(nominal_induct t rule:c.strong_induct,auto simp add: c.fresh u_fresh_ce) + +lemma u_fresh_g: + fixes u::u and t::\ + shows "atom u \ t" + by(induct t rule:\_induct, auto simp add: u_fresh_b u_fresh_c fresh_GCons fresh_GNil) + +lemma u_fresh_t: + fixes u::u and t::\ + shows "atom u \ t" + by(nominal_induct t rule:\.strong_induct,auto simp add: \.fresh u_fresh_c u_fresh_b) + +lemma b_of_c_of_eq: + assumes "atom z \ \" + shows "\ z : b_of \ | c_of \ z \ = \" + using assms proof(nominal_induct \ avoiding: z rule: \.strong_induct) + case (T_refined_type x1a x2a x3a) + + hence " \ z : b_of \ x1a : x2a | x3a \ | c_of \ x1a : x2a | x3a \ z \ = \ z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \" + using b_of.simps c_of.simps c_of_eq by auto + moreover have "\ z : x2a | x3a[x1a::=V_var z]\<^sub>c\<^sub>v \ = \ x1a : x2a | x3a \" using T_refined_type \.fresh by auto + ultimately show ?case by auto +qed + +lemma fresh_d_not_in: + assumes "atom u2 \ \'" + shows "u2 \ fst ` setD \'" + using assms proof(induct \' rule: \_induct) + case DNil + then show ?case by simp +next + case (DCons u t \') + hence *: "atom u2 \ \' \ atom u2 \ (u,t)" + by (simp add: fresh_def supp_DCons) + hence "u2 \ fst ` setD \'" using DCons by auto + moreover have "u2 \ u" using * fresh_Pair + by (metis eq_fst_iff not_self_fresh) + ultimately show ?case by simp +qed + +end \ No newline at end of file diff --git a/thys/MiniSail/Typing.thy b/thys/MiniSail/Typing.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Typing.thy @@ -0,0 +1,586 @@ +(*<*) +theory Typing + imports RCLogic WellformedL +begin + (*>*) + +chapter \Type System\ + +text \The MiniSail type system. We define subtyping judgement first and then typing judgement +for the term forms\ + +section \Subtyping\ + +text \ Subtyping is defined on top of refinement constraint logic (RCL). +A subtyping check is converted into an RCL validity check. \ + +inductive subtype :: "\ \ \ \ \ \ \ \ \ \ bool" ("_ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + subtype_baseI: "\ + atom x \ (\, \, \, z,c,z',c') ; + \; \; \ \\<^sub>w\<^sub>f \ z : b | c \; + \; \; \ \\<^sub>w\<^sub>f \ z' : b | c' \; + \; \ ; (x,b, c[z::=[x]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c'[z'::=[x]\<^sup>v]\<^sub>v +\ \ + \; \; \ \ \ z : b | c \ \ \ z' : b | c' \" + +equivariance subtype +nominal_inductive subtype + avoids subtype_baseI: x +proof(goal_cases) + case (1 \ \ \ z b c z' c' x) + then show ?case using fresh_star_def 1 by force +next + case (2 \ \ \ z b c z' c' x) + then show ?case by auto +qed + +inductive_cases subtype_elims: + "\; \; \ \ \ z : b | c \ \ \ z' : b | c' \" + "\; \; \ \ \\<^sub>1 \ \\<^sub>2" + +section \Literals\ + +text \The type synthesised has the constraint that z equates to the literal\ + +inductive infer_l :: "l \ \ \ bool" (" \ _ \ _" [50, 50] 50) where + infer_trueI: " \ L_true \ \ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_falseI: " \ L_false \ \ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_natI: " \ L_num n \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_num n]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_unitI: " \ L_unit \ \ z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_unit]\<^sup>v]\<^sup>c\<^sup>e \" +| infer_bitvecI: " \ L_bitvec bv \ \ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[L_bitvec bv]\<^sup>v]\<^sup>c\<^sup>e \" + +nominal_inductive infer_l . +equivariance infer_l + +inductive_cases infer_l_elims[elim!]: + "\ L_true \ \" + "\ L_false \ \" + "\ L_num n \ \" + "\ L_unit \ \" + "\ L_bitvec x \ \" + "\ l \ \" + +lemma infer_l_form2[simp]: + shows "\z. \ l \ (\ z : base_for_lit l | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[l]\<^sup>v]\<^sup>c\<^sup>e \)" +proof (nominal_induct l rule: l.strong_induct) + case (L_num x) + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_true + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_false + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case L_unit + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +next + case (L_bitvec x) + then show ?case using infer_l.intros base_for_lit.simps has_fresh_z by metis +qed + +section \Values\ + +inductive infer_v :: "\ \ \ \ \ \ v \ \ \ bool" (" _ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + +infer_v_varI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + Some (b,c) = lookup \ x; + atom z \ x ; atom z \ (\, \, \) +\ \ + \; \; \ \ [x]\<^sup>v \ \ z : b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[x]\<^sup>v]\<^sup>c\<^sup>e \" + +| infer_v_litI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + \ l \ \ +\ \ + \; \; \ \ [l]\<^sup>v \ \" + +| infer_v_pairI: "\ + atom z \ (v1, v2); atom z \ (\, \, \) ; + \; \; \ \ (v1::v) \ t1 ; + \; \ ; \ \ (v2::v) \ t2 +\ \ + \; \; \ \ V_pair v1 v2 \ (\ z : B_pair (b_of t1) (b_of t2) | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \) " + +| infer_v_consI: "\ + AF_typedef s dclist \ set \; + (dc, tc) \ set dclist ; + \; \; \ \ v \ tv ; + \; \; \ \ tv \ tc ; + atom z \ v ; atom z \ (\, \, \) +\ \ + \; \; \ \ V_cons s dc v \ (\ z : B_id s | [[z]\<^sup>v]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \)" + +| infer_v_conspI: "\ + AF_typedef_poly s bv dclist \ set \; + (dc, tc) \ set dclist ; + \; \; \ \ v \ tv; + \; \; \ \ tv \ tc[bv::=b]\<^sub>\\<^sub>b ; + atom z \ (\, \, \, v, b); + atom bv \ (\, \, \, v, b); + \; \ \\<^sub>w\<^sub>f b +\ \ + \; \; \ \ V_consp s dc b v \ (\ z : B_app s b | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_val (V_consp s dc b v)) \)" + +equivariance infer_v +nominal_inductive infer_v + avoids infer_v_conspI: bv and z | infer_v_varI: z | infer_v_pairI: z | infer_v_consI: z +proof(goal_cases) + case (1 \ \ \ b c x z) + hence "atom z \ \ z : b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \" using \.fresh by simp + then show ?case unfolding fresh_star_def using 1 by simp +next + case (2 \ \ \ b c x z) + then show ?case by auto +next + case (3 z v1 v2 \ \ \ t1 t2) + hence "atom z \ \ z : [ b_of t1 , b_of t2 ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ v1 , v2 ]\<^sup>v ]\<^sup>c\<^sup>e \" using \.fresh by simp + then show ?case unfolding fresh_star_def using 3 by simp +next + case (4 z v1 v2 \ \ \ t1 t2) + then show ?case by auto +next + case (5 s dclist \ dc tc \ \ v tv z) + hence "atom z \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_cons s dc v ]\<^sup>c\<^sup>e \" using \.fresh b.fresh pure_fresh by auto + moreover have "atom z \ V_cons s dc v" using v.fresh 5 using v.fresh fresh_prodN pure_fresh by metis + then show ?case unfolding fresh_star_def using 5 by simp +next + case (6 s dclist \ dc tc \ \ v tv z) + then show ?case by auto +next + case (7 s bv dclist \ dc tc \ \ v tv b z) + hence "atom bv \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis + moreover then have "atom bv \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + using \.fresh ce.fresh v.fresh by auto + moreover have "atom z \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh 7 by metis + moreover then have "atom z \ \ z : B_id s | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + using \.fresh ce.fresh v.fresh by auto + ultimately show ?case using fresh_star_def 7 by force +next + case (8 s bv dclist \ dc tc \ \ v tv b z) + then show ?case by auto +qed + +inductive_cases infer_v_elims[elim!]: + "\; \; \ \ V_var x \ \" + "\; \; \ \ V_lit l \ \" + "\; \; \ \ V_pair v1 v2 \ \" + "\; \; \ \ V_cons s dc v \ \" + "\; \; \ \ V_pair v1 v2 \ (\ z : b | c \) " + "\; \; \ \ V_pair v1 v2 \ (\ z : [ b1 , b2 ]\<^sup>b | [[z]\<^sup>v]\<^sup>c\<^sup>e == [[v1,v2]\<^sup>v]\<^sup>c\<^sup>e \) " + "\; \; \ \ V_consp s dc b v \ \ " + +inductive check_v :: "\ \ \ \ \ \ v \ \ \ bool" ("_ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + check_v_subtypeI: "\ \; \; \ \ \1 \ \2; \; \; \ \ v \ \1 \ \ \; \ ; \ \ v \ \2" +equivariance check_v +nominal_inductive check_v . + +inductive_cases check_v_elims[elim!]: + "\; \ ; \ \ v \ \" + +section \Expressions\ + +text \ Type synthesis for expressions \ + +inductive infer_e :: "\ \ \ \ \ \ \ \ \ \ e \ \ \ bool" ("_ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50] 50) where + +infer_e_valI: "\ + (\; \; \ \\<^sub>w\<^sub>f \) ; + (\ \\<^sub>w\<^sub>f (\::\)) ; + (\; \; \ \ v \ \) \ \ + \; \; \; \; \ \ (AE_val v) \ \" + +| infer_e_plusI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_int | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_int | c2 \; + atom z3 \ (AE_op Plus v1 v2); atom z3 \ \ \ \ + \; \; \; \; \ \ AE_op Plus v1 v2 \ \ z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_leqI: "\ + \; \; \ \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_int | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_int | c2 \; + atom z3 \ (AE_op LEq v1 v2); atom z3 \ \ +\ \ + \; \; \; \; \ \ AE_op LEq v1 v2 \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_eqI: "\ + \; \; \ \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : b | c1 \ ; + \; \; \ \ v2 \ \ z2 : b | c2 \; + atom z3 \ (AE_op Eq v1 v2); atom z3 \ \ ; + b \ { B_bool, B_int, B_unit } +\ \ + \; \; \; \; \ \ AE_op Eq v1 v2 \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_appI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f; + \; \; \ \ v \ \ x : b | c \; + atom x \ (\, \, \, \, \,v , \); + \'[x::=v]\<^sub>v = \ +\ \ + \; \; \; \; \ \ AE_app f v \ \" + +| infer_e_appPI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \ \\<^sub>w\<^sub>f b' ; + Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \' s'))) = lookup_fun \ f; + \; \; \ \ v \ \ x : b[bv::=b']\<^sub>b | c[bv::=b']\<^sub>b \; atom x \ \; + (\'[bv::=b']\<^sub>b[x::=v]\<^sub>v) = \ ; + atom bv \ (\, \, \, \, \, b', v, \) +\ \ + \; \; \; \; \ \ AE_appP f b' v \ \" + +| infer_e_fstI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : [b1,b2]\<^sup>b | c \; + atom z \ AE_fst v ; atom z \ \ \ \ + \; \; \; \; \ \ AE_fst v \ \ z : b1 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_fst [v]\<^sup>c\<^sup>e)) \" + +| infer_e_sndI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : [ b1, b2]\<^sup>b | c \; + atom z \ AE_snd v ; atom z \ \ \ \ + \; \; \; \; \ \ AE_snd v \ \ z : b2 | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_snd [v]\<^sup>c\<^sup>e)) \" + +| infer_e_lenI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v \ \ z' : B_bitvec | c \; + atom z \ AE_len v ; atom z \ \\ \ + \; \; \; \; \ \ AE_len v \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e)) \" + +| infer_e_mvarI: "\ + \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \\<^sub>w\<^sub>f \; + (u,\) \ setD \ \ \ + \; \; \; \; \ \ AE_mvar u \ \" + +| infer_e_concatI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\) ; + \; \; \ \ v1 \ \ z1 : B_bitvec | c1 \ ; + \; \; \ \ v2 \ \ z2 : B_bitvec | c2 \; + atom z3 \ (AE_concat v1 v2); atom z3 \ \ \ \ + \; \; \; \; \ \ AE_concat v1 v2 \ \ z3 : B_bitvec | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + +| infer_e_splitI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f (\::\); + \ ; \ ; \ \ v1 \ \ z1 : B_bitvec | c1 \ ; + \ ; \ ; \ \ v2 \ \ z2 : B_int | (CE_op LEq (CE_val (V_lit (L_num 0))) (CE_val (V_var z2))) == (CE_val (V_lit L_true)) AND + (CE_op LEq (CE_val (V_var z2)) (CE_len (CE_val (v1)))) == (CE_val (V_lit L_true)) \; + atom z1 \ (AE_split v1 v2); atom z1 \ \; + atom z2 \ (AE_split v1 v2); atom z2 \ \; + atom z3 \ (AE_split v1 v2); atom z3 \ \ +\ \ + \; \; \; \; \ \ (AE_split v1 v2) \ \ z3 : B_pair B_bitvec B_bitvec | + ((CE_val v1) == (CE_concat (CE_fst (CE_val (V_var z3))) (CE_snd (CE_val (V_var z3))))) + AND (((CE_len (CE_fst (CE_val (V_var z3))))) == (CE_val ( v2))) \" + +equivariance infer_e +nominal_inductive infer_e + avoids infer_e_appI: x |infer_e_appPI: bv | infer_e_splitI: z3 and z1 and z2 +proof(goal_cases) + case (1 \ \ \ \ \ f x b c \' s' v \) + moreover hence "atom x \ [ f v ]\<^sup>e" using fresh_prodN pure_fresh e.fresh by force + ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh by simp +next + case (2 \ \ \ \ \ f x b c \' s' v \) + then show ?case by auto +next + case (3 \ \ \ \ \ b' f bv x b c \' s' v \) + moreover hence "atom bv \ AE_appP f b' v" using fresh_prodN pure_fresh e.fresh by force + ultimately show ?case unfolding fresh_star_def using fresh_prodN e.fresh pure_fresh fresh_Pair by auto +next + case (4 \ \ \ \ \ b' f bv x b c \' s' v \) + then show ?case by auto +next + case (5 \ \ \ \ \ v1 z1 c1 v2 z2 z3) + have "atom z3 \ \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \" + using \.fresh by simp + then show ?case unfolding fresh_star_def fresh_prod7 using wfG_fresh_x2 5 by auto +next + case (6 \ \ \ \ \ v1 z1 c1 v2 z2 z3) + then show ?case by auto +qed + +inductive_cases infer_e_elims[elim!]: + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : B_int | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : B_bool | [[z3]\<^sup>v]\<^sup>c\<^sup>e == (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : B_int | c \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_app f v ) \ \" + "\; \; \; \; \ \ (AE_val v) \ \" + "\; \; \; \; \ \ (AE_fst v) \ \" + "\; \; \; \; \ \ (AE_snd v) \ \" + "\; \; \; \; \ \ (AE_mvar u) \ \" + "\; \; \; \; \ \ (AE_op Plus v1 v2) \ \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \" + "\; \; \; \; \ \ (AE_op LEq v1 v2) \ \ z3 : B_bool | c \" + "\; \; \; \; \ \ (AE_app f v ) \ \[x::=v]\<^sub>\\<^sub>v" + "\; \; \; \; \ \ (AE_op opp v1 v2) \ \" + "\; \; \; \; \ \ (AE_len v) \ \" + "\; \; \; \; \ \ (AE_len v) \ \ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == ((CE_len [v]\<^sup>c\<^sup>e))\ " + "\; \; \; \; \ \ AE_concat v1 v2 \ \" + "\; \; \; \; \ \ AE_concat v1 v2 \ (\ z : b | c \) " + "\; \; \; \; \ \ AE_concat v1 v2 \ (\ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == (CE_concat [v1]\<^sup>c\<^sup>e [v1]\<^sup>c\<^sup>e) \) " + "\; \; \; \; \ \ (AE_appP f b v ) \ \" + "\; \; \; \; \ \ AE_split v1 v2 \ \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \ z3 : b | c \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \ z3 : B_bool | c \" + "\; \; \; \; \ \ (AE_op Eq v1 v2) \ \" +nominal_termination (eqvt) by lexicographic_order + +section \Statements\ + +inductive check_s :: "\ \ \ \ \ \ \ \ \ \ s \ \ \ bool" (" _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) and + check_branch_s :: "\ \ \ \ \ \ \ \ \ \ tyid \ string \ \ \ v \ branch_s \ \ \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) and + check_branch_list :: "\ \ \ \ \ \ \ \ \ \ tyid \ (string * \) list \ v \ branch_list \ \ \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50,50,50] 50) where + check_valI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \ v \ \'; + \; \; \ \ \' \ \ \ \ + \; \; \; \; \ \ (AS_val v) \ \" + +| check_letI: "\ + atom x \ (\, \, \, \, \, e, \); + atom z \ (x, \, \, \, \, \, e, \, s); + \; \; \; \; \ \ e \ \ z : b | c \ ; + \; \ ; \ ; ((x,b,c[z::=V_var x]\<^sub>v)#\<^sub>\\) ; \ \ s \ \ +\ \ + \; \; \; \; \ \ (AS_let x e s) \ \" + +| check_assertI: "\ + atom x \ (\, \, \, \, \, c, \, s); + \; \ ; \ ; ((x,B_bool,c)#\<^sub>\\) ; \ \ s \ \ ; + \; \; \ \ c; + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \ (AS_assert c s) \ \" + +|check_branch_s_branchI: "\ + \; \; \ \\<^sub>w\<^sub>f \ ; + \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + \ ; {||} ; GNil \\<^sub>w\<^sub>f const; + atom x \ (\, \, \, \, \, tid, cons , const, v, \); + \; \; \; ((x,b_of const, ([v]\<^sup>c\<^sup>e == [ V_cons tid cons [x]\<^sup>v]\<^sup>c\<^sup>e ) AND (c_of const x))#\<^sub>\\) ; \ \ s \ \ +\ \ + \; \; \; \; \ ; tid ; cons ; const ; v \ (AS_branch cons x s) \ \" + +|check_branch_list_consI: "\ + \; \; \; \; \; tid; cons; const; v \ cs \ \ ; + \; \; \; \; \; tid; dclist; v \ css \ \ +\ \ + \; \; \; \; \ ; tid ; (cons,const)#dclist ; v \ AS_cons cs css \ \" + +|check_branch_list_finalI: "\ + \; \;\; \; \; tid ; cons ; const ; v \ cs \ \ +\ \ + \; \; \; \; \ ; tid ; [(cons,const)] ; v \ AS_final cs \ \" + +| check_ifI: "\ + atom z \ (\, \, \, \, \, v , s1 , s2 , \ ); + (\; \; \ \ v \ (\ z : B_bool | TRUE \)) ; + \; \; \; \; \ \ s1 \ (\ z : b_of \ | ([v]\<^sup>c\<^sup>e == [[L_true]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \ z) \) ; + \; \; \; \; \ \ s2 \ (\ z : b_of \ | ([v]\<^sup>c\<^sup>e == [[L_false]\<^sup>v]\<^sup>c\<^sup>e) IMP (c_of \ z) \) +\ \ + \; \; \; \; \ \ IF v THEN s1 ELSE s2 \ \" + +| check_let2I: "\ + atom x \ (\, \, \, G, \, t, s1, \) ; + \; \ ; \ ; G; \ \ s1 \ t; + \; \ ; \ ; ((x,b_of t,c_of t x)#\<^sub>\G) ; \ \ s2 \ \ +\ \ + \; \ ; \ ; G ; \ \ (LET x : t = s1 IN s2) \ \" + +| check_varI: "\ + atom u \ (\, \, \, \, \, \', v, \) ; + \; \; \ \ v \ \'; + \; \; \; \ ; ((u,\') #\<^sub>\ \) \ s \ \ +\ \ + \; \; \; \; \ \ (VAR u : \' = v IN s) \ \" + +| check_assignI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + (u,\) \ setD \ ; + \; \; \ \ v \ \; + \; \; \ \ (\ z : B_unit | TRUE \) \ \' +\ \ + \; \; \; \; \ \ (u ::= v) \ \'" + +| check_whileI: "\ + \; \; \; \; \ \ s1 \ \ z : B_bool | TRUE \; + \; \; \; \; \ \ s2 \ \ z : B_unit | TRUE \; + \; \; \ \ (\ z : B_unit | TRUE \) \ \' +\ \ + \; \; \; \; \ \ WHILE s1 DO { s2 } \ \'" + +| check_seqI: "\ + \; \; \; \; \ \ s1 \ \ z : B_unit | TRUE \; + \; \; \; \; \ \ s2 \ \ +\ \ + \; \; \; \; \ \ s1 ;; s2 \ \" + +| check_caseI: "\ + \; \; \; \; \; tid ; dclist ; v \ cs \ \ ; + (AF_typedef tid dclist ) \ set \ ; + \; \; \ \ v \ \ z : B_id tid | TRUE \; + \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \ AS_match v cs \ \" + +equivariance check_s + +text \We only need avoidance for cases where a variable is added to a context\ +nominal_inductive check_s + avoids check_letI: x and z | check_branch_s_branchI: x | check_let2I: x | check_varI: u | check_ifI: z | check_assertI: x +proof(goal_cases) + case (1 x \ \ \ \ \ e \ z s b c) + hence "atom x \ AS_let x e s" using s_branch_s_branch_list.fresh(2) by auto + moreover have "atom z \ AS_let x e s" using s_branch_s_branch_list.fresh(2) 1 fresh_prod8 by auto + then show ?case using fresh_star_def 1 by force +next + case (3 x \ \ \ \ \ c \ s) + hence "atom x \ AS_assert c s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto + then show ?case using fresh_star_def 3 by force +next + case (5 \ \ \ \ \ const x \ tid cons v s) + hence "atom x \ AS_branch cons x s" using fresh_prodN s_branch_s_branch_list.fresh pure_fresh by auto + then show ?case using fresh_star_def 5 by force +next + case (7 z \ \ \ \ \ v s1 s2 \) + hence "atom z \ AS_if v s1 s2 " using s_branch_s_branch_list.fresh by auto + then show ?case using 7 fresh_prodN fresh_star_def by fastforce +next + case (9 x \ \ \ G \ t s1 \ s2) + hence "atom x \ AS_let2 x t s1 s2" using s_branch_s_branch_list.fresh by auto + thus ?case using fresh_star_def 9 by force +next + case (11 u \ \ \ \ \ \' v \ s) + hence "atom u \ AS_var u \' v s" using s_branch_s_branch_list.fresh by auto + then show ?case using fresh_star_def 11 by force + +qed(auto+) + +inductive_cases check_s_elims[elim!]: + "\; \; \; \; \ \ AS_val v \ \" + "\; \; \; \; \ \ AS_let x e s \ \" + "\; \; \; \; \ \ AS_if v s1 s2 \ \" + "\; \; \; \; \ \ AS_let2 x t s1 s2 \ \" + "\; \; \; \; \ \ AS_while s1 s2 \ \" + "\; \; \; \; \ \ AS_var u t v s \ \" + "\; \; \; \; \ \ AS_seq s1 s2 \ \" + "\; \; \; \; \ \ AS_assign u v \ \" + "\; \; \; \; \ \ AS_match v cs \ \" + "\; \; \; \; \ \ AS_assert c s \ \" + +inductive_cases check_branch_s_elims[elim!]: + "\; \; \; \; \; tid ; dclist ; v \ (AS_final cs) \ \" + "\; \; \; \; \; tid ; dclist ; v \ (AS_cons cs css) \ \" + "\; \; \; \; \; tid ; cons ; const ; v \ (AS_branch dc x s ) \ \" + +section \Programs\ + +text \Type check function bodies\ + +inductive check_funtyp :: "\ \ \ \ \ \ fun_typ \ bool" ( " _ ; _ ; _ \ _ " ) where + check_funtypI: "\ + atom x \ (\, \, B , b ); + \; \ ; B ; ((x,b,c) #\<^sub>\ GNil) ; []\<^sub>\ \ s \ \ +\ \ + \; \ ; B \ (AF_fun_typ x b c \ s)" + +equivariance check_funtyp +nominal_inductive check_funtyp + avoids check_funtypI: x +proof(goal_cases) + case (1 x \ \ B b c s \ ) + hence "atom x \ (AF_fun_typ x b c \ s)" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp + then show ?case using fresh_star_def 1 fresh_prodN by fastforce +next + case (2 \ \ x b c s \ f) + then show ?case by auto +qed + +inductive check_funtypq :: "\ \ \ \ fun_typ_q \ bool" ( " _ ; _ \ _ " ) where + check_fundefq_simpleI: "\ + \; \ ; {||} \ (AF_fun_typ x b c t s) +\ \ + \; \ \ ((AF_fun_typ_none (AF_fun_typ x b c t s)))" + +|check_funtypq_polyI: "\ + atom bv \ (\, \, (AF_fun_typ x b c t s)); + \; \; {|bv|} \ (AF_fun_typ x b c t s) +\ \ + \; \ \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" + +equivariance check_funtypq +nominal_inductive check_funtypq + avoids check_funtypq_polyI: bv +proof(goal_cases) + case (1 bv \ \ x b c t s ) + hence "atom bv \ (AF_fun_typ_some bv (AF_fun_typ x b c t s))" using fun_def.fresh fun_typ_q.fresh fun_typ.fresh by simp + thus ?case using fresh_star_def 1 fresh_prodN by fastforce +next + case (2 bv \ \ ft ) + then show ?case by auto +qed + +inductive check_fundef :: "\ \ \ \ fun_def \ bool" ( " _ ; _ \ _ " ) where + check_fundefI: "\ + \; \ \ ft +\ \ + \; \ \ (AF_fundef f ft)" + +equivariance check_fundef +nominal_inductive check_fundef . + +text \Temporarily remove this simproc as it produces untidy eliminations\ +declare[[ simproc del: alpha_lst]] + +inductive_cases check_funtyp_elims[elim!]: + "check_funtyp \ \ B ft" + +inductive_cases check_funtypq_elims[elim!]: + "check_funtypq \ \ (AF_fun_typ_none (AF_fun_typ x b c \ s))" + "check_funtypq \ \ (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" + +inductive_cases check_fundef_elims[elim!]: + "check_fundef \ \ (AF_fundef f ftq)" + +declare[[ simproc add: alpha_lst]] + +nominal_function \_of :: "var_def list \ \" where + "\_of [] = DNil" +| "\_of ((AV_def u t v)#vs) = (u,t) #\<^sub>\ (\_of vs)" + apply auto + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust apply force + using eqvt_def \_of_graph_aux_def neq_Nil_conv old.prod.exhaust + by (metis var_def.strong_exhaust) +nominal_termination (eqvt) by lexicographic_order + +inductive check_prog :: "p \ \ \ bool" ( "\ _ \ _") where + "\ + \; \; {||}; GNil ; \_of \ \ s \ \ +\ \ \ (AP_prog \ \ \ s) \ \" + +inductive_cases check_prog_elims[elim!]: + "\ (AP_prog \ \ \ s) \ \" + +end \ No newline at end of file diff --git a/thys/MiniSail/TypingL.thy b/thys/MiniSail/TypingL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/TypingL.thy @@ -0,0 +1,2800 @@ +(*<*) +theory TypingL + imports Typing RCLogicL "HOL-Eisbach.Eisbach" +begin + (*>*) + +chapter \Typing Lemmas\ + +section \Prelude\ + +text \Needed as the typing elimination rules give us facts for an alpha-equivalent version of a term + and so need to be able to 'jump back' to a typing judgement for the orginal term\ + +lemma \_fresh_c[simp]: + assumes "atom x \ \ z : b | c \" and "atom z \ x" + shows "atom x \ c" + using \.fresh assms fresh_at_base + by (simp add: fresh_at_base(2)) + +lemmas subst_defs = subst_b_b_def subst_b_c_def subst_b_\_def subst_v_v_def subst_v_c_def subst_v_\_def + +lemma wfT_wfT_if1: + assumes "wfT \ \ \ (\ z : b_of t | CE_val v == CE_val (V_lit L_false) IMP c_of t z \)" and "atom z \ (\,t)" + shows "wfT \ \ \ t" + using assms proof(nominal_induct t avoiding: \ z rule: \.strong_induct) + case (T_refined_type z' b' c') + show ?case proof(rule wfT_wfT_if) + show \ \; \; \ \\<^sub>w\<^sub>f \ z : b' | [ v ]\<^sup>c\<^sup>e == [ [ L_false ]\<^sup>v ]\<^sup>c\<^sup>e IMP c'[z'::=[ z]\<^sup>v]\<^sub>c\<^sub>v \ \ + using T_refined_type b_of.simps c_of.simps subst_defs by metis + show \atom z \ (c', \)\ using T_refined_type fresh_prodN \_fresh_c by metis + qed +qed + +lemma fresh_u_replace_true: + fixes bv::bv and \::\ + assumes "atom bv \ \' @ (x, b, c) #\<^sub>\ \" + shows "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" + using fresh_append_g fresh_GCons assms fresh_Pair c.fresh(1) by auto + +lemma wf_replace_true1: + fixes \::\ and \::\ and \::\ and \'::\ and v::v and e::e and c::c and c''::c and c'::c and \::\ and ts::"(string*\) list" and \::\ and b'::b and b::b and s::s + and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list + +shows "\; \; G \\<^sub>w\<^sub>f v : b' \ G = \' @ (x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f v : b'" and + "\; \; G \\<^sub>w\<^sub>f c'' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f c''" and + "\ ; \ \\<^sub>w\<^sub>f G \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ \\<^sub>w\<^sub>f \' @ ((x, b, TRUE) #\<^sub>\ \) " and + "\; \; G \\<^sub>w\<^sub>f \ \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f P \ True" and + "\ ; \ \\<^sub>w\<^sub>f b \ True" and + "\ ; \ ; G \\<^sub>w\<^sub>f ce : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f ce : b'" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b' and c'' and G and \ and ts and P and b and b' and td + arbitrary: \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfB_intI \ \) + then show ?case using wf_intros by metis +next + case (wfB_boolI \ \) + then show ?case using wf_intros by metis +next + case (wfB_unitI \ \) + then show ?case using wf_intros by metis +next + case (wfB_bitvecI \ \) + then show ?case using wf_intros by metis +next + case (wfB_pairI \ \ b1 b2) + then show ?case using wf_intros by metis +next + case (wfB_consI \ s dclist \) + then show ?case using wf_intros by metis +next + case (wfB_appI \ b s bv dclist \) + then show ?case using wf_intros by metis +next + case (wfV_varI \ \ \'' b' c x') + hence wfg: \ \ ; \ \\<^sub>w\<^sub>f \' @ (x, b, TRUE) #\<^sub>\ \ \ by auto + show ?case proof(cases "x=x'") + case True + hence "Some (b, TRUE) = lookup (\' @ (x, b, TRUE) #\<^sub>\ \) x'" using lookup.simps lookup_inside_wf wfg by simp + thus ?thesis using Wellformed.wfV_varI[OF wfg] + by (metis True lookup_inside_wf old.prod.inject option.inject wfV_varI.hyps(1) wfV_varI.hyps(3) wfV_varI.prems) + next + case False + hence "Some (b', c) = lookup (\' @ (x, b, TRUE) #\<^sub>\ \) x'" using lookup_inside2 wfV_varI by metis + then show ?thesis using Wellformed.wfV_varI[OF wfg] + by (metis wfG_elim2 wfG_suffix wfV_varI.hyps(1) wfV_varI.hyps(2) wfV_varI.hyps(3) + wfV_varI.prems Wellformed.wfV_varI wf_replace_inside(1)) + qed +next + case (wfV_litI \ \ \ l) + then show ?case using wf_intros using wf_intros by metis +next + case (wfV_pairI \ \ \ v1 b1 v2 b2) + then show ?case using wf_intros by metis +next + case (wfV_consI s dclist \ dc x b' c \ \ v) + then show ?case using wf_intros by metis +next + case (wfV_conspI s bv dclist \ dc xc bc cc \ b' \'' v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by metis + show \(dc, \ xc : bc | cc \) \ set dclist\ using wfV_conspI by metis + show \ \ ;\ \\<^sub>w\<^sub>f b' \ using wfV_conspI by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : bc[bv::=b']\<^sub>b\<^sub>b \ using wfV_conspI by metis + have "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" using fresh_u_replace_true wfV_conspI by metis + thus \atom bv \ (\, \, \' @ (x, b, TRUE) #\<^sub>\ \, b', v)\ using wfV_conspI fresh_prodN by metis + qed +next + case (wfCE_valI \ \ \ v b) + then show ?case using wf_intros by metis +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_eqI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfCE_lenI \ \ \ v1) + then show ?case using wf_intros by metis +next + case (wfTI z \ \ \'' b' c') + show ?case proof + show \atom z \ (\, \, \' @ (x, b, TRUE) #\<^sub>\ \)\ using wfTI fresh_append_g fresh_GCons fresh_prodN by auto + show \ \ ; \ \\<^sub>w\<^sub>f b' \ using wfTI by metis + show \ \; \; (z, b', TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c' \ using wfTI append_g.simps by metis + qed +next + case (wfC_eqI \ \ \ e1 b e2) + then show ?case using wf_intros by metis +next + case (wfC_trueI \ \ \) + then show ?case using wf_intros by metis +next + case (wfC_falseI \ \ \) + then show ?case using wf_intros by metis +next + case (wfC_conjI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfC_disjI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfC_notI \ \ \ c1) + then show ?case using wf_intros by metis +next + case (wfC_impI \ \ \ c1 c2) + then show ?case using wf_intros by metis +next + case (wfG_nilI \ \) + then show ?case using GNil_append by blast +next + case (wfG_cons1I c \ \ \'' x b) + then show ?case using wf_intros wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix + by (metis (no_types, lifting)) +next + case (wfG_cons2I c \ \ \'' x' b) + then show ?case using wf_intros + by (metis wfG_cons_TRUE2 wfG_elims(2) wfG_replace_inside wfG_suffix) +next + case wfTh_emptyI + then show ?case using wf_intros by metis +next + case (wfTh_consI tdef \) + then show ?case using wf_intros by metis +next + case (wfTD_simpleI \ lst s) + then show ?case using wf_intros by metis +next + case (wfTD_poly \ bv lst s) + then show ?case using wf_intros by metis +next + case (wfTs_nil \ \ \) + then show ?case using wf_intros by metis +next + case (wfTs_cons \ \ \ \ dc ts) + then show ?case using wf_intros by metis +qed + +lemma wf_replace_true2: + fixes \::\ and \::\ and \::\ and \'::\ and v::v and e::e and c::c and c''::c and c'::c and \::\ and ts::"(string*\) list" and \::\ and b'::b and b::b and s::s + and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list + +shows "\ ; \ ; \ ; G ; D \\<^sub>w\<^sub>f e : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \); D \\<^sub>w\<^sub>f e : b'" and + "\ ; \ ; \ ; G ; \ \\<^sub>w\<^sub>f s : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>f s : b'" and + "\ ; \ ; \ ; G ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b'" and + "\ ; \ ; \ ; G ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b' \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b'" and + +"\ \\<^sub>w\<^sub>f \ \ True" and +"\; \; G \\<^sub>w\<^sub>f \ \ G = \' @(x, b, c) #\<^sub>\ \ \ \ ; \ ; \' @ ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and + +"\ ; \ \\<^sub>w\<^sub>f ftq \ True" and +"\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b' and b' and b' and b' and \ and \ and ftq and ft + arbitrary: \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' and \ \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_intros using wf_intros wf_replace_true1 by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfE_appPI \ \ \ \'' \ b' bv v \ f x1 b1 c1 s) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by metis + show \ \ ; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by metis + have "atom bv \ \' @ (x, b, TRUE) #\<^sub>\ \" using fresh_u_replace_true wfE_appPI fresh_prodN by metis + thus \atom bv \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, b', v, (b_of \)[bv::=b']\<^sub>b)\ + using wfE_appPI fresh_prodN by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \ s))) = lookup_fun \ f\ using wfE_appPI by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \ using wfE_appPI wf_replace_true1 by metis + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + then show ?case using wf_intros wf_replace_true1 by metis +next + + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_letI \ \ \ \'' \ e b' x1 s b1) + show ?case proof + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI wf_replace_true1 by metis + have \ \ ; \ ; \ ; ((x1, b', TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b1 \ apply(rule wfS_letI(4)) + using wfS_letI append_g.simps by simp + thus \ \ ; \ ; \ ; (x1, b', TRUE) #\<^sub>\ \'@ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b1 \ using append_g.simps by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_letI by metis + show "atom x1 \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, e, b1)" using fresh_append_g fresh_GCons fresh_prodN wfS_letI by auto + qed +next + case (wfS_assertI \ \ \ x' c \'' \ s b') + show ?case proof + show \ \ ; \ ; \ ; (x', B_bool, c) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ + using wfS_assertI (2)[of "(x', B_bool, c) #\<^sub>\ \'" \] wfS_assertI by simp + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c \ using wfS_assertI wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_assertI by metis + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, c, b', s)\ using wfS_assertI fresh_prodN by simp + qed +next + case (wfS_let2I \ \ \ \'' \ s1 \ x' s2 ba') + show ?case proof + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_replace_true1 by metis + have \ \ ; \ ; \ ; ((x', b_of \, TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : ba' \ + apply(rule wfS_let2I(5)) + using wfS_let2I append_g.simps by auto + thus \ \ ; \ ; \ ; (x', b_of \, TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : ba' \ using wfS_let2I append_g.simps by auto + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, s1, ba', \)\ using fresh_append_g fresh_GCons fresh_prodN wfS_let2I by auto + qed +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_varI \ \ \'' \ v u \ \ b' s) + show ?case proof + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_varI wf_replace_true1 by metis + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI wf_replace_true1 by metis + show \atom u \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, \, v, b')\ using wfS_varI u_fresh_g fresh_prodN by auto + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b' \ using wfS_varI by metis + qed + +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros by metis +next + case (wfS_matchI \ \ \'' v tid dclist \ \ cs b') + show ?case proof + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f v : B_id tid \ using wfS_matchI wf_replace_true1 by auto + show \AF_typedef tid dclist \ set \\ using wfS_matchI by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \ \ ; \ ; \ ; \' @ (x, b, TRUE) #\<^sub>\ \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b' \ using wfS_matchI by auto + qed +next + case (wfS_branchI \ \ \ x' \ \'' \ s b' tid dc) + show ?case proof + have \ \ ; \ ; \ ; ((x', b_of \, TRUE) #\<^sub>\ \') @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ using wfS_branchI append_g.simps by metis + thus \ \ ; \ ; \ ; (x', b_of \, TRUE) #\<^sub>\ \' @ (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b' \ using wfS_branchI append_g.simps append_g.simps by metis + show \atom x' \ (\, \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \, \' @ (x, b, TRUE) #\<^sub>\ \, \)\ using wfS_branchI by auto + show \ \; \; \' @ (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfS_branchI by auto + qed +next + case (wfS_finalI \ \ \ \ \ tid dc t cs b) + then show ?case using wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dc t cs b dclist css) + then show ?case using wf_intros by metis +next + case (wfD_emptyI \ \ \) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfD_cons \ \ \ \ \ u) + then show ?case using wf_intros wf_replace_true1 by metis +next + case (wfPhi_emptyI \) + then show ?case using wf_intros by metis +next + case (wfPhi_consI f \ \ ft) + then show ?case using wf_intros by metis +next + case (wfFTNone \ \ ft) + then show ?case using wf_intros by metis +next + case (wfFTSome \ \ bv ft) + then show ?case using wf_intros by metis +next + case (wfFTI \ B b \ x c s \) + then show ?case using wf_intros by metis +qed + +lemmas wf_replace_true = wf_replace_true1 wf_replace_true2 + +section \Subtyping\ + +lemma subtype_reflI2: + fixes \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" + shows "\; \; \ \ \ \ \" +proof - + obtain z b c where *:"\ = \ z : b | c \ \ atom z \ (\,\,\) \ \; \; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c" + using wfT_elims(1)[OF assms] by metis + obtain x::x where **: "atom x \ (\, \, \, c, z ,c, z , c )" using obtain_fresh by metis + have "\; \; \ \ \ z : b | c \ \ \ z : b | c \" proof + show "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" using * assms by auto + show "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" using * assms by auto + show "atom x \ (\, \, \, z , c , z , c )" using fresh_prod6 fresh_prod5 ** by metis + thus "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var x]\<^sub>v " using wfT_wfC_cons assms * ** subst_v_c_def by simp + qed + thus ?thesis using * by auto +qed + +lemma subtype_reflI: + assumes "\ z1 : b | c1 \ = \ z2 : b | c2 \" and wf1: "\; \;\ \\<^sub>w\<^sub>f (\ z1 : b | c1 \)" + shows "\; \; \ \ (\ z1 : b | c1 \) \ (\ z2 : b | c2 \)" + using assms subtype_reflI2 by metis + +nominal_function base_eq :: "\ \ \ \ \ \ bool" where + "base_eq _ \ z1 : b1| c1 \ \ z2 : b2 | c2 \ = (b1 = b2)" + apply(auto,simp add: eqvt_def base_eq_graph_aux_def ) + by (meson \.exhaust) +nominal_termination (eqvt) by lexicographic_order + +lemma subtype_wfT: + fixes t1::\ and t2::\ + assumes "\; \; \ \ t1 \ t2" + shows "\; \; \ \\<^sub>w\<^sub>f t1 \ \; \; \ \\<^sub>w\<^sub>f t2" + using assms subtype_elims by metis + +lemma subtype_eq_base: + assumes "\; \; \ \ (\ z1 : b1| c1 \) \ (\ z2 : b2 | c2 \)" + shows "b1=b2" + using subtype.simps assms by auto + +lemma subtype_eq_base2: + assumes "\; \; \ \ t1 \ t2" + shows "b_of t1 = b_of t2" + using assms proof(rule subtype.induct[of \ \ \ t1 t2],goal_cases) + case (1 \ \ z1 b c1 z2 c2 x) + then show ?case using subtype_eq_base by auto +qed + +lemma subtype_wf: + fixes \1::\ and \2::\ + assumes "\; \; \ \ \1 \ \2" + shows "\; \; \ \\<^sub>w\<^sub>f \1 \\; \; \ \\<^sub>w\<^sub>f \2" + using assms +proof(rule subtype.induct[of \ \ \ \1 \2],goal_cases) + case (1 \ \G z1 b c1 z2 c2 x) + then show ?case by blast +qed + +lemma subtype_g_wf: + fixes \1::\ and \2::\ and \::\ + assumes "\; \; \ \ \1 \ \2" + shows "\ ; \\\<^sub>w\<^sub>f \" + using assms +proof(rule subtype.induct[of \ \ \ \1 \2],goal_cases) + case (1 \ \ \ z1 b c1 z2 c2 x) + then show ?case using wfX_wfY by auto +qed + +text \ For when we have a particular y that satisfies the freshness conditions that we want the validity check to use \ + +lemma valid_flip_simple: + assumes "\; \; (z, b, c) #\<^sub>\ \ \ c'" and "atom z \ \" and "atom x \ (z, c, z, c', \)" + shows "\; \; (x, b, (z \ x ) \ c) #\<^sub>\ \ \ (z \ x ) \ c'" +proof - + have "(z \ x ) \ \; \; (z \ x ) \ ((z, b, c) #\<^sub>\ \) \ (z \ x ) \ c'" + using True_eqvt valid.eqvt assms beta_flip_eq wfX_wfY by metis + moreover have " \\<^sub>w\<^sub>f \" using valid.simps wfC_wf wfG_wf assms by metis + ultimately show ?thesis + using theta_flip_eq G_cons_flip_fresh3[of x \ z b c] assms fresh_Pair flip_commute by metis +qed + +lemma valid_wf_all: + assumes " \; \; (z0,b,c0)#\<^sub>\G \ c" + shows "wfG \ \ G" and "wfC \ \ ((z0,b,c0)#\<^sub>\G) c" and "atom z0 \ G" + using valid.simps wfC_wf wfG_cons assms by metis+ + +lemma valid_wfT: + fixes z::x + assumes " \; \; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\G \ c[z::=V_var z0]\<^sub>v" and "atom z0 \ (\, \, G,c,c0)" + shows "\; \; G \\<^sub>w\<^sub>f \ z : b | c0 \" and "\; \; G \\<^sub>w\<^sub>f \ z : b | c \" +proof - + have "atom z0 \ c0" using assms fresh_Pair by auto + moreover have *:" \ ; \ \\<^sub>w\<^sub>f (z0,b,c0[z::=V_var z0]\<^sub>c\<^sub>v)#\<^sub>\G" using valid_wf_all wfX_wfY assms subst_v_c_def by metis + ultimately show wft: "\; \; G \\<^sub>w\<^sub>f \ z : b | c0 \" using wfG_wfT[OF *] by auto + + have "atom z0 \ c" using assms fresh_Pair by auto + moreover have wfc: "\; \; (z0,b,c0[z::=V_var z0]\<^sub>v)#\<^sub>\G \\<^sub>w\<^sub>f c[z::=V_var z0]\<^sub>v" using valid_wf_all assms by metis + have "\; \; G \\<^sub>w\<^sub>f \ z0 : b | c[z::=V_var z0]\<^sub>v \" proof + show \atom z0 \ (\, \, G)\ using assms fresh_prodN by simp + show \ \ ; \ \\<^sub>w\<^sub>f b \ using wft wfT_wfB by force + show \ \; \; (z0, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c[z::=[ z0 ]\<^sup>v]\<^sub>v \ using wfc wfC_replace_inside[OF wfc, of GNil z0 b "c0[z::=[ z0 ]\<^sup>v]\<^sub>v" G C_true] wfC_trueI + append_g.simps + by (metis "local.*" wfG_elim2 wf_trans(2)) + qed + moreover have "\ z0 : b | c[z::=V_var z0]\<^sub>v \ = \ z : b | c \" using \atom z0 \ c0\ \.eq_iff Abs1_eq_iff(3) + using calculation(1) subst_v_c_def by auto + ultimately show "\; \; G \\<^sub>w\<^sub>f \ z : b | c \" by auto +qed + +lemma valid_flip: + fixes c::c and z::x and z0::x and xx2::x + assumes " \; \; (xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ \ \ c[z::=V_var xx2]\<^sub>v" and + "atom xx2 \ (c0,\,c,z) " and "atom z0 \ (\,c,z)" + shows " \; \; (z0, b, c0) #\<^sub>\ \ \ c[z::=V_var z0]\<^sub>v" +proof - + + have " \\<^sub>w\<^sub>f \" using assms valid_wf_all wfX_wfY by metis + hence " \ ; \ ; (xx2 \ z0 ) \ ((xx2, b, c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ \) \ ((xx2 \ z0 ) \ c[z::=V_var xx2]\<^sub>v)" + using valid.eqvt True_eqvt assms beta_flip_eq theta_flip_eq by metis + hence " \; \; (((xx2 \ z0 ) \ xx2, b, (xx2 \ z0 ) \ c0[z0::=V_var xx2]\<^sub>v) #\<^sub>\ (xx2 \ z0 ) \ \) \ ((xx2 \ z0 ) \(c[z::=V_var xx2]\<^sub>v))" + using G_cons_flip[of xx2 z0 xx2 b "c0[z0::=V_var xx2]\<^sub>v" \] by auto + moreover have "(xx2 \ z0 ) \ xx2 = z0" by simp + moreover have "(xx2 \ z0 ) \ c0[z0::=V_var xx2]\<^sub>v = c0" + using assms subst_cv_v_flip[of xx2 c0 z0 "V_var z0"] assms fresh_prod4 by auto + moreover have "(xx2 \ z0 ) \ \ = \" proof - + have "atom xx2 \ \" using assms by auto + moreover have "atom z0 \ \" using assms by auto + ultimately show ?thesis using flip_fresh_fresh by auto + qed + moreover have "(xx2 \ z0 ) \ (c[z::=V_var xx2]\<^sub>v) =c[z::=V_var z0]\<^sub>v" + using subst_cv_v_flip3 assms by simp + ultimately show ?thesis by auto +qed + +lemma subtype_valid: + assumes "\; \; \ \ t1 \ t2" and "atom y \ \" and "t1 = \ z1 : b | c1 \" and "t2 = \ z2 : b | c2 \" + shows "\; \; ((y, b, c1[z1::=V_var y]\<^sub>v) #\<^sub>\ \) \ c2[z2::=V_var y]\<^sub>v" + using assms proof(nominal_induct t2 avoiding: y rule: subtype.strong_induct) + case (subtype_baseI x \ \ \ z c z' c' ba) + + hence "(x \ y) \ \ ; (x \ y) \ \ ; (x \ y) \ ((x, ba, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \) \ (x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v" using valid.eqvt + using permute_boolI by blast + moreover have " \\<^sub>w\<^sub>f \" using valid.simps wfC_wf wfG_wf subtype_baseI by metis + ultimately have "\; \; ((y, ba, (x \ y) \ c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \) \ (x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v" + using subtype_baseI theta_flip_eq beta_flip_eq \.eq_iff G_cons_flip_fresh3[of y \ x ba] by (metis flip_commute) + moreover have "(x \ y) \ c[z::=[ x ]\<^sup>v]\<^sub>v = c1[z1::=[ y ]\<^sup>v]\<^sub>v" + by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def) + moreover have "(x \ y) \ c'[z'::=[ x ]\<^sup>v]\<^sub>v = c2[z2::=[ y ]\<^sup>v]\<^sub>v" + by (metis subtype_baseI permute_flip_cancel subst_cv_id subst_cv_v_flip3 subst_cv_var_flip type_eq_subst_eq wfT_fresh_c subst_v_c_def) + + ultimately show ?case using subtype_baseI \.eq_iff by metis +qed + +lemma subtype_valid_simple: + assumes "\; \; \ \ t1 \ t2" and "atom z \ \" and "t1 = \ z : b | c1 \" and "t2 = \ z : b | c2 \" + shows "\; \; ((z, b, c1) #\<^sub>\ \) \ c2" + using subst_v_c_def subst_v_id assms subtype_valid[OF assms] by simp + +lemma obtain_for_t_with_fresh: + assumes "atom x \ t" + shows "\b c. t = \ x : b | c \" +proof - + obtain z1 b1 c1 where *:" t = \ z1 : b1 | c1 \ \ atom z1 \ t" using obtain_fresh_z by metis + then have "t = (x \ z1) \ t" using flip_fresh_fresh assms by metis + also have "... = \ (x \ z1) \ z1 : (x \ z1) \ b1 | (x \ z1) \ c1 \" using * assms by simp + also have "... = \ x : b1 | (x \ z1) \ c1 \" using * assms by auto + finally show ?thesis by auto +qed + +lemma subtype_trans: + assumes "\; \; \ \ \1 \ \2" and "\; \; \ \ \2 \ \3" + shows "\; \; \ \ \1 \ \3" + using assms proof(nominal_induct avoiding: \3 rule: subtype.strong_induct) + case (subtype_baseI x \ \ \ z c z' c' b) + hence "b_of \3 = b" using subtype_eq_base2 b_of.simps by metis + then obtain z'' c'' where t3: "\3 = \ z'' : b | c''\ \ atom z'' \ x" + using obtain_fresh_z2 by metis + hence xf: "atom x \ (z'', c'')" using fresh_prodN subtype_baseI \.fresh by auto + have "\; \; \ \ \ z : b | c \ \ \ z'' : b | c''\" + proof(rule Typing.subtype_baseI) + show \atom x \ (\, \, \, z, c, z'', c'')\ using t3 fresh_prodN subtype_baseI xf by simp + show \ \; \; \ \\<^sub>w\<^sub>f \ z : b | c \ \ using subtype_baseI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ z'' : b | c'' \ \ using subtype_baseI t3 subtype_elims by metis + have " \; \; (x, b, c'[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c''[z''::=[ x ]\<^sup>v]\<^sub>v " + using subtype_valid[OF \\; \; \ \ \ z' : b | c' \ \ \3\ , of x z' b c' z'' c''] subtype_baseI + t3 by simp + thus \\; \; (x, b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ c''[z''::=[ x ]\<^sup>v]\<^sub>v \ + using valid_trans_full[of \ \ x b c z \ c' z' c'' z'' ] subtype_baseI t3 by simp + qed + thus ?case using t3 by simp +qed + +lemma subtype_eq_e: + assumes "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" and "atom z1 \ e1" and "atom z2 \ e2" and "atom z1 \ \" and "atom z2 \ \" + and "wfCE P \ \ e1 b" and "wfCE P \ \ e2 b" + shows "P; \; \ \ \ z1 : b | CE_val (V_var z1) == e1 \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" +proof - + + have "wfCE P \ \ e1 b" and "wfCE P \ \ e2 b" using assms by auto + + have wst1: "wfT P \ \ (\ z1 : b | CE_val (V_var z1) == e1 \)" + using wfC_e_eq wfTI assms wfX_wfB wfG_fresh_x + by (simp add: wfT_e_eq) + + moreover have wst2:"wfT P \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" + using wfC_e_eq wfX_wfB wfTI assms wfG_fresh_x + by (simp add: wfT_e_eq) + + moreover obtain x::x where xf: "atom x \ (P, \ , z1, CE_val (V_var z1) == e1 , z2, CE_val (V_var z2) == e2 , \)" using obtain_fresh by blast + moreover have vld: "P; \ ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ (CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v " (is "P; \ ; ?G \ ?c") + proof - + + have wbg: "P; \ \\<^sub>w\<^sub>f ?G \ P ; \ \\<^sub>w\<^sub>f \ \ toSet \ \ toSet ?G" proof - + have "P; \ \\<^sub>w\<^sub>f ?G" proof(rule wfG_consI) + show "P; \ \\<^sub>w\<^sub>f \" using assms wfX_wfY by metis + show "atom x \ \" using xf by auto + show "P; \ \\<^sub>w\<^sub>f b " using assms(6) wfX_wfB by auto + show "P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v " + using wfC_e_eq[OF assms(6)] wf_subst(2) + by (simp add: \atom x \ \\ assms(2) subst_v_c_def) + qed + moreover hence "P; \ \\<^sub>w\<^sub>f \" using wfG_elims by metis + ultimately show ?thesis using toSet.simps by auto + qed + + have wsc: "wfC P \ ?G ?c" proof - + have "wfCE P \ ?G (CE_val (V_var x)) b" proof + show \ P; \ ; (x, b, (CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f V_var x : b \ using wfV_varI lookup.simps wbg by auto + qed + moreover have "wfCE P \ ?G e2 b" using wf_weakening assms wbg by metis + ultimately have "wfC P \ ?G (CE_val (V_var x) == e2 )" using wfC_eqI by simp + thus ?thesis using subst_cv.simps(6) \atom z2 \ e2\ subst_v_c_def by simp + qed + + moreover have "\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI , rule impI) + fix i + assume as: "wfI P ?G i \ is_satis_g i ?G" + hence "is_satis i ((CE_val (V_var z1) == e1 )[z1::=V_var x]\<^sub>v)" + by (simp add: is_satis_g.simps(2)) + hence "is_satis i (CE_val (V_var x) == e1 )" using subst_cv.simps assms subst_v_c_def by auto + then obtain s1 and s2 where *:"eval_e i (CE_val (V_var x)) s1 \ eval_e i e1 s2 \ s1=s2" using is_satis.simps eval_c_elims by metis + moreover hence "eval_e i e2 s1" proof - + have **:"wfI P ?G i" using as by auto + moreover have "wfCE P \ ?G e1 b \ wfCE P \ ?G e2 b" using assms xf wf_weakening wbg by metis + moreover then obtain s2' where "eval_e i e2 s2'" using assms wfI_wfCE_eval_e ** by metis + ultimately show ?thesis using * assms(1) wfX_wfY by metis + qed + ultimately have "is_satis i (CE_val (V_var x) == e2 )" using is_satis.simps eval_c_eqI by force + thus "is_satis i ((CE_val (V_var z2) == e2 )[z2::=V_var x]\<^sub>v)" using is_satis.simps eval_c_eqI assms subst_cv.simps subst_v_c_def by auto + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, \, z1 , CE_val (V_var z1) == e1, z2, CE_val (V_var z2) == e2 ) " + unfolding fresh_prodN using xf fresh_prod7 \.fresh by fast + ultimately show ?thesis using subtype_baseI[OF _ wst1 wst2 vld] xf by simp +qed + +lemma subtype_eq_e_nil: + assumes "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i e1 s1 \ eval_e i e2 s2 \ s1 = s2" and "supp e1 = {}" and "supp e2 = {}" and "wfTh P" + and "wfCE P \ GNil e1 b" and "wfCE P \ GNil e2 b" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ \ z1 : b | CE_val (V_var z1) == e1 \ \ (\ z2 : b | CE_val (V_var z2) == e2 \)" + apply(rule subtype_eq_e,auto simp add: assms e.fresh) + using assms fresh_def e.fresh supp_GNil by metis+ + +lemma subtype_gnil_fst_aux: + assumes "supp v\<^sub>1 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \ GNil (CE_val v\<^sub>1) b" and "wfCE P \ GNil (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) b" and + "wfCE P \ GNil (CE_val v\<^sub>2) b2" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>1 \) \ (\ z2 : b | CE_val (V_var z2) == CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e \)" +proof - + have "\i s1 s2 G . wfG P \ G \ wfI P G i \ eval_e i ( CE_val v\<^sub>1) s1 \ eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2 \ s1 = s2" proof(rule+) + fix i s1 s2 G + assume as: "wfG P \ G \ wfI P G i \ eval_e i (CE_val v\<^sub>1) s1 \ eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s2" + hence "wfCE P \ G (CE_val v\<^sub>2) b2" using assms wf_weakening + by (metis empty_subsetI toSet.simps(1)) + then obtain s3 where "eval_e i (CE_val v\<^sub>2) s3" using wfI_wfCE_eval_e as by metis + hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s1 s3)" + by (meson as eval_e_elims(1) eval_v_pairI) + hence "eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1" using eval_e_fstI eval_e_valI by metis + show "s1 = s2" using as eval_e_uniqueness + using \eval_e i (CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\ by auto + qed + thus ?thesis using subtype_eq_e_nil ce.supp assms by auto +qed + +lemma subtype_gnil_snd_aux: + assumes "supp v\<^sub>2 = {}" and "supp (V_pair v\<^sub>1 v\<^sub>2) = {}" and "wfTh P" and "wfCE P \ GNil (CE_val v\<^sub>2) b" and + "wfCE P \ GNil (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) b" and + "wfCE P \ GNil (CE_val v\<^sub>1) b2" and "atom z1 \ GNil" and "atom z2 \ GNil" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \) \ (\ z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \)" +proof - + have "\i s1 s2 G. wfG P \ G \ wfI P G i \ eval_e i ( CE_val v\<^sub>2) s1 \ eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2 \ s1 = s2" proof(rule+) + fix i s1 s2 G + assume as: " wfG P \ G \ wfI P G i \ eval_e i (CE_val v\<^sub>2) s1 \ eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s2" + hence "wfCE P \ G (CE_val v\<^sub>1) b2" using assms wf_weakening + by (metis empty_subsetI toSet.simps(1)) + then obtain s3 where "eval_e i (CE_val v\<^sub>1) s3" using wfI_wfCE_eval_e as by metis + hence "eval_v i ((V_pair v\<^sub>1 v\<^sub>2)) (SPair s3 s1)" + by (meson as eval_e_elims eval_v_pairI) + hence "eval_e i (CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e) s1" using eval_e_sndI eval_e_valI by metis + show "s1 = s2" using as eval_e_uniqueness + using \eval_e i (CE_snd [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e) s1\ by auto + qed + thus ?thesis using assms subtype_eq_e_nil by (simp add: ce.supp ce.supp) +qed + +lemma subtype_gnil_fst: + assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f [#1[[v\<^sub>1,v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e : b" + shows "\ ; {||} ; GNil \ (\ z\<^sub>1 : b | [[z\<^sub>1]\<^sup>v]\<^sup>c\<^sup>e == [v\<^sub>1]\<^sup>c\<^sup>e \) \ + (\ z\<^sub>2 : b | [[z\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e == [#1[[v\<^sub>1, v\<^sub>2]\<^sup>v]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e \)" +proof - + obtain b2 where **:" \ ; {||} ; GNil \\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b b2" using wfCE_elims(4)[OF assms ] wfCE_elims by metis + obtain b1' b2' where *:"B_pair b b2 = B_pair b1' b2' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>1 : b1' \ \ ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>2 : b2'" + using wfV_elims(3)[OF **] by metis + + show ?thesis proof(rule subtype_gnil_fst_aux) + show \supp v\<^sub>1 = {}\ using * wfV_supp_nil by auto + show \supp (V_pair v\<^sub>1 v\<^sub>2) = {}\ using ** wfV_supp_nil e.supp by metis + show \\\<^sub>w\<^sub>f \\ using assms wfX_wfY by metis + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>1 : b \ using wfCE_valI "*" by auto + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_fst [V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e : b \ using assms by auto + show \\; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>2 : b2 \using wfCE_valI "*" by auto + show \atom z\<^sub>1 \ GNil\ using fresh_GNil by metis + show \atom z\<^sub>2 \ GNil\ using fresh_GNil by metis + qed +qed + +lemma subtype_gnil_snd: + assumes "wfCE P {||} GNil (CE_snd ([V_pair v\<^sub>1 v\<^sub>2]\<^sup>c\<^sup>e)) b" + shows "P ; {||} ; GNil \ (\ z1 : b | CE_val (V_var z1) == CE_val v\<^sub>2 \) \ (\ z2 : b | CE_val (V_var z2) == CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e \)" +proof - + obtain b1 where **:" P ; {||} ; GNil \\<^sub>w\<^sub>f V_pair v\<^sub>1 v\<^sub>2 : B_pair b1 b " using wfCE_elims assms by metis + obtain b1' b2' where *:"B_pair b1 b = B_pair b1' b2' \ P ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>1 : b1' \ P ; {||} ; GNil \\<^sub>w\<^sub>f v\<^sub>2 : b2'" using wfV_elims(3)[OF **] by metis + + show ?thesis proof(rule subtype_gnil_snd_aux) + show \supp v\<^sub>2 = {}\ using * wfV_supp_nil by auto + show \supp (V_pair v\<^sub>1 v\<^sub>2) = {}\ using ** wfV_supp_nil e.supp by metis + show \\\<^sub>w\<^sub>f P\ using assms wfX_wfY by metis + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>1 : b1 \ using wfCE_valI "*" by simp + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_snd [(V_pair v\<^sub>1 v\<^sub>2)]\<^sup>c\<^sup>e : b \ using assms by auto + show \P; {||}; GNil \\<^sub>w\<^sub>f CE_val v\<^sub>2 : b \using wfCE_valI "*" by simp + show \atom z1 \ GNil\ using fresh_GNil by metis + show \atom z2 \ GNil\ using fresh_GNil by metis + qed +qed + +lemma subtype_fresh_tau: + fixes x::x + assumes "atom x \ t1" and "atom x \ \" and "P; \; \ \ t1 \ t2" + shows "atom x \ t2" +proof - + have wfg: "P; \ \\<^sub>w\<^sub>f \" using subtype_wf wfX_wfY assms by metis + have wft: "wfT P \ \ t2" using subtype_wf wfX_wfY assms by blast + hence "supp t2 \ atom_dom \ \ supp \" using wf_supp + using atom_dom.simps by auto + moreover have "atom x \ atom_dom \" using \atom x \ \\ wfG_atoms_supp_eq wfg fresh_def by blast + ultimately show "atom x \ t2" using fresh_def + by (metis Un_iff contra_subsetD x_not_in_b_set) +qed + +lemma subtype_if_simp: + assumes "wfT P \ GNil (\ z1 : b | CE_val (V_lit l ) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \)" and + "wfT P \ GNil (\ z : b | c \)" and "atom z1 \ c" + shows "P; \ ; GNil \ (\ z1 : b | CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \) \ (\ z : b | c \)" +proof - + obtain x::x where xx: "atom x \ ( P , \ , z1, CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c, GNil)" using obtain_fresh_z by blast + hence xx2: " atom x \ (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c, GNil)" using fresh_prod7 fresh_prod3 by fast + have *:"P; \ ; (x, b, (CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\ GNil \ c[z::=V_var x]\<^sub>v " (is "P; \ ; ?G \ ?c" ) proof - + have "wfC P \ ?G ?c" using wfT_wfC_cons[OF assms(1) assms(2),of x] xx fresh_prod5 fresh_prod3 subst_v_c_def by metis + moreover have "(\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c)" proof(rule allI, rule impI) + fix i + assume as1: "wfI P ?G i \ is_satis_g i ?G" + have "((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" + using assms subst_v_c_def by auto + hence "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var x]\<^sub>v ))" using is_satis_g.simps as1 by presburger + moreover have "is_satis i ((CE_val (V_lit l) == CE_val (V_lit l)))" using is_satis.simps eval_c_eqI[of i "(CE_val (V_lit l))" "eval_l l"] eval_e_uniqueness + eval_e_valI eval_v_litI by metis + ultimately show "is_satis i ?c" using is_satis_mp[of i] by metis + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, GNil, z1 , CE_val (V_lit l) == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z, c )" + unfolding fresh_prod5 \.fresh using xx fresh_prodN x_fresh_b by metis + ultimately show ?thesis using subtype_baseI assms xx xx2 by metis +qed + +lemma subtype_if: + assumes "P; \; \ \ \ z : b | c \ \ \ z' : b | c' \" and + "wfT P \ \ (\ z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \)" and + "wfT P \ \ (\ z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \)" and + "atom z1 \ v" and "atom z \ \" and "atom z1 \ c" and "atom z2 \ c'" and "atom z2 \ v" + shows "P; \ ; \ \ \ z1 : b | CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v \ \ \ z2 : b | CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v \" +proof - + obtain x::x where xx: "atom x \ (P,\,z,c,z', c', z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \)" + using obtain_fresh_z by blast + hence xf: "atom x \ (z, c, z', c', \)" by simp + have xf2: "atom x \ (z1, CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2, CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v , \)" + using xx fresh_prod4 fresh_prodN by metis + + moreover have "P; \ ; (x, b, (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) #\<^sub>\ \ \ (CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v" + (is "P; \ ; ?G \ ?c" ) + proof - + have wbc: "wfC P \ ?G ?c" using assms xx fresh_prod4 fresh_prod2 wfT_wfC_cons assms subst_v_c_def by metis + + moreover have "\i. wfI P ?G i \ is_satis_g i ?G \ is_satis i ?c" proof(rule allI, rule impI) + fix i + assume a1: "wfI P ?G i \ is_satis_g i ?G" + + have *: " is_satis i ((CE_val v == CE_val (V_lit l))) \ is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" + proof + assume a2: "is_satis i ((CE_val v == CE_val (V_lit l)))" + + have "is_satis i ((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v)" + using a1 is_satis_g.simps by simp + moreover have "((CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v ))[z1::=V_var x]\<^sub>v) = (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" + using assms subst_v_c_def by simp + ultimately have "is_satis i (CE_val v == CE_val (V_lit l) IMP ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v))" by argo + + hence "is_satis i ((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v)" using a2 is_satis_mp by auto + moreover have "((c[z::=V_var z1]\<^sub>v )[z1::=V_var x]\<^sub>v) = ((c[z::=V_var x]\<^sub>v ))" using assms by auto + ultimately have "is_satis i ((c[z::=V_var x]\<^sub>v ))" using a2 is_satis.simps by auto + + hence "is_satis_g i ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\ \)" using a1 is_satis_g.simps by meson + moreover have "wfI P ((x,b,(c[z::=V_var x]\<^sub>v )) #\<^sub>\ \) i" proof - + obtain s where "Some s = i x \ wfRCV P s b \ wfI P \ i" using wfI_def a1 by auto + thus ?thesis using wfI_def by auto + qed + ultimately have "is_satis i ((c'[z'::=V_var x]\<^sub>v))" using subtype_valid assms(1) xf valid.simps by simp + + moreover have "(c'[z'::=V_var x]\<^sub>v) = ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" using assms by auto + ultimately show "is_satis i ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v)" by auto + qed + + moreover have "?c = ((CE_val v == CE_val (V_lit l)) IMP ((c'[z'::=V_var z2]\<^sub>v )[z2::=V_var x]\<^sub>v))" + using assms subst_v_c_def by simp + + moreover have "\b1 b2. eval_c i (CE_val v == CE_val (V_lit l) ) b1 \ + eval_c i c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v b2" proof - + + have "wfC P \ ?G (CE_val v == CE_val (V_lit l))" using wbc wfC_elims(7) assms subst_cv.simps subst_v_c_def by fastforce + + moreover have "wfC P \ ?G (c'[z'::=V_var z2]\<^sub>v[z2::=V_var x]\<^sub>v)" proof(rule wfT_wfC_cons) + show \ P; \; \ \\<^sub>w\<^sub>f \ z1 : b | CE_val v == CE_val (V_lit l) IMP (c[z::=V_var z1]\<^sub>v) \ \ using assms subst_v_c_def by auto + have " \ z2 : b | c'[z'::=V_var z2]\<^sub>v \ = \ z' : b | c' \" using assms subst_v_c_def by auto + thus \ P; \; \ \\<^sub>w\<^sub>f \ z2 : b | c'[z'::=V_var z2]\<^sub>v \ \ using assms subtype_elims by metis + show \atom x \ (CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , c'[z'::=V_var z2]\<^sub>v, \)\ using xx fresh_Pair c.fresh by metis + qed + + ultimately show ?thesis using wfI_wfC_eval_c a1 subst_v_c_def by simp + qed + + ultimately show "is_satis i ?c" using is_satis_imp[OF *] by auto + qed + ultimately show ?thesis using valid.simps by simp + qed + moreover have "atom x \ (P, \, \, z1 , CE_val v == CE_val (V_lit l) IMP c[z::=V_var z1]\<^sub>v , z2 , CE_val v == CE_val (V_lit l) IMP c'[z'::=V_var z2]\<^sub>v )" + unfolding fresh_prod5 \.fresh using xx xf2 fresh_prodN x_fresh_b by metis + ultimately show ?thesis using subtype_baseI assms xf2 by metis +qed + +lemma eval_e_concat_eq: + assumes "wfI \ \ i" + shows "\s. eval_e i (CE_val (V_lit (L_bitvec (v1 @ v2))) ) s \ eval_e i (CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e) s" + using eval_e_valI eval_e_concatI eval_v_litI eval_l.simps by metis + +lemma is_satis_eval_e_eq_imp: + assumes "wfI \ \ i" and "eval_e i e1 s" and "eval_e i e2 s" + and "is_satis i (CE_val (V_var x) == e1)" (is "is_satis i ?c1") + shows "is_satis i (CE_val (V_var x) == e2)" +proof - + have *:"eval_c i ?c1 True" using assms is_satis.simps by blast + hence "eval_e i (CE_val (V_var x)) s" using assms is_satis.simps eval_c_elims + by (metis (full_types) eval_e_uniqueness) + thus ?thesis using is_satis.simps eval_c.intros assms by fastforce +qed + +lemma valid_eval_e_eq: + fixes e1::ce and e2::ce + assumes "\\ i. wfI \ \ i \ (\s. eval_e i e1 s \ eval_e i e2 s)" and "\; \; GNil \\<^sub>w\<^sub>f e1 : b " and "\; \; GNil \\<^sub>w\<^sub>f e2 : b" + shows "\; \; (x, b, (CE_val (V_var x) == e1 )) #\<^sub>\ GNil \ (CE_val (V_var x) == e2) " +proof(rule validI) + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) == e2" + proof + have " \ ; \ ; (x, b, TRUE )#\<^sub>\GNil \\<^sub>w\<^sub>f CE_val (V_var x) == e1" using assms wfC_eqI wfE_valI wfV_varI wfX_wfY + by (simp add: fresh_GNil wfC_e_eq) + hence "\ ; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil" using wfG_consI fresh_GNil wfX_wfY assms wfX_wfB by metis + thus "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) : b " using wfCE_valI wfV_varI wfX_wfY + lookup.simps assms wfX_wfY by simp + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f e2 : b " using assms wf_weakening wfX_wfY + by (metis (full_types) \\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil \\<^sub>w\<^sub>f CE_val (V_var x) : b\ empty_iff subsetI toSet.simps(1)) + qed + show " \i. wfI \ ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) i \ is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) \ is_satis i (CE_val (V_var x) == e2 )" + proof(rule,rule) + fix i + assume "wfI \ ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil) i \ is_satis_g i ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ GNil)" + moreover then obtain s where "eval_e i e1 s \ eval_e i e2 s" using assms by auto + ultimately show "is_satis i (CE_val (V_var x) == e2 )" using assms is_satis_eval_e_eq_imp is_satis_g.simps by meson + qed +qed + +lemma subtype_concat: + assumes " \\<^sub>w\<^sub>f \" + shows "\; \; GNil \ \ z : B_bitvec | CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) \ \ + \ z : B_bitvec | CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e \" (is "\; \; GNil \ ?t1 \ ?t2") +proof - + obtain x::x where x: "atom x \ (\, \, GNil, z , CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))), + z , CE_val (V_var z) == CE_concat [V_lit (L_bitvec v1)]\<^sup>c\<^sup>e [V_lit (L_bitvec v2)]\<^sup>c\<^sup>e )" + (is "?xfree" ) + using obtain_fresh by auto + + have wb1: "\; \; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_bitvec (v1 @ v2))) : B_bitvec" using wfX_wfY wfCE_valI wfV_litI assms base_for_lit.simps wfG_nilI by metis + hence wb2: "\; \; GNil \\<^sub>w\<^sub>f CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e : B_bitvec" + using wfCE_concatI wfX_wfY wfV_litI base_for_lit.simps wfCE_valI by metis + + show ?thesis proof + show " \; \; GNil \\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil wb1 wb2 by metis + show " \; \; GNil \\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil wb1 wb2 by metis + show ?xfree using x by auto + show "\; \; (x, B_bitvec, (CE_val (V_var z) == CE_val (V_lit (L_bitvec (v1 @ v2))) )[z::=V_var x]\<^sub>v) #\<^sub>\ + GNil \ (CE_val (V_var z) == CE_concat [(V_lit (L_bitvec v1))]\<^sup>c\<^sup>e [(V_lit (L_bitvec v2))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v " + using valid_eval_e_eq eval_e_concat_eq wb1 wb2 subst_v_c_def by fastforce + qed +qed + +lemma subtype_len: + assumes " \\<^sub>w\<^sub>f \" + shows "\; \; GNil \ \ z' : B_int | CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) \ \ + \ z : B_int | CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e \" (is "\; \; GNil \ ?t1 \ ?t2") +proof - + + have *: "\ \\<^sub>w\<^sub>f [] \ \; \; GNil \\<^sub>w\<^sub>f []\<^sub>\ " using assms wfG_nilI wfD_emptyI wfPhi_emptyI by auto + obtain x::x where x: "atom x \ (\, \, GNil, z' , CE_val (V_var z') == + CE_val (V_lit (L_num (int (length v)))) , z, CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )" + (is "atom x \ ?F") + using obtain_fresh by metis + then show ?thesis proof + have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_val (V_lit (L_num (int (length v)))) : B_int" + using wfCE_valI * wfV_litI base_for_lit.simps + by (metis wfE_valI wfX_wfY) + + thus "\; \; GNil \\<^sub>w\<^sub>f ?t1" using wfT_e_eq fresh_GNil by auto + + have "\ ; \ ; GNil \\<^sub>w\<^sub>f CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e : B_int" + using wfE_valI * wfV_litI base_for_lit.simps wfE_valI wfX_wfY + by (metis wfCE_lenI wfCE_valI) + + thus "\; \; GNil \\<^sub>w\<^sub>f ?t2" using wfT_e_eq fresh_GNil by auto + + show "\; \; (x, B_int, (CE_val (V_var z') == CE_val (V_lit (L_num (int (length v)))) )[z'::=V_var x]\<^sub>v) #\<^sub>\ GNil \ (CE_val (V_var z) == CE_len [(V_lit (L_bitvec v))]\<^sup>c\<^sup>e )[z::=V_var x]\<^sub>v" + (is "\; \; ?G \ ?c" ) using valid_len assms subst_v_c_def by auto + qed +qed + +lemma subtype_base_fresh: + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "\; \; \ \\<^sub>w\<^sub>f \ z : b | c' \ " and + "atom z \ \" and "\; \; (z, b, c) #\<^sub>\ \ \ c'" + shows "\; \; \ \ \ z : b | c \ \ \ z : b | c' \" +proof - + obtain x::x where *:"atom x \ ((\ , \ , z, c, z, c', \) , (\, \, \, \ z : b | c \, \ z : b | c' \))" using obtain_fresh by metis + moreover hence "atom x \ \" using fresh_Pair by auto + moreover hence "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \ c'[z::=V_var x]\<^sub>v" using assms valid_flip_simple * subst_v_c_def by auto + ultimately show ?thesis using subtype_baseI assms \.fresh fresh_Pair by metis +qed + +lemma subtype_bop_arith: + assumes "wfG \ \ \" and "(opp = Plus \ ll = (L_num (n1+n2))) \ (opp = LEq \ ll = ( if n1\n2 then L_true else L_false))" + and "(opp = Plus \ b = B_int) \ (opp = LEq \ b = B_bool)" + shows "\; \; \ \ (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit (ll))) \) \ + \ z : b | C_eq (CE_val (V_var z)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?T1 \ ?T2") +proof - + obtain x::x where xf: "atom x \ (z, CE_val (V_var z) == CE_val (V_lit (ll)) , z, CE_val (V_var z) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e , \)" + using obtain_fresh by blast + + have "\; \; \ \ (\ x : b | C_eq (CE_val (V_var x)) (CE_val (V_lit (ll))) \) \ + \ x : b | C_eq (CE_val (V_var x)) (CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?S1 \ ?S2") + proof(rule subtype_base_fresh) + + show "atom x \ \" using xf fresh_Pair by auto + + show "wfT \ \ \ (\ x : b | CE_val (V_var x) == CE_val (V_lit ll) \)" (is "wfT \ \ ?A ?B") + proof(rule wfT_e_eq) + have " \; \; \ \\<^sub>w\<^sub>f (V_lit ll) : b" using wfV_litI base_for_lit.simps assms by metis + thus " \; \; \ \\<^sub>w\<^sub>f CE_val (V_lit ll) : b" using wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + consider "opp = Plus" | "opp = LEq" using opp.exhaust assms by blast + then show "wfT \ \ \ (\ x : b | CE_val (V_var x) == CE_op opp [(V_lit (L_num n1))]\<^sup>c\<^sup>e [(V_lit (L_num n2))]\<^sup>c\<^sup>e \)" (is "wfT \ \ ?A ?C") + proof(cases) + case 1 + then show "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \" + using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms + by (metis \atom x \ \\ wfT_e_eq) + next + case 2 + then show "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ opp [ [ L_num n1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ L_num n2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ " + using wfCE_valI wfCE_plusI assms wfV_litI base_for_lit.simps assms + + by (metis \atom x \ \\ wfCE_leqI wfT_e_eq) + qed + + show "\; \ ; (x, b, (CE_val (V_var x) == CE_val (V_lit (ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op opp [V_lit (L_num n1)]\<^sup>c\<^sup>e [V_lit (L_num n2)]\<^sup>c\<^sup>e)" (is "\; \; ?G \ ?c") + using valid_arith_bop assms xf by simp + + qed + moreover have "?S1 = ?T1 " using type_l_eq by auto + moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp + by (metis ms_fresh_all(4)) + ultimately show ?thesis by auto + +qed + +lemma subtype_bop_eq: + assumes "wfG \ \ \" and "base_for_lit l1 = base_for_lit l2" + shows "\; \; \ \ (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_val (V_lit (if l1 = l2 then L_true else L_false))) \) \ + \ z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?T1 \ ?T2") +proof - + let ?ll = "if l1 = l2 then L_true else L_false" + obtain x::x where xf: "atom x \ (z, CE_val (V_var z) == CE_val (V_lit (if l1 = l2 then L_true else L_false)) , z, CE_val (V_var z) == CE_op Eq [(V_lit l1)]\<^sup>c\<^sup>e [(V_lit l2)]\<^sup>c\<^sup>e , \, (\, \, \))" + using obtain_fresh by blast + + have "\; \; \ \ (\ x : B_bool | C_eq (CE_val (V_var x)) (CE_val (V_lit (?ll))) \) \ + \ x : B_bool | C_eq (CE_val (V_var x)) (CE_op Eq [(V_lit (l1))]\<^sup>c\<^sup>e [(V_lit (l2))]\<^sup>c\<^sup>e) \" (is "\; \; \ \ ?S1 \ ?S2") + proof(rule subtype_base_fresh) + + show "atom x \ \" using xf fresh_Pair by auto + + show "wfT \ \ \ (\ x : B_bool | CE_val (V_var x) == CE_val (V_lit ?ll) \)" (is "wfT \ \ ?A ?B") + proof(rule wfT_e_eq) + have " \; \; \ \\<^sub>w\<^sub>f (V_lit ?ll) : B_bool" using wfV_litI base_for_lit.simps assms by metis + thus " \; \; \ \\<^sub>w\<^sub>f CE_val (V_lit ?ll) : B_bool" using wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + show " \ ; \ ; \ \\<^sub>w\<^sub>f \ x : B_bool | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e \ " + proof(rule wfT_e_eq) + show "\ ; \ ; \ \\<^sub>w\<^sub>f [ eq [ [ l1 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ l2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e : B_bool" + apply(rule wfCE_eqI, rule wfCE_valI) + apply(rule wfV_litI, simp add: assms) + using wfV_litI assms wfCE_valI by auto + show "atom x \ \" using xf fresh_Pair by auto + qed + + show "\; \ ; (x, B_bool, (CE_val (V_var x) == CE_val (V_lit (?ll)) )) #\<^sub>\ \ + \ (CE_val (V_var x) == CE_op Eq [V_lit (l1)]\<^sup>c\<^sup>e [V_lit (l2)]\<^sup>c\<^sup>e)" (is "\; \; ?G \ ?c") + using valid_eq_bop assms xf by auto + + qed + moreover have "?S1 = ?T1 " using type_l_eq by auto + moreover have "?S2 = ?T2" using type_e_eq ce.fresh v.fresh supp_l_empty fresh_def empty_iff fresh_e_opp + by (metis ms_fresh_all(4)) + ultimately show ?thesis by auto + +qed + +lemma subtype_top: + assumes "wfT \ \ G (\ z : b | c \)" + shows "\; \; G \ (\ z : b | c \) \ (\ z : b | TRUE \)" +proof - + obtain x::x where *: "atom x \ (\, \, G, z , c, z , TRUE)" using obtain_fresh by blast + then show ?thesis proof(rule subtype_baseI) + show \ \; \; G \\<^sub>w\<^sub>f \ z : b | c \ \ using assms by auto + show \ \; \;G \\<^sub>w\<^sub>f \ z : b | TRUE \ \ using wfT_TRUE assms wfX_wfY b_of.simps wfT_wf + by (metis wfX_wfB(8)) + hence "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ G" using wfT_wf_cons3 assms fresh_Pair * subst_v_c_def by auto + thus \\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ G \ (TRUE)[z::=V_var x]\<^sub>v \ using valid_trueI subst_cv.simps subst_v_c_def by metis + qed +qed + +lemma if_simp: + "(if x = x then e1 else e2) = e1" + by auto + +lemma subtype_split: + assumes "split n v (v1,v2)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; GNil \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ [ L_bitvec + v1 ]\<^sup>v , [ L_bitvec + v2 ]\<^sup>v ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | [ [ L_bitvec + v ]\<^sup>v ]\<^sup>c\<^sup>e == [ [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ [ L_num + n ]\<^sup>v ]\<^sup>c\<^sup>e \" + (is "\ ;?B ; GNil \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c1 \ \ \ z : [ B_bitvec , B_bitvec ]\<^sup>b | ?c2 \") +proof - + obtain x::x where xf:"atom x \ (\, ?B, GNil, z, ?c1,z, ?c2 )" using obtain_fresh by auto + then show ?thesis proof(rule subtype_baseI) + show *: \\ ; ?B ; (x, [ B_bitvec , B_bitvec ]\<^sup>b, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ + GNil \ (?c2)[z::=[ x ]\<^sup>v]\<^sub>v \ + unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp + using valid_split[OF assms, of x] by simp + show \ \ ; ?B ; GNil \\<^sub>w\<^sub>f \ z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c1 \ \ using valid_wfT[OF *] xf fresh_prodN by metis + show \ \ ; ?B ; GNil \\<^sub>w\<^sub>f \ z : [ B_bitvec , B_bitvec ]\<^sup>b| ?c2 \ \ using valid_wfT[OF *] xf fresh_prodN by metis + qed +qed + +lemma subtype_range: + fixes n::int and \::\ + assumes "0 \ n \ n \ int (length v)" and "\ ; {||} \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; \ \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ + \ z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) AND ( + [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) \" + (is "\ ; ?B ; \ \ \ z : B_int | ?c1 \ \ \ z : B_int | ?c2 AND ?c3 \") +proof - + obtain x::x where *:\atom x \ (\, ?B, \, z, ?c1 , z, ?c2 AND ?c3)\ using obtain_fresh by auto + moreover have **:\\ ; ?B ; (x, B_int, (?c1)[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \ \ (?c2 AND ?c3)[z::=[ x ]\<^sup>v]\<^sub>v\ + unfolding subst_v_c_def subst_cv.simps subst_cev.simps subst_vv.simps if_simp using valid_range_length[OF assms(1)] assms fresh_prodN * by simp + + moreover hence \ \ ; ?B ; \ \\<^sub>w\<^sub>f \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ using + valid_wfT * fresh_prodN by metis + moreover have \ \ ; ?B ; \ \\<^sub>w\<^sub>f \ z : B_int | ?c2 AND ?c3 \ \ + using valid_wfT[OF **] * fresh_prodN by metis + ultimately show ?thesis using subtype_baseI by auto +qed + +lemma check_num_range: + assumes "0 \ n \ n \ int (length v)" and "\\<^sub>w\<^sub>f \" + shows "\ ; {||} ; GNil \ ([ L_num n ]\<^sup>v) \ \ z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) AND + [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \" + using assms subtype_range check_v.intros infer_v_litI wfG_nilI + by (meson infer_natI) + +section \Literals\ + +nominal_function type_for_lit :: "l \ \" where + "type_for_lit (L_true) = (\ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_true]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_false) = (\ z : B_bool | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit L_false]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_num n) = (\ z : B_int | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_num n)]\<^sup>c\<^sup>e \)" +| "type_for_lit (L_unit) = (\ z : B_unit | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_unit )]\<^sup>c\<^sup>e \)" (* could have done { z : unit | True } but want the uniformity *) +| "type_for_lit (L_bitvec v) = (\ z : B_bitvec | [[z]\<^sup>v]\<^sup>c\<^sup>e == [V_lit (L_bitvec v)]\<^sup>c\<^sup>e \)" + by (auto simp: eqvt_def type_for_lit_graph_aux_def, metis l.strong_exhaust,(simp add: permute_int_def flip_bitvec0)+ ) +nominal_termination (eqvt) by lexicographic_order + + +nominal_function type_for_var :: "\ \ \ \ x \ \" where + "type_for_var G \ x = (case lookup G x of + None \ \ + | Some (b,c) \ (\ x : b | c \)) " + apply auto unfolding eqvt_def apply(rule allI) unfolding type_for_var_graph_aux_def eqvt_def by simp +nominal_termination (eqvt) by lexicographic_order + +lemma infer_l_form: + fixes l::l and tm::"'a::fs" + assumes "\ l \ \" + shows "\z b. \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \) \ atom z \ tm" +proof - + obtain z' and b where t:"\ = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \)" using infer_l_elims assms using infer_l.simps type_for_lit.simps + type_for_lit.cases by blast + obtain z::x where zf: "atom z \ tm" using obtain_fresh by metis + have "\ = \ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \" using type_e_eq ce.fresh v.fresh l.fresh + by (metis t type_l_eq) + thus ?thesis using zf by auto +qed + +lemma infer_l_form3: + fixes l::l + assumes "\ l \ \" + shows "\z. \ = (\ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" + using infer_l_elims using assms using infer_l.simps type_for_lit.simps base_for_lit.simps by auto + +lemma infer_l_form4[simp]: + fixes \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f \ " + shows "\z. \ l \ (\ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" + using assms infer_l_form2 infer_l_form3 by metis + +lemma infer_v_unit_form: + fixes v::v + assumes "P ; \ ; \ \ v \ (\ z1 : B_unit | c1 \)" and "supp v = {}" + shows "v = V_lit L_unit" + using assms proof(nominal_induct \ v "\ z1 : B_unit | c1 \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ c x z) + then show ?case using supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \\ l \ \ z1 : B_unit | c1 \\ show ?case by(nominal_induct "\ z1 : B_unit | c1 \" rule: infer_l.strong_induct,auto) +qed + +lemma base_for_lit_wf: + assumes "\\<^sub>w\<^sub>f \" + shows "\ ; \ \\<^sub>w\<^sub>f base_for_lit l" + using base_for_lit.simps using wfV_elims wf_intros assms l.exhaust by metis + +lemma infer_l_t_wf: + fixes \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f \ \ atom z \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \ z : base_for_lit l | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \" +proof + show "atom z \ (\, \, \)" using wfG_fresh_x assms by auto + show "\ ; \ \\<^sub>w\<^sub>f base_for_lit l" using base_for_lit_wf assms wfX_wfY by metis + thus "\; \; (z, base_for_lit l, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var z) == CE_val (V_lit l)" using wfC_v_eq wfV_litI assms wfX_wfY by metis +qed + +lemma infer_l_wf: + fixes l::l and \::\ and \::\ and \::\ + assumes "\ l \ \" and "\ ; \ \\<^sub>w\<^sub>f \" + shows "\\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f \" +proof - + show *:"\ ; \ \\<^sub>w\<^sub>f \" using assms infer_l_elims by auto + thus "\\<^sub>w\<^sub>f \" using wfX_wfY by auto + show *:"\ ; \ ; \ \\<^sub>w\<^sub>f \" using infer_l_t_wf assms infer_l_form3 * + by (metis \\\<^sub>w\<^sub>f \\ fresh_GNil wfG_nilI wfT_weakening_nil) +qed + +lemma infer_l_uniqueness: + fixes l::l + assumes "\ l \ \" and "\ l \ \'" + shows "\ = \'" + using assms +proof - + obtain z and b where zt: "\ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_lit l)) \)" using infer_l_form assms by blast + obtain z' and b where z't: "\' = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_lit l)) \)" using infer_l_form assms by blast + thus ?thesis using type_l_eq zt z't assms infer_l.simps infer_l_elims l.distinct + by (metis infer_l_form3) +qed + +section \Values\ + +lemma type_v_eq: + assumes "\ z1 : b1 | c1 \ = \ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \" and "atom z \ x" + shows "b = b1" and "c1 = C_eq (CE_val (V_var z1)) (CE_val (V_var x))" + using assms by (auto,metis Abs1_eq_iff \.eq_iff assms c.fresh ce.fresh type_e_eq v.fresh) + +lemma infer_var2 [elim]: + assumes "P; \ ; G \ V_var x \ \" + shows "\b c. Some (b,c) = lookup G x" + using assms infer_v_elims lookup_iff by (metis (no_types, lifting)) + +lemma infer_var3 [elim]: + assumes "\; \; \ \ V_var x \ \" + shows "\z b c. Some (b,c) = lookup \ x \ \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var x)) \) \ atom z \ x \ atom z \ (\, \, \)" + using infer_v_elims(1)[OF assms(1)] by metis + +lemma infer_bool_options2: + fixes v::v + assumes "\; \; \ \ v \ \ z : b | c \" and "supp v = {} \ b = B_bool" + shows "v = V_lit L_true \ (v = (V_lit L_false))" + using assms +proof(nominal_induct "\ z : b | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \ \ l \ \ z : b | c \\ show ?case proof(nominal_induct "\ z : b | c \" rule: infer_l.strong_induct) + case (infer_trueI z) + then show ?case by auto + next + case (infer_falseI z) + then show ?case by auto + next + case (infer_natI n z) + then show ?case using infer_v_litI by simp + next + case (infer_unitI z) + then show ?case using infer_v_litI by simp + next + case (infer_bitvecI bv z) + then show ?case using infer_v_litI by simp + qed +qed(auto+) + +lemma infer_bool_options: + fixes v::v + assumes "\; \; \ \ v \ \ z : B_bool | c \" and "supp v = {}" + shows "v = V_lit L_true \ (v = (V_lit L_false))" + using infer_bool_options2 assms by blast + +lemma infer_int2: + fixes v::v + assumes "\; \; \ \ v \ \ z : b | c \" + shows "supp v = {} \ b = B_int \ (\n. v = V_lit (L_num n))" + using assms +proof(nominal_induct "\ z : b | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ \ l) + from \ \ l \ \ z : b | c \\ show ?case proof(nominal_induct "\ z : b | c \" rule: infer_l.strong_induct) + case (infer_trueI z) + then show ?case by auto + next + case (infer_falseI z) + then show ?case by auto + next + case (infer_natI n z) + then show ?case using infer_v_litI by simp + next + case (infer_unitI z) + then show ?case using infer_v_litI by simp + next + case (infer_bitvecI bv z) + then show ?case using infer_v_litI by simp + qed +qed(auto+) + +lemma infer_bitvec: + fixes \::\ and v::v + assumes "\; \; \ \ v \ \ z' : B_bitvec | c' \" and "supp v = {}" + shows "\bv. v = V_lit (L_bitvec bv)" + using assms proof(nominal_induct v rule: v.strong_induct) + case (V_lit l) + then show ?case by(nominal_induct l rule: l.strong_induct,force+) +next + case (V_consp s dc b v) + then show ?case using infer_v_elims(7)[OF V_consp(2)] \.eq_iff by auto +next + case (V_var x) + then show ?case using supp_at_base by auto +qed(force+) + +lemma infer_int: + assumes "infer_v \ \ \ v (\ z : B_int | c \)" and "supp v= {}" + shows "\n. V_lit (L_num n) = v" + using assms infer_int2 by (metis (no_types, lifting)) + +lemma infer_lit: + assumes "infer_v \ \ \ v (\ z : b | c \)" and "supp v= {}" and "b \ { B_bool , B_int , B_unit }" + shows "\l. V_lit l = v" + using assms proof(nominal_induct v rule: v.strong_induct) + case (V_lit x) + then show ?case by (simp add: supp_at_base) +next + case (V_var x) + then show ?case + by (simp add: supp_at_base) +next + case (V_pair x1a x2a) + then show ?case using supp_at_base by auto +next + case (V_cons x1a x2a x3) + then show ?case using supp_at_base by auto +next + case (V_consp x1a x2a x3 x4) + then show ?case using supp_at_base by auto +qed + +lemma infer_v_form[simp]: + fixes v::v + assumes "\; \; \ \ v \ \" + shows "\z b. \ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val v)\) \ atom z \ v \ atom z \ (\, \, \)" + using assms +proof(nominal_induct rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b c x z) + then show ?case by force +next + case (infer_v_litI \ \ \ l \) + then obtain z and b where "\ = \ z : b | CE_val (V_var z) == CE_val (V_lit l) \ \atom z \ (\, \, \) " + using infer_l_form by metis + moreover hence "atom z \ (V_lit l)" using supp_l_empty v.fresh(1) fresh_prod2 fresh_def by blast + ultimately show ?case by metis +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case by force +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + moreover hence "atom z \ (V_cons s dc v)" using + Un_commute b.supp(3) fresh_def sup_bot.right_neutral supp_b_empty v.supp(4) pure_supp by metis + ultimately show ?case using fresh_prodN by metis +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + moreover hence "atom z \ (V_consp s dc b v)" unfolding v.fresh using pure_fresh fresh_prodN * by metis + ultimately show ?case using fresh_prodN by metis +qed + +lemma infer_v_form2: + fixes v::v + assumes "\; \; \ \ v \ (\ z : b | c \)" and "atom z \ v" + shows "c = C_eq (CE_val (V_var z)) (CE_val v)" + using assms +proof - + obtain z' and b' where "(\ z : b | c \) = (\ z' : b' | CE_val (V_var z') == CE_val v \) \ atom z' \ v" + using infer_v_form assms by meson + thus ?thesis using Abs1_eq_iff(3) \.eq_iff type_e_eq + by (metis assms(2) ce.fresh(1)) +qed + +lemma infer_v_form3: + fixes v::v + assumes "\; \; \ \ v \ \" and "atom z \ (v,\)" + shows "\; \; \ \ v \ \ z : b_of \ | C_eq (CE_val (V_var z)) (CE_val v)\" +proof - + obtain z' and b' where "\ = \ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\ \ atom z' \ v \ atom z' \ (\, \, \)" + using infer_v_form assms by metis + moreover hence "\ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\ = \ z : b' | C_eq (CE_val (V_var z)) (CE_val v)\" + using assms type_e_eq fresh_Pair ce.fresh by auto + ultimately show ?thesis using b_of.simps assms by auto +qed + +lemma infer_v_form4: + fixes v::v + assumes "\; \; \ \ v \ \" and "atom z \ (v,\)" and "b = b_of \" + shows "\; \; \ \ v \ \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\" + using assms infer_v_form3 by simp + +lemma infer_v_v_wf: + fixes v::v + shows "\; \ ; G \ v \ \ \ \; \; G \\<^sub>w\<^sub>f v : (b_of \)" +proof(induct rule: infer_v.induct) + case (infer_v_varI \ \ \ b c x z) + then show ?case using wfC_elims wf_intros by auto +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case using wfC_elims wf_intros by auto +next + case (infer_v_litI \ \ \ l \) + hence "b_of \ = base_for_lit l" using infer_l_form3 b_of.simps by metis + then show ?case using wfV_litI infer_l_wf infer_v_litI wfG_b_weakening + by (metis fempty_fsubsetI) +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + then show ?case using wfC_elims wf_intros + by (metis (no_types, lifting) b_of.simps has_fresh_z2 subtype_eq_base2) +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + obtain z1 b1 c1 where t:"tc = \ z1 : b1 | c1 \" using obtain_fresh_z by metis + show ?case unfolding b_of.simps proof(rule wfV_conspI) + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, \ z1 : b1 | c1 \ ) \ set dclist\ using infer_v_conspI t by auto + show \ \ ; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + show \atom bv \ (\, \, \, b, v)\ using infer_v_conspI by auto + have " b1[bv::=b]\<^sub>b\<^sub>b = b_of tv" using subtype_eq_base2[OF infer_v_conspI(5)] b_of.simps t subst_tb.simps by auto + thus \ \; \; \ \\<^sub>w\<^sub>f v : b1[bv::=b]\<^sub>b\<^sub>b \ using infer_v_conspI by auto + qed +qed + +lemma infer_v_t_form_wf: + assumes " wfB \ \ b" and "wfV \ \ \ v b" and "atom z \ \" + shows "wfT \ \ \ \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\" + using wfT_v_eq assms by auto + +lemma infer_v_t_wf: + fixes v::v + assumes "\; \; G \ v \ \" + shows "wfT \ \ G \ \ wfB \ \ (b_of \) " +proof - + obtain z and b where "\ = \ z : b | CE_val (V_var z) == CE_val v \ \ atom z \ v \ atom z \ (\, \, G)" using infer_v_form assms by metis + moreover have "wfB \ \ b" using infer_v_v_wf b_of.simps wfX_wfB(1) assms + using calculation by fastforce + ultimately show "wfT \ \ G \ \ wfB \ \ (b_of \)" using infer_v_v_wf infer_v_t_form_wf assms by fastforce +qed + +lemma infer_v_wf: + fixes v::v + assumes "\; \; G \ v \ \" + shows "\; \; G \\<^sub>w\<^sub>f v : (b_of \)" and "wfT \ \ G \" and "wfTh \" and "wfG \ \ G" +proof - + show "\; \; G \\<^sub>w\<^sub>f v : b_of \ " using infer_v_v_wf assms by auto + show "\; \; G \\<^sub>w\<^sub>f \" using infer_v_t_wf assms by auto + thus "\ ; \ \\<^sub>w\<^sub>f G" using wfX_wfY by auto + thus "\\<^sub>w\<^sub>f \" using wfX_wfY by auto +qed + +lemma check_bool_options: + assumes "\; \; \ \ v \ \ z : B_bool | TRUE \" and "supp v = {}" + shows "v = V_lit L_true \ v = V_lit L_false" +proof - + obtain t1 where " \; \; \ \ t1 \ \ z : B_bool | TRUE \ \ \; \; \ \ v \ t1" using check_v_elims + using assms by blast + thus ?thesis using infer_bool_options assms + by (metis \.exhaust b_of.simps subtype_eq_base2) +qed + +lemma check_v_wf: + fixes v::v and \::\ and \::\ + assumes "\; \; \ \ v \ \" + shows " \ ; \ \\<^sub>w\<^sub>f \" and "\; \;\ \\<^sub>w\<^sub>f v : b_of \" and "\; \;\ \\<^sub>w\<^sub>f \" +proof - + obtain \' where *: "\; \; \ \ \' \ \ \ \; \; \ \ v \ \'" using check_v_elims assms by auto + thus "\ ; \ \\<^sub>w\<^sub>f \ " and "\; \;\ \\<^sub>w\<^sub>f v : b_of \" and "\; \; \ \\<^sub>w\<^sub>f \" + using infer_v_wf infer_v_v_wf subtype_eq_base2 * subtype_wf by metis+ +qed + +lemma infer_v_form_fresh: + fixes v::v and t::"'a::fs" + assumes "\; \; \ \ v \ \" + shows "\z b. \ = \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\ \ atom z \ (t,v)" +proof - + obtain z' and b' where "\ = \ z' : b' | C_eq (CE_val (V_var z')) (CE_val v)\" using infer_v_form assms by blast + moreover then obtain z and b and c where "\ = \ z : b | c \ \ atom z \ (t,v)" using obtain_fresh_z by metis + ultimately have "\ = \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\ \ atom z \ (t,v) " + using assms infer_v_form2 by auto + thus ?thesis by blast +qed + +text \ More generally, if support of a term is empty then any G will do \ + +lemma infer_v_form_consp: + assumes "\; \; \ \ V_consp s dc b v \ \" + shows "b_of \ = B_app s b" + using assms proof(nominal_induct "V_consp s dc b v" \ rule: infer_v.strong_induct) + case (infer_v_conspI bv dclist \ tc \ \ tv z) + then show ?case using b_of.simps by metis +qed + +lemma lookup_in_rig_b: + assumes "Some (b2, c2) = lookup (\[x\c']) x'" and + "Some (b1, c1) = lookup \ x'" + shows "b1 = b2" + using assms lookup_in_rig[OF assms(2)] + by (metis option.inject prod.inject) + +lemma infer_v_uniqueness_rig: + fixes x::x and c::c + assumes "infer_v P B G v \" and "infer_v P B (replace_in_g G x c') v \'" + shows "\ = \'" + using assms +proof(nominal_induct "v" arbitrary: \' \ rule: v.strong_induct) + case (V_lit l) + hence "infer_l l \" and "infer_l l \'" using assms(1) infer_v_elims(2) by auto + then show ?case using infer_l_uniqueness by presburger +next + case (V_var y) + + obtain b and c where bc: "Some (b,c) = lookup G y" + using assms(1) infer_v_elims(2) using V_var.prems(1) lookup_iff by force + then obtain c'' where bc':"Some (b,c'') = lookup (replace_in_g G x c') y" + using lookup_in_rig by blast + + obtain z where "\ = (\ z : b | C_eq (CE_val (V_var z)) (CE_val (V_var y)) \)" using infer_v_elims(1)[of P B G y \] V_var + bc option.inject prod.inject lookup_in_g by metis + moreover obtain z' where "\' = (\ z' : b | C_eq (CE_val (V_var z')) (CE_val (V_var y)) \)" using infer_v_elims(1)[of P B _ y \'] V_var + option.inject prod.inject lookup_in_rig by (metis bc') + ultimately show ?case using type_e_eq + by (metis V_var.prems(1) V_var.prems(2) \.eq_iff ce.fresh(1) finite.emptyI fresh_atom_at_base + fresh_finite_insert infer_v_elims(1) v.fresh(2)) +next + case (V_pair v1 v2) + obtain z and z1 and z2 and t1 and t2 and c1 and c2 where + t1: "\ = (\ z : [ b_of t1 , b_of t2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v1 v2) \) \ + atom z \ (v1, v2) \ P ; B ; G \ v1 \ t1 \ P ; B ; G \ v2 \ t2" + using infer_v_elims(3)[OF V_pair(3)] by metis + moreover obtain z' and z1' and z2' and t1' and t2' and c1' and c2' where + t2: "\' = (\ z' : [ b_of t1' , b_of t2' ]\<^sup>b | CE_val (V_var z') == CE_val (V_pair v1 v2) \) \ + atom z' \ (v1, v2) \ P ; B ; (replace_in_g G x c') \ v1 \ t1' \ + P ; B ; (replace_in_g G x c') \ v2 \ t2'" + using infer_v_elims(3)[OF V_pair(4)] by metis + ultimately have "t1 = t1' \ t2 = t2'" using V_pair.hyps(1) V_pair.hyps(2) \.eq_iff by blast + then show ?case using t1 t2 by simp +next + case (V_cons s dc v) + obtain x and z and tc and dclist where t1: "\ = (\ z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \) \ + AF_typedef s dclist \ set P \ + (dc, tc) \ set dclist \ atom z \ v" + using infer_v_elims(4)[OF V_cons(2)] by metis + moreover obtain x' and z' and tc' and dclist' where t2: "\' = (\ z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \) + \ AF_typedef s dclist' \ set P \ (dc, tc') \ set dclist' \ atom z' \ v" + using infer_v_elims(4)[OF V_cons(3)] by metis + moreover have a: "AF_typedef s dclist' \ set P" and b:"(dc,tc') \ set dclist'" and c:"AF_typedef s dclist \ set P" and + d:"(dc, tc) \ set dclist" using t1 t2 by auto + ultimately have "tc = tc'" using wfTh_dc_t_unique2 infer_v_wf(3)[OF V_cons(2)] by metis + + moreover have "atom z \ CE_val (V_cons s dc v) \ atom z' \ CE_val (V_cons s dc v)" + using e.fresh(1) v.fresh(4) t1 t2 pure_fresh by auto + ultimately have "(\ z : B_id s | CE_val (V_var z) == CE_val (V_cons s dc v) \) = (\ z' : B_id s | CE_val (V_var z') == CE_val (V_cons s dc v) \)" + using type_e_eq by metis + thus ?case using t1 t2 by simp +next + case (V_consp s dc b v) + from V_consp(2) V_consp show ?case proof(nominal_induct "V_consp s dc b v" \ arbitrary: v rule:infer_v.strong_induct) + + case (infer_v_conspI bv dclist \ tc \ \ v tv z) + + obtain z3 and b3 where *:"\' = \ z3 : b3 | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \ \ atom z3 \ V_consp s dc b v" + using infer_v_form[OF \\; \; \[x\c'] \ V_consp s dc b v \ \'\ ] by metis + moreover then have "b3 = B_app s b" using infer_v_form_consp b_of.simps * infer_v_conspI by metis + + moreover have "\ z3 : B_app s b | [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \ = \ z : B_app s b | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ V_consp s dc b v ]\<^sup>c\<^sup>e \" + proof - + have "atom z3 \ [V_consp s dc b v]\<^sup>c\<^sup>e" using * ce.fresh by auto + moreover have "atom z \ [V_consp s dc b v]\<^sup>c\<^sup>e" using * infer_v_conspI ce.fresh v.fresh pure_fresh by metis + ultimately show ?thesis using type_e_eq infer_v_conspI v.fresh ce.fresh by metis + qed + ultimately show ?case using * by auto + qed +qed + +lemma infer_v_uniqueness: + assumes "infer_v P B G v \" and "infer_v P B G v \'" + shows "\ = \'" +proof - + obtain x::x where "atom x \ G" using obtain_fresh by metis + hence "G [ x \ C_true ] = G" using replace_in_g_forget assms infer_v_wf by fast + thus ?thesis using infer_v_uniqueness_rig assms by metis +qed + +lemma infer_v_tid_form: + fixes v::v + assumes "\ ; B ; \ \ v \ \ z : B_id tid | c \" and "AF_typedef tid dclist \ set \" and "supp v = {}" + shows "\dc v' t. v = V_cons tid dc v' \ (dc , t ) \ set dclist" + using assms proof(nominal_induct v "\ z : B_id tid | c \" rule: infer_v.strong_induct) + case (infer_v_varI \ \ c x z) + then show ?case using v.supp supp_at_base by auto +next + case (infer_v_litI \ \ l) + then show ?case by auto +next + case (infer_v_consI dclist1 \ dc tc \ \ v tv z) + hence "supp v = {}" using v.supp by simp + then obtain dca and v' where *:"V_cons tid dc v = V_cons tid dca v'" using infer_v_consI by auto + hence "dca = dc" using v.eq_iff(4) by auto + hence "V_cons tid dc v = V_cons tid dca v' \ (dca, tc) \ set dclist1" using infer_v_consI * by auto + moreover have "dclist = dclist1" using wfTh_dclist_unique infer_v_consI wfX_wfY \dca=dc\ + proof - + show ?thesis + by (meson \AF_typedef tid dclist1 \ set \\ \\; \; \ \ v \ tv\ infer_v_consI.prems infer_v_wf(4) wfTh_dclist_unique wfX_wfY) + qed + ultimately show ?case by auto +qed + +lemma check_v_tid_form: + assumes "\ ; B ; \ \ v \ \ z : B_id tid | TRUE \" and "AF_typedef tid dclist \ set \" and "supp v = {}" + shows "\dc v' t. v = V_cons tid dc v' \ (dc , t ) \ set dclist" + using assms proof(nominal_induct v "\ z : B_id tid | TRUE \" rule: check_v.strong_induct) + case (check_v_subtypeI \ \ \ \1 v) + then obtain z and c where "\1 = \ z : B_id tid | c \" using subtype_eq_base2 b_of.simps + by (metis obtain_fresh_z2) + then show ?case using infer_v_tid_form check_v_subtypeI by simp +qed + +lemma check_v_num_leq: + fixes n::int and \::\ + assumes "0 \ n \ n \ int (length v)" and " \\<^sub>w\<^sub>f \ " and "\ ; {||} \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; \ \ [ L_num n ]\<^sup>v \ \ z : B_int | ([ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e) + AND ([ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ [ L_bitvec v ]\<^sup>v ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e ) \" +proof - + have "\ ; {||} ; \ \ [ L_num n ]\<^sup>v \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \" + using infer_v_litI infer_natI wfG_nilI assms by auto + thus ?thesis using subtype_range[OF assms(1) ] assms check_v_subtypeI by metis +qed + +lemma check_int: + assumes "check_v \ \ \ v (\ z : B_int | c \)" and "supp v = {}" + shows "\n. V_lit (L_num n) = v" + using assms infer_int check_v_elims by (metis b_of.simps infer_v_form subtype_eq_base2) + +definition sble :: "\ \ \ \ bool" where + "sble \ \ = (\i. i \ \ \ \ ; \ \ i)" + +lemma check_v_range: + assumes "\ ; B ; \ \ v2 \ \ z : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e AND + [ leq [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \ " + (is "\ ; ?B ; \ \ v2 \ \ z : B_int | ?c1 \") + and "v1 = V_lit (L_bitvec bv) \ v2 = V_lit (L_num n) " and "atom z \ \" and "sble \ \" + shows "0 \ n \ n \ int (length bv)" +proof - + have "\ ; ?B ; \ \ \ z : B_int | [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e \ \ \ z : B_int | ?c1 \" + using check_v_elims assms + by (metis infer_l_uniqueness infer_natI infer_v_elims(2)) + moreover have "atom z \ \" using fresh_GNil assms by simp + ultimately have "\ ; ?B ; ((z, B_int, [ [ z ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ L_num n ]\<^sup>v ]\<^sup>c\<^sup>e ) #\<^sub>\ \) \ ?c1" + using subtype_valid_simple by auto + thus ?thesis using assms valid_range_length_inv check_v_wf wfX_wfY sble_def by metis +qed + +section \Expressions\ + +lemma infer_e_plus[elim]: + fixes v1::v and v2::v + assumes "\ ; \ ; \ ; \ ; \ \ AE_op Plus v1 v2 \ \" + shows "\z . (\ z : B_int | C_eq (CE_val (V_var z)) (CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims assms by metis + +lemma infer_e_leq[elim]: + assumes "\ ; \ ; \ ; \ ; \ \ AE_op LEq v1 v2 \ \" + shows "\z . (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims assms by metis + +lemma infer_e_eq[elim]: + assumes "\ ; \ ; \ ; \ ; \ \ AE_op Eq v1 v2 \ \" + shows "\z . (\ z : B_bool | C_eq (CE_val (V_var z)) (CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e) \ = \)" + using infer_e_elims(25)[OF assms] by metis + +lemma infer_e_e_wf: + fixes e::e + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b_of \" + using assms proof(nominal_induct \ avoiding: \ rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_v_wf wf_intros by metis +next + case (infer_e_plusI \ \ \ \' \ v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_leqI \ \ \ \' v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_eqI \ \ \ \' v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_appI \ \ \ \ \ f x b c \' s' v \'') + have "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_app f v : b_of \' " proof + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_appI by auto + show \Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f\ using infer_e_appI by auto + show "\; \;\ \\<^sub>w\<^sub>f v : b" using infer_e_appI check_v_wf b_of.simps by metis + qed + moreover have "b_of \' = b_of (\'[x::=v]\<^sub>v)" using subst_tbase_eq subst_v_\_def by auto + ultimately show ?case using infer_e_appI subst_v_c_def subst_b_\_def by auto +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \'' s' v \') + + have "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_appP f b' v : (b_of \'')[bv::=b']\<^sub>b " proof + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \'' s'))) = lookup_fun \ f\ using * infer_e_appPI by metis + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI by auto + show "\; \;\ \\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b)" using infer_e_appPI check_v_wf b_of.simps subst_b_b_def by metis + have "atom bv \ (b_of \'')[bv::=b']\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def infer_e_appPI by metis + thus "atom bv \ (\, \, \, \, \, b', v, (b_of \'')[bv::=b']\<^sub>b)" using infer_e_appPI fresh_prodN subst_b_b_def by metis + qed + moreover have "b_of \' = (b_of \'')[bv::=b']\<^sub>b" + using \\''[bv::=b']\<^sub>b[x::=v]\<^sub>v = \'\ b_of_subst_bb_commute subst_tbase_eq subst_b_b_def subst_v_\_def subst_b_\_def by auto + ultimately show ?case using infer_e_appI by auto +next + case (infer_e_fstI \ \ \ \' \ v z' b1 b2 c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_sndI \ \ \ \' \ v z' b1 b2 c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_lenI \ \ \ \' \ v z' c z) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_concatI \ \ \ \' \ v1 z1 c1 v2 z2 c2 z3) + then show ?case using b_of.simps infer_v_v_wf wf_intros by metis +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + have " \ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec" + proof + show "\ \\<^sub>w\<^sub>f \" using infer_e_splitI by auto + show "\; \; \ \\<^sub>w\<^sub>f \" using infer_e_splitI by auto + show "\; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec" using infer_e_splitI b_of.simps infer_v_wf by metis + show "\; \; \ \\<^sub>w\<^sub>f v2 : B_int" using infer_e_splitI b_of.simps check_v_wf by metis + qed + then show ?case using b_of.simps by auto +qed + +lemma infer_e_t_wf: + fixes e::e and \::\ and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\; \;\ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \" + using assms proof(induct rule: infer_e.induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_t_wf by auto +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence "\; \; \ \\<^sub>w\<^sub>f CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_int" using wfCE_plusI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_plusI by auto +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence " \; \; \ \\<^sub>w\<^sub>f CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_leqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_leqI by auto +next + case (infer_e_eqI \ \ \ \ \ v1 z1 b c1 v2 z2 c2 z3) + hence " \; \; \ \\<^sub>w\<^sub>f CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e : B_bool" using wfCE_eqI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_eqI by auto +next + case (infer_e_appI \ \ \ \ \ f x b c \ s' v \') + show ?case proof + show "\ \\<^sub>w\<^sub>f \" using infer_e_appI by auto + show "\; \; \ \\<^sub>w\<^sub>f \'" proof - + have *:" \; \; \ \\<^sub>w\<^sub>f v : b " using infer_e_appI check_v_wf(2) b_of.simps by metis + moreover have *:"\; \; (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \" proof(rule wf_weakening1(4)) + show \ \; \; (x,b,c)#\<^sub>\GNil \\<^sub>w\<^sub>f \ \ using wfPhi_f_simple_wfT wfD_wf infer_e_appI wb_b_weakening by fastforce + have "\ ; \ ; \ \\<^sub>w\<^sub>f \ x : b | c \" using infer_e_appI check_v_wf(3) by auto + thus \ \ ; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \ \ using infer_e_appI + wfT_wfC[THEN wfG_consI[rotated 3] ] * wfT_wf_cons fresh_prodN by simp + show \toSet ((x,b,c)#\<^sub>\GNil) \ toSet ((x, b, c) #\<^sub>\ \)\ using toSet.simps by auto + qed + moreover have "((x, b, c) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + + ultimately show ?thesis using infer_e_appI wf_subst1(4)[OF *, of GNil x b c \ v] subst_v_\_def by auto + qed + qed +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v \) + + have "\; \; ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v \\<^sub>w\<^sub>f (\'[bv::=b']\<^sub>b)[x::=v]\<^sub>\\<^sub>v" + proof(rule wf_subst(4)) + show \ \; \; (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ \\<^sub>w\<^sub>f \'[bv::=b']\<^sub>b \ + proof(rule wf_weakening1(4)) + have \ \ ; {|bv|} ; (x,b,c)#\<^sub>\GNil \\<^sub>w\<^sub>f \' \ using wfPhi_f_poly_wfT infer_e_appI infer_e_appPI by simp + thus \ \; \; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\GNil \\<^sub>w\<^sub>f \'[bv::=b']\<^sub>b \ + using wfT_subst_wfT infer_e_appPI wb_b_weakening subst_b_\_def subst_v_\_def by presburger + have "\; \; \ \\<^sub>w\<^sub>f \ x : b[bv::=b']\<^sub>b\<^sub>b | c[bv::=b']\<^sub>c\<^sub>b \" + using infer_e_appPI check_v_wf(3) subst_b_b_def subst_b_c_def by metis + thus \ \ ; \ \\<^sub>w\<^sub>f (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ \ + using infer_e_appPI wfT_wfC[THEN wfG_consI[rotated 3] ] * wfX_wfY wfT_wf_cons wb_b_weakening by metis + show \toSet ((x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b)#\<^sub>\GNil) \ toSet ((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)\ using toSet.simps by auto + qed + show \(x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \ = GNil @ (x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \\ using append_g.simps by auto + show \ \; \; \ \\<^sub>w\<^sub>f v :b[bv::=b']\<^sub>b\<^sub>b \ using infer_e_appPI check_v_wf(2) b_of.simps subst_b_b_def by metis + qed + moreover have "((x, b[bv::=b']\<^sub>b\<^sub>b, c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + ultimately show ?case using infer_e_appPI subst_v_\_def by simp +next + case (infer_e_fstI \ \ \ \ \ v z' b1 b2 c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_fst [v]\<^sup>c\<^sup>e: b1" using wfCE_fstI wfD_emptyI wfPhi_emptyI infer_v_v_wf + b_of.simps using wfCE_valI by fastforce + then show ?case using wfT_e_eq infer_e_fstI by auto +next + case (infer_e_sndI \ \ \ \ \ v z' b1 b2 c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_snd [v]\<^sup>c\<^sup>e: b2" using wfCE_sndI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_sndI by auto +next + case (infer_e_lenI \ \ \ \ \ v z' c z) + hence "\; \; \ \\<^sub>w\<^sub>f CE_len [v]\<^sup>c\<^sup>e: B_int" using wfCE_lenI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_lenI by auto +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using wfD_wfT by blast +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + hence "\; \; \ \\<^sub>w\<^sub>f CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e: B_bitvec" using wfCE_concatI wfD_emptyI wfPhi_emptyI infer_v_v_wf wfCE_valI + by (metis b_of.simps infer_v_wf) + then show ?case using wfT_e_eq infer_e_concatI by auto +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + + hence wfg: " \ ; \ \\<^sub>w\<^sub>f (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \" + using infer_v_wf wfG_cons2I wfB_pairI wfB_bitvecI by simp + have wfz: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e : [ B_bitvec , B_bitvec ]\<^sup>b " + apply(rule wfCE_valI , rule wfV_varI) + using wfg apply simp + using lookup.simps(2)[of z3 "[ B_bitvec , B_bitvec ]\<^sup>b" TRUE \ z3] by simp + have 1: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v2 ]\<^sup>c\<^sup>e : B_int " + using check_v_wf[OF infer_e_splitI(4)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce + have 2: "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e : B_bitvec" + using infer_v_wf[OF infer_e_splitI(3)] wf_weakening(1)[OF _ wfg] b_of.simps toSet.simps wfCE_valI by fastforce + + have "\; \; \ \\<^sub>w\<^sub>f \ z3 : [ B_bitvec , B_bitvec ]\<^sup>b | [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \" + proof + show "atom z3 \ (\, \, \)" using infer_e_splitI wfTh_x_fresh wfX_wfY fresh_prod3 wfG_fresh_x by metis + show "\ ; \ \\<^sub>w\<^sub>f [ B_bitvec , B_bitvec ]\<^sup>b " using wfB_pairI wfB_bitvecI infer_e_splitI wfX_wfY by metis + show "\; \; (z3, [ B_bitvec , B_bitvec ]\<^sup>b, TRUE) #\<^sub>\ + \ \\<^sub>w\<^sub>f [ v1 ]\<^sup>c\<^sup>e == [ [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e @@ [#2[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e AND [| [#1[ [ z3 ]\<^sup>v ]\<^sup>c\<^sup>e]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e" + using wfg wfz 1 2 wf_intros by meson + qed + thus ?case using infer_e_splitI by auto +qed + +lemma infer_e_wf: + fixes e::e and \::\ and \::\ and \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \" and "\ ; \ \\<^sub>w\<^sub>f \" and "\; \; \ \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>f \" and "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : (b_of \)" + using infer_e_t_wf infer_e_e_wf wfE_wf assms by metis+ + +lemma infer_e_fresh: + fixes x::x + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "atom x \ \" + shows "atom x \ (e,\)" +proof - + have "atom x \ e" using infer_e_e_wf[THEN wfE_x_fresh,OF assms(1)] assms(2) by auto + moreover have "atom x \ \" using assms infer_e_wf wfT_x_fresh by metis + ultimately show ?thesis using fresh_Pair by auto +qed + +inductive check_e :: "\ \ \ \ \ \ \ \ \ \ e \ \ \ bool" (" _ ; _ ; _ ; _ ; _ \ _ \ _" [50, 50, 50] 50) where + check_e_subtypeI: "\ infer_e T P B G D e \' ; subtype T B G \' \ \ \ check_e T P B G D e \" +equivariance check_e +nominal_inductive check_e . + +inductive_cases check_e_elims[elim!]: + "check_e F D B G \ (AE_val v) \" + "check_e F D B G \ e \" + +lemma infer_e_fst_pair: + fixes v1::v + assumes "\ ; \ ; {||} ; GNil ; \ \ [#1[ v1 , v2 ]\<^sup>v]\<^sup>e \ \" + shows "\\'. \ ; \ ; {||} ; GNil ; \ \ [v1]\<^sup>e \ \' \ + \ ; {||} ; GNil \ \' \ \" +proof - + obtain z' and b1 and b2 and c and z where ** : "\ = (\ z : b1 | C_eq (CE_val (V_var z)) (CE_fst [(V_pair v1 v2)]\<^sup>c\<^sup>e) \) \ + wfD \ {||} GNil \ \ wfPhi \ \ \ + (\ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \) \ atom z \ V_pair v1 v2 " + using infer_e_elims assms by metis + hence *:" \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \" by auto + + obtain t1a and t2a where + *: "\ ; {||} ; GNil \ v1 \ t1a \ \ ; {||} ; GNil \ v2 \ t2a \ B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)" + using infer_v_elims(5)[OF *] by metis + + hence suppv: "supp v1 = {} \ supp v2 = {} \ supp (V_pair v1 v2) = {}" using ** infer_v_v_wf wfV_supp atom_dom.simps toSet.simps supp_GNil + by (meson wfV_supp_nil) + + obtain z1 and b1' where "t1a = \ z1 : b1' | [ [ z1 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v1 ]\<^sup>c\<^sup>e \ " + using infer_v_form[of \ "{||}" GNil v1 t1a] * by auto + moreover hence "b1' = b1" using * b_of.simps by simp + ultimately have "\ ; {||} ; GNil \ v1 \ \ z1 : b1 | CE_val (V_var z1) == CE_val v1 \" using * by auto + moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e : b1" using wfCE_fstI infer_v_wf(1) ** b_of.simps wfCE_valI by metis + moreover hence st: "\ ; {||} ; GNil \ \ z1 : b1 | CE_val (V_var z1) == CE_val v1 \ \ (\ z : b1 | CE_val (V_var z) == CE_fst [V_pair v1 v2]\<^sup>c\<^sup>e \)" + using subtype_gnil_fst infer_v_v_wf by auto + moreover have "wfD \ {||} GNil \ \ wfPhi \ \" using ** by auto + ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis +qed + +lemma infer_e_snd_pair: + assumes "\ ; \ ; {||} ; GNil ; \ \ AE_snd (V_pair v1 v2) \ \" + shows "\\'. \ ; \ ; {||} ; GNil ; \ \ AE_val v2 \ \' \ \ ; {||} ; GNil \ \' \ \" +proof - + obtain z' and b1 and b2 and c and z where ** : "(\ = (\ z : b2 | C_eq (CE_val (V_var z)) (CE_snd [(V_pair v1 v2)]\<^sup>c\<^sup>e) \)) \ + (wfD \ {||} GNil \) \ (wfPhi \ \) \ + \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \ \ atom z \ V_pair v1 v2 " + using infer_e_elims(9)[OF assms(1)] by metis + hence *:" \ ; {||} ; GNil \ V_pair v1 v2 \ \ z' : B_pair b1 b2 | c \" by auto + + obtain t1a and t2a where + *: "\ ; {||} ; GNil \ v1 \ t1a \ \ ; {||} ; GNil \ v2 \ t2a \ B_pair b1 b2 = B_pair (b_of t1a) (b_of t2a)" + using infer_v_elims(5)[OF *] by metis + + hence suppv: "supp v1 = {} \ supp v2 = {} \ supp (V_pair v1 v2) = {}" using infer_v_v_wf wfV.simps v.supp by (meson "**" wfV_supp_nil) + + obtain z2 and b2' where "t2a = \ z2 : b2' | [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e == [ v2 ]\<^sup>c\<^sup>e \ " + using infer_v_form[of \ "{||}" GNil v2 t2a] * by auto + moreover hence "b2' = b2" using * b_of.simps by simp + + ultimately have "\ ; {||} ; GNil \ v2 \ \ z2 : b2 | CE_val (V_var z2) == CE_val v2 \" using * by auto + moreover have "\ ; {||} ; GNil \\<^sub>w\<^sub>f CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e : b2" using wfCE_sndI infer_v_wf(1) ** b_of.simps wfCE_valI by metis + moreover hence st: "\ ; {||} ; GNil \ \ z2 : b2 | CE_val (V_var z2) == CE_val v2 \ \ (\ z : b2 | CE_val (V_var z) == CE_snd [V_pair v1 v2]\<^sup>c\<^sup>e \)" + using subtype_gnil_snd infer_v_v_wf by auto + moreover have "wfD \ {||} GNil \ \ wfPhi \ \" using ** by metis + ultimately show ?thesis using wfX_wfY ** infer_e_valI by metis +qed + +section \Statements\ + +lemma check_s_v_unit: + assumes "\; \; \ \ (\ z : B_unit | TRUE \) \ \" and "wfD \ \ \ \" and "wfPhi \ \" + shows "\ ; \ ; \ ; \ ; \ \ AS_val (V_lit L_unit ) \ \" +proof - + have "wfG \ \ \" using assms subtype_g_wf by meson + moreover hence "wfTh \" using wfG_wf by simp + moreover obtain z'::x where "atom z' \ \" using obtain_fresh by auto + ultimately have *:"\; \; \ \ V_lit L_unit \ \ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \" + using infer_v_litI infer_unitI by simp + moreover have "wfT \ \ \ (\ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \)" using infer_v_t_wf + by (meson calculation) + moreover then have "\; \; \ \ (\ z' : B_unit | CE_val (V_var z') == CE_val (V_lit L_unit) \) \ \" using subtype_trans subtype_top assms + type_for_lit.simps(4) wfX_wfY by metis + ultimately show ?thesis using check_valI assms * by auto +qed + +lemma check_s_check_branch_s_wf: + fixes s::s and cs::branch_s and \::\ and \::\ and \::\ and \::\ and v::v and \::\ and css::branch_list + shows "\ ; \ ; B ; \ ; \ \ s \ \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " and + "check_branch_s \ \ B \ \ tid cons const v cs \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " + "check_branch_list \ \ B \ \ tid dclist v css \ \ \ ; B \\<^sub>w\<^sub>f \ \ wfTh \ \ wfD \ B \ \ \ wfT \ B \ \ \ wfPhi \ \ " +proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ B \ \ \ v \' \) + then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI + by (metis subtype_eq_base2) +next + case (check_letI x \ \ \ \ \ e \ z s b c) + then have *:"wfT \ \ ((x, b , c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \" by force + moreover have "atom x \ \" using check_letI fresh_prodN by force + ultimately have "\; \;\ \\<^sub>w\<^sub>f \" using wfT_restrict2 by force + then show ?case using check_letI infer_e_wf wfS_letI wfX_wfY wfG_elims by metis +next + case (check_assertI x \ \ \ \ \ c \ s) + then have *:"wfT \ \ ((x, B_bool , c) #\<^sub>\ \) \" by force + moreover have "atom x \ \" using check_assertI fresh_prodN by force + ultimately have "\; \;\ \\<^sub>w\<^sub>f \" using wfT_restrict2 by force + then show ?case using check_assertI wfS_assertI wfX_wfY wfG_elims by metis +next + case (check_branch_s_branchI \ \ \ \ \ cons const x v \ s tid) + then show ?case using wfX_wfY by metis +next + case (check_branch_list_consI \ \ \ \ \ tid dclist' v cs \ css ) + then show ?case using wfX_wfY by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist' v cs \ ) + then show ?case using wfX_wfY by metis +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + hence *:"wfT \ \ \ (\ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \)" (is "wfT \ \ \ ?tau") by auto + hence "wfT \ \ \ \" using wfT_wfT_if1 check_ifI fresh_prodN by metis + hence " \; \; \ \\<^sub>w\<^sub>f \ " using check_ifI b_of_c_of_eq fresh_prodN by auto + thus ?case using check_ifI by metis +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + then have "wfT \ \ ((x, b_of t, (c_of t x)) #\<^sub>\ G) \" by fastforce + moreover have "atom x \ \" using check_let2I by force + ultimately have "wfT \ \ G \" using wfT_restrict2 by metis + then show ?case using check_let2I by argo +next + case (check_varI u \ P G v \' \ s \) + then show ?case using wfG_elims wfD_elims + list.distinct list.inject by metis +next + case (check_assignI \ \ \ \ \ u \ v z \') + obtain z'::x where *:"atom z' \ \" using obtain_fresh by metis + moreover have "\ z : B_unit | TRUE \ = \ z' : B_unit | TRUE \" by auto + moreover hence "wfT \ \ \ \ z' : B_unit |TRUE \" using wfT_TRUE check_assignI check_v_wf * wfB_unitI wfG_wf by metis + ultimately show ?case using check_v.cases infer_v_wf subtype_wf check_assignI wfT_wf check_v_wf wfG_wf + by (meson subtype_wf) +next + case (check_whileI \ \ G P s1 z s2 \') + then show ?case using subtype_wf subtype_wf by auto +next + case (check_seqI \ G P s1 z s2 \) + then show ?case by fast +next + case (check_caseI \ \ \ \ \ dclist cs \ tid v z) + then show ?case by fast +qed + +lemma check_s_check_branch_s_wfS: + fixes s::s and cs::branch_s and \::\ and \::\ and \::\ and \::\ and v::v and \::\ and css::branch_list + shows "\ ; \ ; B ; \ ; \ \ s \ \ \ \ ; \ ; B ; \ ; \ \\<^sub>w\<^sub>f s : b_of \ " and + "check_branch_s \ \ B \ \ tid cons const v cs \ \ wfCS \ \ B \ \ tid cons const cs (b_of \)" + "check_branch_list \ \ B \ \ tid dclist v css \ \ wfCSS \ \ B \ \ tid dclist css (b_of \)" +proof(induct rule: check_s_check_branch_s_check_branch_list.inducts) + case (check_valI \ \ \ \ \ v \' \) + then show ?case using infer_v_wf infer_v_wf subtype_wf wfX_wfY wfS_valI + by (metis subtype_eq_base2) +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof + show \ \ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b \ using infer_e_wf check_letI b_of.simps by metis + show \ \ ; \ ; \ ; (x, b, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ + using check_letI b_of.simps wf_replace_true2(2)[OF check_letI(5)] append_g.simps by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using infer_e_wf check_letI b_of.simps by metis + show \atom x \ (\, \, \, \, \, e, b_of \)\ + apply(simp add: fresh_prodN, intro conjI) + using check_letI(1) fresh_prod7 by simp+ + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + + show \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ using check_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f c \ using check_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using check_assertI by auto + next + show \atom x \ (\, \, \, \, \, c, b_of \, s)\ using check_assertI by auto + qed +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + show \ \ ; \ ; \ ; (x, b_of const, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b_of \ \ + using wf_replace_true append_g.simps check_branch_s_branchI by metis + show \atom x \ (\, \, \, \, \, \, const)\ + using wf_replace_true append_g.simps check_branch_s_branchI fresh_prodN by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wf_replace_true append_g.simps check_branch_s_branchI by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid cons const v cs \ dclist css) + then show ?case using wf_intros by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid cons const v cs \) + then show ?case using wf_intros by metis +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case using wfS_ifI check_v_wf check_ifI b_of.simps by metis +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show \ \ ; \ ; \ ; G ; \ \\<^sub>w\<^sub>f s1 : b_of t \ using check_let2I b_of.simps by metis + show \ \; \; G \\<^sub>w\<^sub>f t \ using check_let2I check_s_check_branch_s_wf by metis + show \ \ ; \ ; \ ; (x, b_of t, TRUE) #\<^sub>\ G ; \ \\<^sub>w\<^sub>f s2 : b_of \ \ + using check_let2I(5) wf_replace_true2(2) append_g.simps check_let2I by metis + show \atom x \ (\, \, \, G, \, s1, b_of \, t)\ + apply(simp add: fresh_prodN, intro conjI) + using check_let2I(1) fresh_prod7 by simp+ + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using check_v_wf check_varI by metis + show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \' \ using check_v_wf check_varI by metis + show \atom u \ (\, \, \, \, \, \', v, b_of \)\ using check_varI fresh_prodN u_fresh_b by metis + show \ \ ; \ ; \ ; \ ; (u, \') #\<^sub>\\ \\<^sub>w\<^sub>f s : b_of \ \ using check_varI by metis + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + then show ?case using wf_intros check_v_wf subtype_eq_base2 b_of.simps by metis +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + thus ?case using wf_intros b_of.simps check_v_wf subtype_eq_base2 b_of.simps by metis +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + thus ?case using wf_intros b_of.simps by metis +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f v : B_id tid \ using check_caseI check_v_wf b_of.simps by metis + show \AF_typedef tid dclist \ set \\ using check_caseI by metis + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using check_caseI check_s_check_branch_s_wf by metis + show \ \ \\<^sub>w\<^sub>f \ \ using check_caseI check_s_check_branch_s_wf by metis + show \ \ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b_of \ \ using check_caseI by metis + qed +qed + +lemma check_s_wf: + fixes s::s + assumes "\ ; \ ; B ; \ ; \ \ s \ \" + shows "\ ; B \\<^sub>w\<^sub>f \ \ wfT \ B \ \ \ wfPhi \ \ \ wfTh \ \ wfD \ B \ \ \ wfS \ \ B \ \ s (b_of \)" + using check_s_check_branch_s_wf check_s_check_branch_s_wfS assms by meson + +lemma check_s_flip_u1: + fixes s::s and u::u and u'::u + assumes "\ ; \ ; \ ; \ ; \ \ s \ \" + shows "\ ; \ ; \ ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \" +proof - + have "(u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ ; (u \ u') \ \ \ (u \ u') \ s \ (u \ u') \ \" + using check_s.eqvt assms by blast + thus ?thesis using check_s_wf[OF assms] flip_u_eq phi_flip_eq by metis +qed + +lemma check_s_flip_u2: + fixes s::s and u::u and u'::u + assumes "\ ; \ ; B ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \" + shows "\ ; \ ; B ; \ ; \ \ s \ \" +proof - + have "\ ; \ ; B ; \ ; ( u \ u') \ ( u \ u') \ \ \ ( u \ u') \ (u \ u') \ s \ \" + using check_s_flip_u1 assms by blast + thus ?thesis using permute_flip_cancel by force +qed + +lemma check_s_flip_u: + fixes s::s and u::u and u'::u + shows "\ ; \ ; B ; \ ; ( u \ u') \ \ \ (u \ u') \ s \ \ = (\ ; \ ; B ; \ ; \ \ s \ \)" + using check_s_flip_u1 check_s_flip_u2 by metis + +lemma check_s_abs_u: + fixes s::s and s'::s and u::u and u'::u and \'::\ + assumes "[[atom u]]lst. s = [[atom u']]lst. s'" and "atom u \ \" and "atom u' \ \" + and "\ ; B ; \ \\<^sub>w\<^sub>f \'" + and "\ ; \ ; B ; \ ; ( u , \') #\<^sub>\\ \ s \ \" + shows "\ ; \ ; B ; \ ; ( u', \' ) #\<^sub>\\ \ s' \ \" +proof - + have "\ ; \ ; B ; \ ; ( u' \ u) \ (( u , \') #\<^sub>\\) \ ( u' \ u) \ s \ \" + using assms check_s_flip_u by metis + moreover have " ( u' \ u) \ (( u , \') #\<^sub>\ \) = ( u' , \') #\<^sub>\\" proof - + have " ( u' \ u) \ (( u , \') #\<^sub>\ \) = ((u' \ u) \ u, (u' \ u) \ \') #\<^sub>\ (u' \ u) \ \" + using DCons_eqvt Pair_eqvt by auto + also have "... = ( u' , (u' \ u) \ \') #\<^sub>\ \" + using assms flip_fresh_fresh by auto + also have "... = ( u' , \') #\<^sub>\ \" using + u_not_in_t fresh_def flip_fresh_fresh assms by metis + finally show ?thesis by auto + qed + moreover have "( u' \ u) \ s = s'" using assms Abs1_eq_iff(3)[of u' s' u s] by auto + ultimately show ?thesis by auto +qed + +section \Additional Elimination and Intros\ + +subsection \Values\ + +nominal_function b_for :: "opp \ b" where + "b_for Plus = B_int" +| "b_for LEq = B_bool" | "b_for Eq = B_bool" + apply(auto,simp add: eqvt_def b_for_graph_aux_def ) + by (meson opp.exhaust) +nominal_termination (eqvt) by lexicographic_order + + +lemma infer_v_pair2I: + fixes v\<^sub>1::v and v\<^sub>2::v + assumes "\; \; \ \ v\<^sub>1 \ \\<^sub>1" and "\; \; \ \ v\<^sub>2 \ \\<^sub>2" + shows "\\. \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ b_of \ = B_pair (b_of \\<^sub>1) (b_of \\<^sub>2)" +proof - + obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\\<^sub>1 = (\ z1 : b1 | c1 \) \ \\<^sub>2 = (\ z2 : b2 | c2 \)" + using \.exhaust by meson + obtain z::x where "atom z \ ( v\<^sub>1, v\<^sub>2,\, \,\)" using obtain_fresh + by blast + hence "atom z \ ( v\<^sub>1, v\<^sub>2) \ atom z \ (\, \,\)" using fresh_prodN by metis + hence " \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ z : [ b_of \\<^sub>1 , b_of \\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \" + using assms infer_v_pairI zbc by metis + moreover obtain \ where "\ = (\ z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by blast + moreover hence "b_of \ = B_pair (b_of \\<^sub>1) (b_of \\<^sub>2)" using b_of.simps zbc by presburger + ultimately show ?thesis using b_of.simps by metis +qed + +lemma infer_v_pair2I_zbc: + fixes v\<^sub>1::v and v\<^sub>2::v + assumes "\; \; \ \ v\<^sub>1 \ \\<^sub>1" and "\; \; \ \ v\<^sub>2 \ \\<^sub>2" + shows "\z \. \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \) \ atom z \ (v\<^sub>1,v\<^sub>2) \ atom z \ \" +proof - + obtain z1 and b1 and c1 and z2 and b2 and c2 where zbc: "\\<^sub>1 = (\ z1 : b1 | c1 \) \ \\<^sub>2 = (\ z2 : b2 | c2 \)" + using \.exhaust by meson + obtain z::x where * : "atom z \ ( v\<^sub>1, v\<^sub>2,\,\ , \ )" using obtain_fresh + by blast + hence vinf: " \; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ z :[ b_of \\<^sub>1 , b_of \\<^sub>2 ]\<^sup>b | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \" + using assms infer_v_pairI by simp + moreover obtain \ where "\ = (\ z : B_pair b1 b2 | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by blast + moreover have "b_of \\<^sub>1 = b1 \ b_of \\<^sub>2 = b2" using zbc b_of.simps by auto + ultimately have "\; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \ \ \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \)" by auto + thus ?thesis using * fresh_prod2 fresh_prod3 by metis +qed + +lemma infer_v_pair2E: + assumes "\; \; \ \ V_pair v\<^sub>1 v\<^sub>2 \ \" + shows "\\\<^sub>1 \\<^sub>2 z . \; \; \ \ v\<^sub>1 \ \\<^sub>1 \ \; \; \ \ v\<^sub>2 \ \\<^sub>2 \ + \ = (\ z : B_pair (b_of \\<^sub>1) (b_of \\<^sub>2) | C_eq (CE_val (V_var z)) (CE_val (V_pair v\<^sub>1 v\<^sub>2)) \) \ atom z \ (v\<^sub>1, v\<^sub>2)" +proof - + obtain z and t1 and t2 where + "\ = (\ z : B_pair (b_of t1) (b_of t2) | CE_val (V_var z) == CE_val (V_pair v\<^sub>1 v\<^sub>2) \) \ + atom z \ (v\<^sub>1, v\<^sub>2) \ \; \; \ \ v\<^sub>1 \ t1 \ \; \; \ \ v\<^sub>2 \ t2 " using infer_v_elims(3)[OF assms ] by metis + thus ?thesis using b_of.simps by metis +qed + +subsection \Expressions\ + +lemma infer_e_app2E: + fixes \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ AE_app f v \ \" + shows "\x b c s' \'. wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ + \; \; \ \ v \ \ x : b | c \ \ \ = \'[x::=v]\<^sub>\\<^sub>v \ atom x \ (\, \, \, \, \, v, \)" + using infer_e_elims(6)[OF assms] b_of.simps subst_defs by metis + +lemma infer_e_appP2E: + fixes \::\ and \::\ + assumes "\ ; \ ; \ ; \ ; \ \ AE_appP f b v \ \" + shows "\bv x ba c s' \'. wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ \ ; \ \\<^sub>w\<^sub>f b \ + (\; \; \ \ v \ \ x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \) \ (\ = \'[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v) \ atom x \ \ \ atom bv \ v" +proof - + obtain bv x ba c s' \' where *:"wfD \ \ \ \ \ Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x ba c \' s'))) = lookup_fun \ f \ \ \\<^sub>w\<^sub>f \ \ \ ; \ \\<^sub>w\<^sub>f b \ + (\; \; \ \ v \ \ x : ba[bv::=b]\<^sub>b\<^sub>b | c[bv::=b]\<^sub>c\<^sub>b \) \ (\ = \'[bv::=b]\<^sub>\\<^sub>b[x::=v]\<^sub>\\<^sub>v) \ atom x \ \ \ atom bv \ (\, \, \, \, \, b, v, \)" + using infer_e_elims(21)[OF assms] subst_defs by metis + moreover then have "atom bv \ v" using fresh_prodN by metis + ultimately show ?thesis by metis +qed + +section \Weakening\ + +text \ Lemmas showing that typing judgements hold when a context is extended \ + +lemma subtype_weakening: + fixes \'::\ + assumes "\; \; \ \ \1 \ \2" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \; \' \ \1 \ \2" + using assms proof(nominal_induct \2 avoiding: \' rule: subtype.strong_induct) + + case (subtype_baseI x \ \ \ z c z' c' b) + show ?case proof + show *:"\; \; \' \\<^sub>w\<^sub>f \ z : b | c \" using wfT_weakening subtype_baseI by metis + show "\; \; \' \\<^sub>w\<^sub>f \ z' : b | c' \" using wfT_weakening subtype_baseI by metis + show "atom x \ (\, \, \', z , c , z' , c' )" using subtype_baseI fresh_Pair by metis + have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \ toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using subtype_baseI toSet.simps by blast + moreover have "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \'" using wfT_wf_cons3[OF *, of x] subtype_baseI fresh_Pair subst_defs by metis + ultimately show "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \' \ c'[z'::=V_var x]\<^sub>v" using valid_weakening subtype_baseI by metis + qed +qed + +method many_rules uses add = ( (rule+),((simp add: add)+)?) + +lemma infer_v_g_weakening: + fixes e::e and \'::\ and v::v + assumes "\; \ ; \ \ v \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \ ; \' \ v \ \" + using assms proof(nominal_induct avoiding: \' rule: infer_v.strong_induct) + case (infer_v_varI \ \ \ b c x' z) + show ?case proof + show \ \ ; \ \\<^sub>w\<^sub>f \' \ using infer_v_varI by auto + show \Some (b, c) = lookup \' x'\ using infer_v_varI lookup_weakening by metis + show \atom z \ x'\ using infer_v_varI by auto + show \atom z \ (\, \, \')\ using infer_v_varI by auto + qed +next + case (infer_v_litI \ \ \ l \) + then show ?case using infer_v.intros by simp +next + case (infer_v_pairI z v1 v2 \ \ \ t1 t2) + then show ?case using infer_v.intros by simp +next + case (infer_v_consI s dclist \ dc tc \ \ v tv z) + show ?case proof + show \AF_typedef s dclist \ set \\ using infer_v_consI by auto + show \(dc, tc) \ set dclist\ using infer_v_consI by auto + show \ \; \; \' \ v \ tv\ using infer_v_consI by auto + show \\; \; \' \ tv \ tc\ using infer_v_consI subtype_weakening by auto + show \atom z \ v\ using infer_v_consI by auto + show \atom z \ (\, \, \')\ using infer_v_consI by auto + qed +next + case (infer_v_conspI s bv dclist \ dc tc \ \ v tv b z) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using infer_v_conspI by auto + show \(dc, tc) \ set dclist\ using infer_v_conspI by auto + show \ \; \; \' \ v \ tv\ using infer_v_conspI by auto + show \\; \; \' \ tv \ tc[bv::=b]\<^sub>\\<^sub>b\ using infer_v_conspI subtype_weakening by auto + show \atom z \ (\, \, \', v, b)\ using infer_v_conspI by auto + show \atom bv \ (\, \, \', v, b)\ using infer_v_conspI by auto + show \ \ ; \ \\<^sub>w\<^sub>f b \ using infer_v_conspI by auto + qed +qed + +lemma check_v_g_weakening: + fixes e::e and \'::\ + assumes "\; \ ; \ \ v \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\; \ ; \' \ v \ \" + using subtype_weakening infer_v_g_weakening check_v_elims check_v_subtypeI assms by metis + +lemma infer_e_g_weakening: + fixes e::e and \'::\ + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "toSet \ \ toSet \'" and "\ ; \ \\<^sub>w\<^sub>f \'" + shows "\ ; \ ; \ ; \'; \ \ e \ \" + using assms proof(nominal_induct \ avoiding: \' rule: infer_e.strong_induct) + case (infer_e_valI \ \ \ \' \ v \) + then show ?case using infer_v_g_weakening wf_weakening infer_e.intros by metis +next + case (infer_e_plusI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + moreover hence *:"\ z3 : B_int | CE_val (V_var z3) == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_plusI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op Plus v1 v2 \ \ z' : B_int | CE_val (V_var z') == CE_op Plus [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_plusI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_plusI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : B_int | c1 \\ using infer_v_g_weakening infer_e_plusI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : B_int | c2 \\ using infer_v_g_weakening infer_e_plusI by auto + show \atom z' \ AE_op Plus v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + qed + thus ?case using * by metis + +next + case (infer_e_leqI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + + moreover hence *:"\ z3 : B_bool | CE_val (V_var z3) == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_leqI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op LEq v1 v2 \ \ z' : B_bool | CE_val (V_var z') == CE_op LEq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_leqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_leqI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : B_int | c1 \\ using infer_v_g_weakening infer_e_leqI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : B_int | c2 \\ using infer_v_g_weakening infer_e_leqI by auto + show \atom z' \ AE_op LEq v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + qed + thus ?case using * by metis +next + case (infer_e_eqI \ \ \ \ \ v1 z1 bb c1 v2 z2 c2 z3) + obtain z'::x where z': "atom z' \ v1 \ atom z' \ v2 \ atom z' \ \'" using obtain_fresh fresh_prod3 by metis + + moreover hence *:"\ z3 : B_bool | CE_val (V_var z3) == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = (\ z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \)" + using infer_e_eqI type_e_eq ce.fresh fresh_e_opp by auto + + have "\ ; \ ; \ ; \' ; \ \ AE_op Eq v1 v2 \ \ z' : B_bool | CE_val (V_var z') == CE_op Eq [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_eqI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using infer_e_eqI by auto + show \ \ ; \ ; \' \ v1 \ \ z1 : bb | c1 \\ using infer_v_g_weakening infer_e_eqI by auto + show \ \ ; \ ; \' \ v2 \ \ z2 : bb | c2 \\ using infer_v_g_weakening infer_e_eqI by auto + show \atom z' \ AE_op Eq v1 v2\ using z' by auto + show \atom z' \ \'\ using z' by auto + show "bb \ {B_bool, B_int, B_unit}" using infer_e_eqI by auto + qed + thus ?case using * by metis +next + case (infer_e_appI \ \ \ \ \ f x b c \' s' v \) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \ " using wf_weakening infer_e_appI by auto + show "\ \\<^sub>w\<^sub>f \" using wf_weakening infer_e_appI by auto + show "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \' s'))) = lookup_fun \ f" using wf_weakening infer_e_appI by auto + show "\; \; \' \ v \ \ x : b | c \" using wf_weakening infer_e_appI check_v_g_weakening by auto + show "atom x \ (\, \, \, \', \, v, \)" using wf_weakening infer_e_appI by auto + show "\'[x::=v]\<^sub>v = \" using wf_weakening infer_e_appI by auto + qed + +next + case (infer_e_appPI \ \ \ \ \ b' f bv x b c \' s' v \) + + hence *:"\ ; \ ; \ ; \ ; \ \ AE_appP f b' v \ \" using Typing.infer_e_appPI by auto + + obtain x'::x where x':"atom x' \ (s', c, \', (\',v,\)) \ (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \' s'))) = (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \ x) \ c) ((x' \ x) \ \') ((x' \ x) \ s'))))" + using obtain_fresh_fun_def[of s' c \' "(\',v,\)" f x b] + by (metis fun_def.eq_iff fun_typ_q.eq_iff(2)) + + hence **: " \ x : b | c \ = \ x' : b | (x' \ x) \ c \" + using fresh_PairD(1) fresh_PairD(2) type_eq_flip by blast + have "atom x' \ \" using x' infer_e_appPI fresh_weakening fresh_Pair by metis + + show ?case proof(rule Typing.infer_e_appPI[where x=x' and bv=bv and b=b and c="(x' \ x) \ c" and \'="(x' \ x) \ \'"and s'="((x' \ x) \ s')"]) + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_appPI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_appPI by auto + show "\ ; \ \\<^sub>w\<^sub>f b'" using infer_e_appPI by auto + show "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x' b ((x' \ x) \ c) ((x' \ x) \ \') ((x' \ x) \ s')))) = lookup_fun \ f" using x' infer_e_appPI by argo + show "\; \; \' \ v \ \ x' : b[bv::=b']\<^sub>b | ((x' \ x) \ c)[bv::=b']\<^sub>b \" using ** + \.eq_iff check_v_g_weakening infer_e_appPI.hyps infer_e_appPI.prems(1) infer_e_appPI.prems subst_defs + subst_tb.simps by metis + show "atom x' \ \'" using x' fresh_prodN by metis + + have "atom x \ (v, \) \ atom x' \ (v, \)" using x' infer_e_fresh[OF *] e.fresh(2) fresh_Pair infer_e_appPI \atom x' \ \\ e.fresh by metis + moreover then have "((x' \ x) \ \')[bv::=b']\<^sub>\\<^sub>b = (x' \ x) \ (\'[bv::=b']\<^sub>\\<^sub>b)" using subst_b_x_flip + by (metis subst_b_\_def) + ultimately show "((x' \ x) \ \')[bv::=b']\<^sub>b[x'::=v]\<^sub>v = \" + using infer_e_appPI subst_tv_flip subst_defs by metis + + show "atom bv \ (\, \, \, \', \, b', v, \)" using infer_e_appPI fresh_prodN by metis + qed + +next + case (infer_e_fstI \ \ \ \ \ v z'' b1 b2 c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : b1 | CE_val (V_var z) == CE_fst [v]\<^sup>c\<^sup>e \ = \ z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_fstI v.fresh e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_fst v \ \ z' : b1 | CE_val (V_var z') == CE_fst [v]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_fstI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_fstI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_pair b1 b2 | c \" using infer_v_g_weakening infer_e_fstI by metis + show "atom z' \ AE_fst v" using * ** e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_sndI \ \ \ \ \ v z'' b1 b2 c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : b2 | CE_val (V_var z) == CE_snd [v]\<^sup>c\<^sup>e \ = \ z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_sndI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_snd v \ \ z' : b2 | CE_val (V_var z') == CE_snd [v]\<^sup>c\<^sup>e \" proof + show \ \ ; \ ;\' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_sndI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_sndI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_pair b1 b2 | c \" using infer_v_g_weakening infer_e_sndI by metis + show "atom z' \ AE_snd v" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using ** by metis +next + case (infer_e_lenI \ \ \ \ \ v z'' c z) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v \ atom z' \ c" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z : B_int | CE_val (V_var z) == CE_len [v]\<^sup>c\<^sup>e \ = \ z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_lenI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_len v \ \ z' : B_int | CE_val (V_var z') == CE_len [v]\<^sup>c\<^sup>e \" proof + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_lenI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_lenI by auto + show "\ ; \ ; \' \ v \ \ z'' : B_bitvec | c \" using infer_v_g_weakening infer_e_lenI by metis + show "atom z' \ AE_len v" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_mvarI \ \ \ \ u \) + then show ?case using wf_weakening infer_e.intros by metis +next + case (infer_e_concatI \ \ \ \ \ v1 z1 c1 v2 z2 c2 z3) + + obtain z'::x where *: "atom z' \ \' \ atom z' \ v1 \ atom z' \ v2" using obtain_fresh_z fresh_Pair by metis + hence **:"\ z3 : B_bitvec | CE_val (V_var z3) == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \ = \ z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" + using type_e_eq infer_e_concatI e.fresh ce.fresh obtain_fresh_z fresh_Pair by metis + + have "\ ; \ ; \ ; \' ; \ \ AE_concat v1 v2 \ \ z' : B_bitvec | CE_val (V_var z') == CE_concat [v1]\<^sup>c\<^sup>e [v2]\<^sup>c\<^sup>e \" proof + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_concatI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wf_weakening infer_e_concatI by auto + show "\ ; \ ; \' \ v1 \ \ z1 : B_bitvec | c1 \" using infer_v_g_weakening infer_e_concatI by metis + show "\ ; \ ; \' \ v2 \ \ z2 : B_bitvec | c2 \" using infer_v_g_weakening infer_e_concatI by metis + show "atom z' \ AE_concat v1 v2" using * e.supp by auto + show "atom z' \ \'" using * by auto + qed + thus ?case using * ** by metis +next + case (infer_e_splitI \ \ \ \ \ v1 z1 c1 v2 z2 z3) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \" using infer_e_splitI wf_weakening by auto + show "\ \\<^sub>w\<^sub>f \ " using infer_e_splitI wf_weakening by auto + show "\; \; \' \ v1 \ \ z1 : B_bitvec | c1 \" using infer_v_g_weakening infer_e_splitI by metis + show "\; \; \' \ v2 \ \ z2 : B_int | [ leq [ [ L_num 0 ]\<^sup>v ]\<^sup>c\<^sup>e [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e + AND [ leq [ [ z2 ]\<^sup>v ]\<^sup>c\<^sup>e [| [ v1 ]\<^sup>c\<^sup>e |]\<^sup>c\<^sup>e ]\<^sup>c\<^sup>e == [ [ L_true ]\<^sup>v ]\<^sup>c\<^sup>e \" + using check_v_g_weakening infer_e_splitI by metis + show "atom z1 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z1 \ \'" using infer_e_splitI by auto + show "atom z2 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z2 \ \'" using infer_e_splitI by auto + show "atom z3 \ AE_split v1 v2" using infer_e_splitI by auto + show "atom z3 \ \'" using infer_e_splitI by auto + qed +qed + +text \ Special cases proved explicitly, other cases at the end with method + \ + +lemma infer_e_d_weakening: + fixes e::e + assumes "\ ; \ ; \ ; \ ; \ \ e \ \" and "setD \ \ setD \'" and "wfD \ \ \ \'" + shows "\; \ ; \ ; \ ; \' \ e \ \" + using assms by(nominal_induct \ avoiding: \' rule: infer_e.strong_induct,auto simp add:infer_e.intros) + +lemma wfG_x_fresh_in_v_simple: + fixes x::x and v :: v + assumes "\; \; \ \ v \ \" and "atom x \ \" + shows "atom x \ v" + using wfV_x_fresh infer_v_wf assms by metis + +lemma check_s_g_weakening: + fixes v::v and s::s and cs::branch_s and x::x and c::c and b::b and \'::\ and \::\ and css::branch_list + shows "check_s \ \ \ \ \ s t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_s \ \ \ \' \ s t" and + "check_branch_s \ \ \ \ \ tid cons const v cs t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_branch_s \ \ \ \' \ tid cons const v cs t" and + "check_branch_list \ \ \ \ \ tid dclist v css t \ toSet \ \ toSet \' \ \ ; \ \\<^sub>w\<^sub>f \' \ check_branch_list \ \ \ \' \ tid dclist v css t" +proof(nominal_induct t and t and t avoiding: \' rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \' \ v \' \) + then show ?case using Typing.check_valI infer_v_g_weakening wf_weakening subtype_weakening by metis +next + case (check_letI x \ \ \ \ \ e \ z s b c) + hence xf:"atom x \ \'" by metis + show ?case proof + show "atom x \ (\, \, \, \', \, e, \)" using check_letI using fresh_prod4 xf by metis + show "\ ; \ ; \ ; \' ; \ \ e \ \ z : b | c \" using infer_e_g_weakening check_letI by metis + show "atom z \ (x, \, \, \, \', \, e, \, s)" + by(unfold fresh_prodN,auto simp add: check_letI fresh_prodN) + have "toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \) \ toSet ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using check_letI toSet.simps + by (metis Un_commute Un_empty_right Un_insert_right insert_mono) + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \')" using check_letI wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \' ; \ \ s \ \" using check_letI by metis + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show "atom x \ (\, \, \, \', \, t, s1, \)" using check_let2I using fresh_prod4 by auto + show "\ ; \ ; \ ; \' ; \ \ s1 \ t" using check_let2I by metis + have "toSet ((x, b_of t, c_of t x) #\<^sub>\ G) \ toSet ((x, b_of t, c_of t x ) #\<^sub>\ \')" using check_let2I by auto + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b_of t, c_of t x) #\<^sub>\ \')" using check_let2I wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b_of t, c_of t x) #\<^sub>\ \' ; \ \ s2 \ \" using check_let2I by metis + qed +next + case (check_branch_list_consI \ \ \ \ \ tid dclist' v cs \ css dclist) + thus ?case using Typing.check_branch_list_consI by metis +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist' v cs \ dclist) + thus ?case using Typing.check_branch_list_finalI by metis +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \ " using wf_weakening2(6) check_branch_s_branchI by metis + show "\\<^sub>w\<^sub>f \" using check_branch_s_branchI by auto + show "\; \; \' \\<^sub>w\<^sub>f \ " using check_branch_s_branchI wfT_weakening \wfG \ \ \'\ by presburger + + show "\ ; {||} ; GNil \\<^sub>w\<^sub>f const " using check_branch_s_branchI by auto + show "atom x \ (\, \, \, \', \, tid, cons, const, v, \)" using check_branch_s_branchI by auto + have "toSet ((x, b_of const, CE_val v == CE_val(V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \) \ toSet ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x) #\<^sub>\ \')" + using check_branch_s_branchI by auto + moreover hence "\ ; \ \\<^sub>w\<^sub>f ((x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \')" + using check_branch_s_branchI wfG_cons_weakening check_s_wf by metis + ultimately show "\ ; \ ; \ ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \' ; \ \ s \ \" + using check_branch_s_branchI using fresh_dom_free by auto + qed +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof + show \atom z \ (\, \, \, \', \, v, s1, s2, \)\ using fresh_prodN check_ifI by auto + show \\; \; \' \ v \ \ z : B_bool | TRUE \\ using check_v_g_weakening check_ifI by auto + show \ \ ; \ ; \ ; \' ; \ \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; \' ; \ \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + case (check_whileI \ G P s1 z s2 \') + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_seqI \ G P s1 z s2 \) + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_varI u \ \ \ \ \ \' v \ s) + thus ?case using check_v_g_weakening check_s_check_branch_s_check_branch_list.intros by auto +next + case (check_assignI \ \ \ \ \ u \ v z \') + show ?case proof + show \\ \\<^sub>w\<^sub>f \ \ using check_assignI by auto + show \\; \; \' \\<^sub>w\<^sub>f \\ using check_assignI wf_weakening by auto + show \(u, \) \ setD \\ using check_assignI by auto + show \\; \; \' \ v \ \\ using check_assignI check_v_g_weakening by auto + show \\; \; \' \ \ z : B_unit | TRUE \ \ \'\ using subtype_weakening check_assignI by auto + qed +next + case (check_caseI \ \ \ dclist cs \ tid v z) + + then show ?case using check_s_check_branch_s_check_branch_list.intros check_v_g_weakening subtype_weakening wf_weakening + by (meson infer_v_g_weakening) +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + show \atom x \ (\, \, \, \', \, c, \, s)\ using check_assertI by auto + + have " \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \" using check_assertI check_s_wf by metis + hence *:" \ ; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \'" using wfG_cons_weakening check_assertI by metis + moreover have "toSet ((x, B_bool, c) #\<^sub>\ \) \ toSet ((x, B_bool, c) #\<^sub>\ \')" using check_assertI by auto + thus \ \ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \' ; \ \ s \ \\ using check_assertI(11) [OF _ *] by auto + + show \\; \; \' \ c \ using check_assertI valid_weakening by metis + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using check_assertI wf_weakening by metis + qed +qed + +lemma wfG_xa_fresh_in_v: + fixes c::c and \::\ and G::\ and v::v and xa::x + assumes "\; \; \ \ v \ \" and "G=( \'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and "\ ; \ \\<^sub>w\<^sub>f G" + shows "atom xa \ v" +proof - + have "\; \; \ \\<^sub>w\<^sub>f v : b_of \" using infer_v_wf assms by metis + hence "supp v \ atom_dom \ \ supp \" using wfV_supp by simp + moreover have "atom xa \ atom_dom G" + using assms wfG_atoms_supp_eq[OF assms(4)] fresh_def by metis + ultimately show ?thesis using fresh_def + using assms infer_v_wf wfG_atoms_supp_eq + fresh_GCons fresh_append_g subsetCE + by (metis wfG_x_fresh_in_v_simple) +qed + +lemma fresh_z_subst_g: + fixes G::\ + assumes "atom z' \ (x,v)" and \atom z' \ G\ + shows "atom z' \ G[x::=v]\<^sub>\\<^sub>v" +proof - + have "atom z' \ v" using assms fresh_prod2 by auto + thus ?thesis using fresh_subst_gv assms by metis +qed + +lemma wfG_xa_fresh_in_subst_v: + fixes c::c and v::v and x::x and \::\ and G::\ and xa::x + assumes "\; \; \ \ v \ \" and "G=( \'@ (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and "\ ; \ \\<^sub>w\<^sub>f G" + shows "atom xa \ (subst_gv G x v)" +proof - + have "atom xa \ v" using wfG_xa_fresh_in_v assms by metis + thus ?thesis using fresh_subst_gv assms by metis +qed + +subsection \Weakening Immutable Variable Context\ + +declare check_s_check_branch_s_check_branch_list.intros[simp] +declare check_s_check_branch_s_check_branch_list.intros[intro] + +lemma check_s_d_weakening: + fixes s::s and v::v and cs::branch_s and css::branch_list + shows " \ ; \ ; \ ; \ ; \ \ s \ \ \ setD \ \ setD \' \ wfD \ \ \ \' \ \ ; \ ; \ ; \ ; \' \ s \ \" and + "check_branch_s \ \ \ \ \ tid cons const v cs \ \ setD \ \ setD \' \ wfD \ \ \ \' \ check_branch_s \ \ \ \ \' tid cons const v cs \" and + "check_branch_list \ \ \ \ \ tid dclist v css \ \ setD \ \ setD \' \ wfD \ \ \ \' \ check_branch_list \ \ \ \ \' tid dclist v css \" +proof(nominal_induct \ and \ and \ avoiding: \' arbitrary: v rule: check_s_check_branch_s_check_branch_list.strong_induct) + case (check_valI \ \ \ \ \ v \' \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by blast +next + case (check_letI x \ \ \ \ \ e \ z s b c) + show ?case proof + show "atom x \ (\, \, \, \, \', e, \)" using check_letI by auto + show "atom z \ (x, \, \, \, \, \', e, \, s)" using check_letI by auto + show "\ ; \ ; \ ; \ ; \' \ e \ \ z : b | c \" using check_letI infer_e_d_weakening by auto + have "\ ; \ \\<^sub>w\<^sub>f (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \" using check_letI check_s_wf by metis + moreover have "\; \; \ \\<^sub>w\<^sub>f \'" using check_letI check_s_wf by metis + ultimately have "\; \; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) toSet.simps by fast + thus "\ ; \ ; \ ; (x, b, c[z::=V_var x]\<^sub>v) #\<^sub>\ \ ; \' \ s \ \" using check_letI by simp + qed +next + case (check_branch_s_branchI \ \ \ \ \ const x \ tid cons v s) + moreover have "\ ;\ \\<^sub>w\<^sub>f (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \" + using check_s_wf[OF check_branch_s_branchI(16) ] by metis + moreover hence " \ ;\ ; (x, b_of const, CE_val v == CE_val (V_cons tid cons (V_var x)) AND c_of const x ) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" + using wf_weakening2(6) check_branch_s_branchI by fastforce + ultimately show ?case + using check_s_check_branch_s_check_branch_list.intros by simp +next + case (check_branch_list_consI \ \ \ \ \ tid dclist v cs \ css) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_branch_list_finalI \ \ \ \ \ tid dclist v cs \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_ifI z \ \ \ \ \ v s1 s2 \) + show ?case proof + show \atom z \ (\, \, \, \, \', v, s1, s2, \)\ using fresh_prodN check_ifI by auto + show \\; \; \ \ v \ \ z : B_bool | TRUE \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \' \ s1 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_true) IMP c_of \ z \\ using check_ifI by auto + show \ \ ; \ ; \ ; \ ; \' \ s2 \ \ z : b_of \ | CE_val v == CE_val (V_lit L_false) IMP c_of \ z \\ using check_ifI by auto + qed +next + case (check_assertI x \ \ \ \ \ c \ s) + show ?case proof + show "atom x \ (\, \, \, \, \', c, \,s)" using fresh_prodN check_assertI by auto + show *:" \; \; \ \\<^sub>w\<^sub>f \' " using check_assertI by auto + hence "\; \; (x, B_bool, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \' " using wf_weakening2(6)[OF *, of " (x, B_bool, c) #\<^sub>\ \"] check_assertI check_s_wf toSet.simps by auto + thus "\ ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \' \ s \ \" + using check_assertI(11)[OF \setD \ \ setD \'\] by simp + (* wf_weakening2(6) check_s_wf check_assertI *) + show "\; \; \ \ c " using fresh_prodN check_assertI by auto + + qed +next + case (check_let2I x \ \ \ G \ t s1 \ s2) + show ?case proof + show "atom x \ (\, \, \, G, \', t , s1, \)" using check_let2I by auto + + show "\ ; \ ; \ ; G ; \' \ s1 \ t" using check_let2I infer_e_d_weakening by auto + have "\; \; (x, b_of t, c_of t x ) #\<^sub>\ G \\<^sub>w\<^sub>f \'" using check_let2I wf_weakening2(6) check_s_wf by fastforce + thus "\ ; \ ; \ ; (x, b_of t, c_of t x) #\<^sub>\ G ; \' \ s2 \ \" using check_let2I by simp + qed +next + case (check_varI u \ \ \ \ \ \' v \ s) + show ?case proof + show "atom u \ (\, \, \, \, \', \', v, \)" using check_varI by auto + show "\; \; \ \ v \ \'" using check_varI by auto + have "setD ((u, \') #\<^sub>\\) \ setD ((u, \') #\<^sub>\\')" using setD.simps check_varI by auto + moreover have "u \ fst ` setD \'" using check_varI(1) setD.simps fresh_DCons by (simp add: fresh_d_not_in) + moreover hence "\; \; \ \\<^sub>w\<^sub>f (u, \') #\<^sub>\\' " using wfD_cons fresh_DCons setD.simps check_varI check_v_wf by metis + ultimately show "\ ; \ ; \ ; \ ; (u, \') #\<^sub>\\' \ s \ \" using check_varI by auto + qed +next + case (check_assignI \ \ \ \ \ u \ v z \') + moreover hence "(u, \) \ setD \'" by auto + ultimately show ?case using Typing.check_assignI by simp +next + case (check_whileI \ \ \ \ \ s1 z s2 \') + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_seqI \ \ \ \ \ s1 z s2 \) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson +next + case (check_caseI \ \ \ \ \ tid dclist v cs \ z) + then show ?case using check_s_check_branch_s_check_branch_list.intros by meson + +qed + +lemma valid_ce_eq: + fixes v::v and ce2::ce + assumes "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" and "wfV \ \ GNil v b" and "wfCE \ \ ((x, b, TRUE) #\<^sub>\ GNil) ce2 b'" and "wfCE \ \ GNil ce1 b'" + shows \\; \; (x, b, ([[x ]\<^sup>v]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\ GNil \ ce1 == ce2 \ + unfolding valid.simps proof + have wfg: "\ ; \ \\<^sub>w\<^sub>f (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" + using wfG_cons1I wfG_nilI wfX_wfY assms wf_intros + by (meson fresh_GNil wfC_e_eq wfG_intros2) + + show wf: \\; \; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1 == ce2 \ + apply(rule wfC_eqI[where b=b']) + using wfg toSet.simps assms wfCE_weakening apply simp + + using wfg assms wf_replace_inside1(8) assms + using wfC_trueI wf_trans(8) by auto + + show \\i. ((\ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \ i) \ (i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) \ + (i \ (ce1 == ce2 )))\ proof(rule+,goal_cases) + fix i + assume as:"\ ; (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil \ i \ i \ (x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil" + have 1:"wfV \ \ ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) v b" + using wf_weakening assms append_g.simps toSet.simps wf wfX_wfY + by (metis empty_subsetI) + hence "\s. i \ v \ ~ s" using eval_v_exist[OF _ 1] as by auto + then obtain s where iv:"i\ v \ ~ s" .. + + hence ix:"i x = Some s" proof - + have "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e" using is_satis_g.simps as by auto + hence "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ ~ True" using is_satis.simps by auto + hence "i \ [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e \ ~ s" using + iv eval_e_elims + by (metis eval_c_elims(7) eval_e_uniqueness eval_e_valI) + thus ?thesis using eval_v_elims(2) eval_e_elims(1) by metis + qed + + have 1:"wfCE \ \ ((x, b, [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e ) #\<^sub>\ GNil) ce1 b'" + using wfCE_weakening assms append_g.simps toSet.simps wf wfX_wfY + by (metis empty_subsetI) + hence "\s1. i \ ce1 \ ~ s1" using eval_e_exist assms as by auto + then obtain s1 where s1: "i\ce1\ ~ s1" .. + + moreover have "i\ ce2 \ ~ s1" proof - + have "i\ ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v \ ~ s1" using assms s1 by auto + moreover have "ce1 = ce2[x::=v]\<^sub>c\<^sub>e\<^sub>v" using subst_v_ce_def assms subst_v_simple_commute by auto + ultimately have "i(x \ s) \ ce2 \ ~ s1" + using ix subst_e_eval_v[of i ce1 s1 "ce2[z::=[ x ]\<^sup>v]\<^sub>v" x v s] iv s1 by auto + moreover have "i(x \ s) = i" using ix by auto + ultimately show ?thesis by auto + qed + ultimately show "i \ ce1 == ce2 \ ~ True " using eval_c_eqI by metis + qed +qed + + +lemma check_v_top: + fixes v::v + assumes "\; \; GNil \ v \ \" and "ce1 = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" and "\; \; GNil \\<^sub>w\<^sub>f \ z : b_of \ | ce1 == ce2 \" + and "supp ce1 \ supp \" + shows "\; \; GNil \ v \ \ z : b_of \ | ce1 == ce2 \" +proof - + obtain t where t: "\; \; GNil \ v \ t \ \; \; GNil \ t \ \" + using assms check_v_elims by metis + + then obtain z' and b' where *:"t = \ z' : b' | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ atom z' \ v \ atom z' \ (\, \,GNil)" + using assms infer_v_form by metis + have beq: "b_of t = b_of \" using subtype_eq_base2 b_of.simps t by auto + obtain x::x where xf: \atom x \ (\, \, GNil, z', [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e , z, ce1 == ce2 )\ + using obtain_fresh by metis + + have "\; \; (x, b_of \, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v )" + using wfT_wfC2[OF assms(3), of x] subst_cv.simps(6) subst_v_c_def subst_v_ce_def fresh_GNil by simp + + then obtain b2 where b2: "\; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ + \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2" using wfC_elims(3) + beq by metis + + from xf have "\; \; GNil \ \ z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ \ z : b_of t | ce1 == ce2 \" + proof + show \ \; \; GNil \\<^sub>w\<^sub>f \ z' : b_of t | [ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e \ \ using b_of.simps assms infer_v_wf t * by auto + show \ \; \; GNil \\<^sub>w\<^sub>f \ z : b_of t | ce1 == ce2 \ \ using beq assms by auto + have \\; \; (x, b_of t, ([ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )) #\<^sub>\ GNil \ (ce1[z::=[ x ]\<^sup>v]\<^sub>v == ce2[z::=[ x ]\<^sup>v]\<^sub>v ) \ + proof(rule valid_ce_eq) + show \ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\ proof - + have "atom z \ ce1" using assms fresh_def x_not_in_b_set by fast + hence "ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce1" + using forget_subst_v by auto + also have "... = ce2[z::=v]\<^sub>c\<^sub>e\<^sub>v" using assms by auto + also have "... = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v" proof - + have "atom x \ ce2" using xf fresh_prodN c.fresh by metis + thus ?thesis using subst_v_simple_commute subst_v_ce_def by simp + qed + finally show ?thesis by auto + qed + show \ \; \; GNil \\<^sub>w\<^sub>f v : b_of t \ using infer_v_wf t by simp + show \ \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce2[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ using b2 by auto + + have " \; \; (x, b_of t, TRUE) #\<^sub>\ GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2" using b2 by auto + moreover have "atom x \ ce1[z::=[ x ]\<^sup>v]\<^sub>v" + using fresh_subst_v_if assms fresh_def + using \\; \; GNil \\<^sub>w\<^sub>f v : b_of t\ \ce1[z::=[ x ]\<^sup>v]\<^sub>v = ce2[z::=[ x ]\<^sup>v]\<^sub>v[x::=v]\<^sub>c\<^sub>e\<^sub>v\ + fresh_GNil subst_v_ce_def wfV_x_fresh by auto + ultimately show \ \; \; GNil \\<^sub>w\<^sub>f ce1[z::=[ x ]\<^sup>v]\<^sub>v : b2 \ using + wf_restrict(8) by force + qed + moreover have v: "v[z'::=[ x ]\<^sup>v]\<^sub>v\<^sub>v = v" + using forget_subst assms infer_v_wf wfV_supp x_not_in_b_set + by (simp add: "local.*") + ultimately show "\; \; (x, b_of t, ([ [ z' ]\<^sup>v ]\<^sup>c\<^sup>e == [ v ]\<^sup>c\<^sup>e )[z'::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ GNil \ (ce1 == ce2 )[z::=[ x ]\<^sup>v]\<^sub>v" + unfolding subst_cv.simps subst_v_c_def subst_cev.simps subst_vv.simps + using subst_v_ce_def by simp + qed + thus ?thesis using b_of.simps assms * check_v_subtypeI t b_of.simps subtype_eq_base2 by metis +qed + +end + diff --git a/thys/MiniSail/Wellformed.thy b/thys/MiniSail/Wellformed.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/Wellformed.thy @@ -0,0 +1,581 @@ +(*<*) +theory Wellformed + imports IVSubst BTVSubst +begin + (*>*) + +chapter \Wellformed Terms\ + +text \We require that expressions and values are well-formed. This includes a notion +of well-sortedness. We identify a sort with a basic type and +define the judgement as two clusters of mutually recursive +inductive predicates. Some of the proofs are across all of the predicates and +although they seemed at first to be daunting, they have all +worked out well.\ + +named_theorems ms_wb "Facts for helping with well-sortedness" + +section \Definitions\ + +inductive wfV :: "\ \ \ \ \ \ v \ b \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfC :: "\ \ \ \ \ \ c \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfG :: "\ \ \ \ \ \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfT :: "\ \ \ \ \ \ \ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfTs :: "\ \ \ \ \ \ (string*\) list \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfTh :: "\ \ bool" (" \\<^sub>w\<^sub>f _ " [50] 50) and + wfB :: "\ \ \ \ b \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfCE :: "\ \ \ \ \ \ ce \ b \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfTD :: "\ \ type_def \ bool" (" _ \\<^sub>w\<^sub>f _ " [50,50] 50) + where + +wfB_intI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_int" +| wfB_boolI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_bool" +| wfB_unitI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_unit" +| wfB_bitvecI: "\\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f B_bitvec" +| wfB_pairI: "\ \; \ \\<^sub>w\<^sub>f b1 ; \; \ \\<^sub>w\<^sub>f b2 \ \ \; \ \\<^sub>w\<^sub>f B_pair b1 b2" + +| wfB_consI: "\ + \\<^sub>w\<^sub>f \; + (AF_typedef s dclist) \ set \ +\ \ + \; \ \\<^sub>w\<^sub>f B_id s" + +| wfB_appI: "\ + \\<^sub>w\<^sub>f \; + \; \ \\<^sub>w\<^sub>f b; + (AF_typedef_poly s bv dclist) \ set \ +\ \ + \; \ \\<^sub>w\<^sub>f B_app s b" + +| wfV_varI: "\ \; \ \\<^sub>w\<^sub>f \ ; Some (b,c) = lookup \ x \ \ \; \; \ \\<^sub>w\<^sub>f V_var x : b" +| wfV_litI: "\; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f V_lit l : base_for_lit l" + +| wfV_pairI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : b1 ; + \; \; \ \\<^sub>w\<^sub>f v2 : b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f (V_pair v1 v2) : B_pair b1 b2" + +| wfV_consI: "\ + AF_typedef s dclist \ set \; + (dc, \ x : b' | c \) \ set dclist ; + \; \; \ \\<^sub>w\<^sub>f v : b' +\ \ + \; \; \ \\<^sub>w\<^sub>f V_cons s dc v : B_id s" + +| wfV_conspI: "\ + AF_typedef_poly s bv dclist \ set \; + (dc, \ x : b' | c \) \ set dclist ; + \ ; \ \\<^sub>w\<^sub>f b; + atom bv \ (\, \, \, b , v); + \; \; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b +\ \ + \; \; \ \\<^sub>w\<^sub>f V_consp s dc b v : B_app s b" + +| wfCE_valI : "\ + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_val v : b" + +| wfCE_plusI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op Plus v1 v2 : B_int" + +| wfCE_leqI:"\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op LEq v1 v2 : B_bool" + +| wfCE_eqI:"\ + \; \; \ \\<^sub>w\<^sub>f v1 : b; + \; \; \ \\<^sub>w\<^sub>f v2 : b +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_op Eq v1 v2 : B_bool" + +| wfCE_fstI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_fst v1 : b1" + +| wfCE_sndI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_snd v1 : b2" + +| wfCE_concatI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec ; + \; \; \ \\<^sub>w\<^sub>f v2 : B_bitvec +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_concat v1 v2 : B_bitvec" + +| wfCE_lenI: "\ + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec +\ \ + \; \; \ \\<^sub>w\<^sub>f CE_len v1 : B_int" + +| wfTI : "\ + atom z \ (\, \, \) ; + \; \ \\<^sub>w\<^sub>f b; + \; \ ; (z,b,C_true) #\<^sub>\ \ \\<^sub>w\<^sub>f c +\ \ + \; \; \ \\<^sub>w\<^sub>f \ z : b | c \" + +| wfC_eqI: "\ + \; \; \ \\<^sub>w\<^sub>f e1 : b ; + \; \; \ \\<^sub>w\<^sub>f e2 : b \ \ + \; \; \ \\<^sub>w\<^sub>f C_eq e1 e2" +| wfC_trueI: " \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f C_true " +| wfC_falseI: " \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f C_false " + +| wfC_conjI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_conj c1 c2" +| wfC_disjI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_disj c1 c2" +| wfC_notI: "\ \; \; \ \\<^sub>w\<^sub>f c1 \ \ \; \; \ \\<^sub>w\<^sub>f C_not c1" +| wfC_impI: "\ \; \; \ \\<^sub>w\<^sub>f c1 ; + \; \; \ \\<^sub>w\<^sub>f c2 \ \ \; \; \ \\<^sub>w\<^sub>f C_imp c1 c2" + +| wfG_nilI: " \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f GNil" +| wfG_cons1I: "\ c \ { TRUE, FALSE } ; + \; \ \\<^sub>w\<^sub>f \ ; + atom x \ \ ; + \ ; \ ; (x,b,C_true)#\<^sub>\\ \\<^sub>w\<^sub>f c ; wfB \ \ b + \ \ \; \ \\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\\)" +| wfG_cons2I: "\ c \ { TRUE, FALSE } ; + \; \ \\<^sub>w\<^sub>f \ ; + atom x \ \ ; + wfB \ \ b + \ \ \; \ \\<^sub>w\<^sub>f ((x,b,c)#\<^sub>\\)" + +| wfTh_emptyI: " \\<^sub>w\<^sub>f []" + +| wfTh_consI: "\ + (name_of_type tdef) \ name_of_type ` set \ ; + \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f tdef \ \ \\<^sub>w\<^sub>f tdef#\" + +| wfTD_simpleI: "\ + \ ; {||} ; GNil \\<^sub>w\<^sub>f lst + \ \ + \ \\<^sub>w\<^sub>f (AF_typedef s lst )" + +| wfTD_poly: "\ + \ ; {|bv|} ; GNil \\<^sub>w\<^sub>f lst + \ \ + \ \\<^sub>w\<^sub>f (AF_typedef_poly s bv lst)" + +| wfTs_nil: "\; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f []::(string*\) list" +| wfTs_cons: "\ \; \; \ \\<^sub>w\<^sub>f \ ; + dc \ fst ` set ts; + \; \; \ \\<^sub>w\<^sub>f ts::(string*\) list \ \ \; \; \ \\<^sub>w\<^sub>f ((dc,\)#ts)" + +inductive_cases wfC_elims: + "\; \; \ \\<^sub>w\<^sub>f C_true" + "\; \; \ \\<^sub>w\<^sub>f C_false" + "\; \; \ \\<^sub>w\<^sub>f C_eq e1 e2" + "\; \; \ \\<^sub>w\<^sub>f C_conj c1 c2" + "\; \; \ \\<^sub>w\<^sub>f C_disj c1 c2" + "\; \; \ \\<^sub>w\<^sub>f C_not c1" + "\; \; \ \\<^sub>w\<^sub>f C_imp c1 c2" + +inductive_cases wfV_elims: + "\; \; \ \\<^sub>w\<^sub>f V_var x : b" + "\; \; \ \\<^sub>w\<^sub>f V_lit l : b" + "\; \; \ \\<^sub>w\<^sub>f V_pair v1 v2 : b" + "\; \; \ \\<^sub>w\<^sub>f V_cons tyid dc v : b" + "\; \; \ \\<^sub>w\<^sub>f V_consp tyid dc b v : b'" + +inductive_cases wfCE_elims: + " \; \; \ \\<^sub>w\<^sub>f CE_val v : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op Plus v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op LEq v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_fst v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_snd v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_concat v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_len v1 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op opp v1 v2 : b" + " \; \; \ \\<^sub>w\<^sub>f CE_op Eq v1 v2 : b" + +inductive_cases wfT_elims: + "\; \ ; \ \\<^sub>w\<^sub>f \::\" + "\; \ ; \ \\<^sub>w\<^sub>f \ z : b | c \" + +inductive_cases wfG_elims: + "\; \ \\<^sub>w\<^sub>f GNil" + "\; \ \\<^sub>w\<^sub>f (x,b,c)#\<^sub>\\" + "\; \ \\<^sub>w\<^sub>f (x,b,TRUE)#\<^sub>\\" + "\; \ \\<^sub>w\<^sub>f (x,b,FALSE)#\<^sub>\\" + +inductive_cases wfTh_elims: + " \\<^sub>w\<^sub>f []" + " \\<^sub>w\<^sub>f td#\" + +inductive_cases wfTD_elims: + "\ \\<^sub>w\<^sub>f (AF_typedef s lst )" + "\ \\<^sub>w\<^sub>f (AF_typedef_poly s bv lst )" + +inductive_cases wfTs_elims: + "\; \ ; GNil \\<^sub>w\<^sub>f ([]::((string*\) list))" + "\; \ ; GNil \\<^sub>w\<^sub>f ((t#ts)::((string*\) list))" + +inductive_cases wfB_elims: + " \; \ \\<^sub>w\<^sub>f B_pair b1 b2" + " \; \ \\<^sub>w\<^sub>f B_id s" + " \; \ \\<^sub>w\<^sub>f B_app s b" + +equivariance wfV + +text \This setup of 'avoiding' is not complete and for some of lemmas, such as weakening, +do it the hard way\ +nominal_inductive wfV + avoids wfV_conspI: bv | wfTI: z +proof(goal_cases) + case (1 s bv dclist \ dc x b' c \ b \ v) + + moreover hence "atom bv \ V_consp s dc b v" using v.fresh fresh_prodN pure_fresh by metis + moreover have "atom bv \ B_app s b" using b.fresh fresh_prodN pure_fresh 1 by metis + ultimately show ?case using b.fresh v.fresh pure_fresh fresh_star_def fresh_prodN by fastforce +next + case (2 s bv dclist \ dc x b' c \ b \ v) + then show ?case by auto +next + case (3 z \ \ \ b c) + then show ?case using \.fresh fresh_star_def fresh_prodN by fastforce +next + case (4 z \ \ \ b c) + then show ?case by auto +qed + +inductive + wfE :: "\ \ \ \ \ \ \ \ \ \ e \ b \ bool" (" _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfS :: "\ \ \ \ \ \ \ \ \ \ s \ b \ bool" (" _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50] 50) and + wfCS :: "\ \ \ \ \ \ \ \ \ \ tyid \ string \ \ \ branch_s \ b \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and + wfCSS :: "\ \ \ \ \ \ \ \ \ \ tyid \ (string * \) list \ branch_list \ b \ bool" (" _ ; _ ; _ ; _ ; _ ; _ ; _ \\<^sub>w\<^sub>f _ : _ " [50,50,50,50,50] 50) and + wfPhi :: "\ \ \ \ bool" (" _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfD :: "\ \ \ \ \ \ \ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50,50] 50) and + wfFTQ :: "\ \ \ \ fun_typ_q \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ " [50] 50) and + wfFT :: "\ \ \ \ \ \ fun_typ \ bool" (" _ ; _ ; _ \\<^sub>w\<^sub>f _ " [50] 50) where + +wfE_valI : "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_val v : b" + +| wfE_plusI: "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op Plus v1 v2 : B_int" + +| wfE_leqI:"\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_int; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op LEq v1 v2 : B_bool" + +| wfE_eqI:"\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : b; + \; \; \ \\<^sub>w\<^sub>f v2 : b; + b \ { B_bool, B_int, B_unit } +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_op Eq v1 v2 : B_bool" + +| wfE_fstI: "\ + \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 + \ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_fst v1 : b1" + +| wfE_sndI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_pair b1 b2 +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_snd v1 : b2" + +| wfE_concatI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec; + \; \; \ \\<^sub>w\<^sub>f v2 : B_bitvec +\ \ + \ ; \ ; \ ; \; \ \\<^sub>w\<^sub>f AE_concat v1 v2 : B_bitvec" + +| wfE_splitI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec; + \; \; \ \\<^sub>w\<^sub>f v2 : B_int +\ \ + \ ; \ ; \ ; \; \ \\<^sub>w\<^sub>f AE_split v1 v2 : B_pair B_bitvec B_bitvec" + +| wfE_lenI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \; \ \\<^sub>w\<^sub>f v1 : B_bitvec +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_len v1 : B_int" + +| wfE_appI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f ; + \; \; \ \\<^sub>w\<^sub>f v : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_app f v : b_of \" + +| wfE_appPI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \; \ \\<^sub>w\<^sub>f b'; + atom bv \ (\, \, \, \, \, b', v, (b_of \)[bv::=b']\<^sub>b); + Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f; + \; \; \ \\<^sub>w\<^sub>f v : (b[bv::=b']\<^sub>b) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f (AE_appP f b' v) : ((b_of \)[bv::=b']\<^sub>b)" + +| wfE_mvarI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f \; + (u,\) \ setD \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AE_mvar u : b_of \" + +| wfS_valI: "\ + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f v : b ; + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f (AS_val v) : b" + +| wfS_letI: "\ + wfE \ \ \ \ \ e b' ; + \ ; \ ; \ ; (x,b',C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b; + \; \; \ \\<^sub>w\<^sub>f \ ; + atom x \ (\, \, \, \, \, e, b) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f LET x = e IN s : b" + +| wfS_assertI: "\ + \ ; \ ; \ ; (x,B_bool,c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b; + \; \; \ \\<^sub>w\<^sub>f c ; + \; \; \ \\<^sub>w\<^sub>f \ ; + atom x \ (\, \, \, \, \, c, b, s) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f ASSERT c IN s : b" + +| wfS_let2I: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : b_of \ ; + \; \; \ \\<^sub>w\<^sub>f \; + \ ; \ ; \ ; (x,b_of \,C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : b ; + atom x \ (\, \, \, \, \, s1, b,\) +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f LET x : \ = s1 IN s2 : b" + +| wfS_ifI: "\ + \; \; \ \\<^sub>w\<^sub>f v : B_bool; + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : b ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f IF v THEN s1 ELSE s2 : b" + +| wfS_varI : "\ + wfT \ \ \ \ ; + \; \; \ \\<^sub>w\<^sub>f v : b_of \; + atom u \ (\, \, \, \, \, \, v, b); + \ ; \ ; \ ; \ ; (u,\)#\<^sub>\\ \\<^sub>w\<^sub>f s : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f VAR u : \ = v IN s : b " + +| wfS_assignI: "\ + (u,\) \ setD \ ; + \; \; \ \\<^sub>w\<^sub>f \ ; + \ \\<^sub>w\<^sub>f \ ; + \; \; \ \\<^sub>w\<^sub>f v : b_of \ +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f u ::= v : B_unit" + +| wfS_whileI: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : B_bool ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f WHILE s1 DO { s2 } : b" + +| wfS_seqI: "\ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 : B_unit ; + \; \; \; \; \ \\<^sub>w\<^sub>f s2 : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f s1 ;; s2 : b" + +| wfS_matchI: "\ + wfV \ \ \ v (B_id tid) ; + (AF_typedef tid dclist ) \ set \; + wfD \ \ \ \ ; + \ \\<^sub>w\<^sub>f \ ; + \ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f cs : b +\ \ + \; \; \; \; \ \\<^sub>w\<^sub>f AS_match v cs : b " + +| wfS_branchI: "\ + \ ; \ ; \ ; (x,b_of \,C_true) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b ; + atom x \ (\, \, \, \, \, \,\); + \; \; \ \\<^sub>w\<^sub>f \ +\ \ + \ ; \ ; \ ; \ ; \ ; tid ; dc ; \ \\<^sub>w\<^sub>f dc x \ s : b" + +| wfS_finalI: "\ + \; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b + \ \ + \ ; \ ; \ ; \ ; \ ; tid ; [(dc,t)] \\<^sub>w\<^sub>f AS_final cs : b " + +| wfS_cons: "\ + \; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b; + \; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b + \ \ + \ ; \ ; \ ; \ ; \ ; tid ; (dc,t)#dclist \\<^sub>w\<^sub>f AS_cons cs css : b " + +| wfD_emptyI: "\; \ \\<^sub>w\<^sub>f \ \ \ ; \ ; \ \\<^sub>w\<^sub>f []\<^sub>\" + +| wfD_cons: "\ + \ ; \ ; \ \\<^sub>w\<^sub>f \::\ ; + \ ; \ ; \ \\<^sub>w\<^sub>f \; + u \ fst ` setD \ +\ \ + \; \; \ \\<^sub>w\<^sub>f ((u,\) #\<^sub>\ \)" + +| wfPhi_emptyI: " \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f []" + +| wfPhi_consI: "\ + f \ name_of_fun ` set \; + \ ; \ \\<^sub>w\<^sub>f ft; + \ \\<^sub>w\<^sub>f \ +\ \ + \ \\<^sub>w\<^sub>f ((AF_fundef f ft)#\)" + +| wfFTNone: " \ ; \ ; {||} \\<^sub>w\<^sub>f ft \ \ ; \ \\<^sub>w\<^sub>f AF_fun_typ_none ft" +| wfFTSome: " \ ; \ ; {| bv |} \\<^sub>w\<^sub>f ft \ \ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv ft" + +| wfFTI: "\ + \ ; B \\<^sub>w\<^sub>f b; + supp s \ {atom x} \ supp B ; + supp c \ { atom x } ; + \ ; B ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \; + \ \\<^sub>w\<^sub>f \ +\ \ + \ ; \ ; B \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" + +inductive_cases wfE_elims: + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_val v : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op Plus v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op LEq v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_fst v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_snd v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_concat v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_len v1 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op opp v1 v2 : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_app f v: b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_appP f b' v: b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_mvar u : b" + "\; \; \; \; \ \\<^sub>w\<^sub>f AE_op Eq v1 v2 : b" + +inductive_cases wfCS_elims: + "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f (cs::branch_s) : b" + "\; \; \; \; \ ; tid ; dc \\<^sub>w\<^sub>f (cs::branch_list) : b" + +inductive_cases wfPhi_elims: + " \ \\<^sub>w\<^sub>f []" + " \ \\<^sub>w\<^sub>f ((AF_fundef f ft)#\)" + " \ \\<^sub>w\<^sub>f (fd#\::\)" + +declare[[ simproc del: alpha_lst]] + +inductive_cases wfFTQ_elims: + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_none ft" + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv ft" + "\ ; \ \\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \ s)" + +inductive_cases wfFT_elims: + "\ ; \ ; \ \\<^sub>w\<^sub>f AF_fun_typ x b c \ s" + +declare[[ simproc add: alpha_lst]] + +inductive_cases wfD_elims: + "\ ; \ ; (\::\) \\<^sub>w\<^sub>f []\<^sub>\" + "\ ; \ ; (\::\) \\<^sub>w\<^sub>f (u,\) #\<^sub>\ \::\" + +equivariance wfE + +nominal_inductive wfE + avoids wfE_appPI: bv | wfS_varI: u | wfS_letI: x | wfS_let2I: x | wfS_branchI: x | wfS_assertI: x + +proof(goal_cases) + case (1 \ \ \ \ \ b' bv v \ f x b c s) + moreover hence "atom bv \ AE_appP f b' v" using pure_fresh fresh_prodN e.fresh by auto + ultimately show ?case using fresh_star_def by fastforce +next + case (2 \ \ \ \ \ b' bv v \ f x b c s) + then show ?case by auto +next + case (3 \ \ \ \ \ e b' x s b) + moreover hence "atom x \ LET x = e IN s" using fresh_prodN by auto + ultimately show ?case using fresh_prodN fresh_star_def by fastforce +next + case (4 \ \ \ \ \ e b' x s b) + then show ?case by auto +next + case (5 \ \ \ x c \ \ s b) + hence "atom x \ ASSERT c IN s" using s_branch_s_branch_list.fresh by auto + then show ?case using fresh_prodN fresh_star_def 5 by fastforce +next + case (6 \ \ \ x c \ \ s b) + then show ?case by auto +next + case (7 \ \ \ \ \ s1 \ x s2 b) + hence "atom x \ \ \ atom x \ s1" using fresh_prodN by metis + moreover hence "atom x \ LET x : \ = s1 IN s2" + using s_branch_s_branch_list.fresh(3)[of "atom x" x "\" s1 s2 ] fresh_prodN by simp + ultimately show ?case using fresh_prodN fresh_star_def 7 by fastforce +next + case (8 \ \ \ \ \ s1 \ x s2 b) + then show ?case by auto +next + case (9 \ \ \ \ v u \ \ b s) + moreover hence " atom u \ AS_var u \ v s" using fresh_prodN s_branch_s_branch_list.fresh by simp + ultimately show ?case using fresh_star_def fresh_prodN s_branch_s_branch_list.fresh by fastforce +next + case (10 \ \ \ \ v u \ \ b s) + then show ?case by auto +next + case (11 \ \ \ x \ \ \ s b tid dc) + moreover have "atom x \ (dc x \ s)" using pure_fresh s_branch_s_branch_list.fresh by auto + ultimately show ?case using fresh_prodN fresh_star_def pure_fresh by fastforce +next + case (12 \ \ \ x \ \ \ s b tid dc) + then show ?case by auto +qed + +inductive wfVDs :: "var_def list \ bool" where + +wfVDs_nilI: "wfVDs []" + +| wfVDs_consI: "\ + atom u \ ts; + wfV ([]::\) {||} GNil v (b_of \); + wfT ([]::\) {||} GNil \; + wfVDs ts +\ \ wfVDs ((AV_def u \ v)#ts)" + +equivariance wfVDs +nominal_inductive wfVDs . + +end \ No newline at end of file diff --git a/thys/MiniSail/WellformedL.thy b/thys/MiniSail/WellformedL.thy new file mode 100644 --- /dev/null +++ b/thys/MiniSail/WellformedL.thy @@ -0,0 +1,4755 @@ +(*<*) +theory WellformedL + imports Wellformed "SyntaxL" +begin +(*>*) + +chapter \Wellformedness Lemmas\ + +section \Prelude\ + +lemma b_of_subst_bb_commute: + "(b_of (\[bv::=b]\<^sub>\\<^sub>b)) = (b_of \)[bv::=b]\<^sub>b\<^sub>b" +proof - + obtain z' and b' and c' where "\ = \ z' : b' | c' \ " using obtain_fresh_z by metis + moreover hence "(b_of (\[bv::=b]\<^sub>\\<^sub>b)) = b_of \ z' : b'[bv::=b]\<^sub>b\<^sub>b | c' \" using subst_tb.simps by simp + ultimately show ?thesis using subst_tv.simps subst_tb.simps by simp +qed + +lemmas wf_intros = wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.intros wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.intros + +lemmas freshers = fresh_prodN b.fresh c.fresh v.fresh ce.fresh fresh_GCons fresh_GNil fresh_at_base + +section \Strong Elimination\ + +text \Inversion/elimination for well-formed polymorphic constructors \ +lemma wf_strong_elim: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" + and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s and tm::"'a::fs" + and cs::branch_s and css::branch_list and \::\ + shows "\; \; \ \\<^sub>w\<^sub>f (V_consp tyid dc b v) : b'' \ (\ bv dclist x b' c. b'' = B_app tyid b \ + AF_typedef_poly tyid bv dclist \ set \ \ + (dc, \ x : b' | c \) \ set dclist \ + \; \ \\<^sub>w\<^sub>f b \ atom bv \ (\, \, \, b, v) \ \; \; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ atom bv \ tm)" and + "\; \; \ \\<^sub>w\<^sub>f c \ True" and + "\; \ \\<^sub>w\<^sub>f \ \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ True" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \True" and + "\; \ \\<^sub>w\<^sub>f b \ True " and + "\; \; \ \\<^sub>w\<^sub>f ce : b' \ True" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + "V_consp tyid dc b v" b'' and c and \ and \ and ts and \ and b and b' and td + avoiding: tm + +rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_conspI bv dclist \ x b' c \ \) + then show ?case by force +qed(auto+) + +section \Context Extension\ + +definition wfExt :: "\ \ \ \ \ \ \ \ bool" (" _ ; _ \\<^sub>w\<^sub>f _ < _ " [50,50,50] 50) where + "wfExt T B G1 G2 = (wfG T B G2 \ wfG T B G1 \ toSet G1 \ toSet G2)" + +section \Context\ + +lemma wfG_cons[ms_wb]: + fixes \::\ + assumes "P; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" + shows "P; \ \\<^sub>w\<^sub>f \ \ atom z \ \ \ wfB P \ b" + using wfG_elims(2)[OF assms] by metis + +lemma wfG_cons2[ms_wb]: + fixes \::\ + assumes "P; \ \\<^sub>w\<^sub>f zbc #\<^sub>\\" + shows "P; \ \\<^sub>w\<^sub>f \" +proof - + obtain z and b and c where zbc: "zbc=(z,b,c)" using prod_cases3 by blast + hence "P; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" using assms by auto + thus ?thesis using zbc wfG_cons assms by simp +qed + +lemma wf_g_unique: + fixes \::\ + assumes "\; \ \\<^sub>w\<^sub>f \" and "(x,b,c) \ toSet \" and "(x,b',c') \ toSet \" + shows "b=b' \ c=c'" +using assms proof(induct \ rule: \.induct) + case GNil + then show ?case by simp +next + case (GCons a \) + consider "(x,b,c)=a \ (x,b',c')=a" | "(x,b,c)=a \ (x,b',c')\a" | "(x,b,c)\a \ (x,b',c')=a" | "(x,b,c)\ a \ (x,b',c')\a" by blast + then show ?case proof(cases) + case 1 + then show ?thesis by auto + next + case 2 + hence "atom x \ \" using wfG_elims(2) GCons by blast + moreover have "(x,b',c') \ toSet \" using GCons 2 by force + ultimately show ?thesis using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \.distinct subst_gv.simps 2 GCons by metis + next + case 3 + hence "atom x \ \" using wfG_elims(2) GCons by blast + moreover have "(x,b,c) \ toSet \" using GCons 3 by force + ultimately show ?thesis + using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \.distinct subst_gv.simps 3 GCons by metis + next + case 4 + then obtain x'' and b'' and c''::c where xbc: "a=(x'',b'',c'')" + using prod_cases3 by blast + hence "\; \ \\<^sub>w\<^sub>f ((x'',b'',c'') #\<^sub>\\)" using GCons wfG_elims by blast + hence "\; \ \\<^sub>w\<^sub>f \ \ (x, b, c) \ toSet \ \ (x, b', c') \ toSet \" using GCons wfG_elims 4 xbc + prod_cases3 set_GConsD using forget_subst_gv fresh_GCons fresh_GNil fresh_gamma_elem \.distinct subst_gv.simps 4 GCons by meson + thus ?thesis using GCons by auto + qed +qed + +lemma lookup_if1: + fixes \::\ + assumes "\; \ \\<^sub>w\<^sub>f \" and "Some (b,c) = lookup \ x" + shows "(x,b,c) \ toSet \ \ (\b' c'. (x,b',c') \ toSet \ \ b'=b \ c'=c)" +using assms proof(induct \ rule: \.induct) + case GNil + then show ?case by auto +next + case (GCons xbc \) + then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')" + using prod_cases3 by blast + then show ?case using wf_g_unique GCons lookup_in_g xbc + lookup.simps set_GConsD wfG.cases + insertE insert_is_Un toSet.simps wfG_elims by metis +qed + +lemma lookup_if2: + assumes "wfG P \ \" and "(x,b,c) \ toSet \ \ (\b' c'. (x,b',c') \ toSet \ \ b'=b \ c'=c)" + shows "Some (b,c) = lookup \ x" +using assms proof(induct \ rule: \.induct) + case GNil + then show ?case by auto +next + case (GCons xbc \) + then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')" + using prod_cases3 by blast + then show ?case proof(cases "x=x'") + case True + then show ?thesis using lookup.simps GCons xbc by simp + next + case False + then show ?thesis using lookup.simps GCons xbc toSet.simps Un_iff set_GConsD wfG_cons2 + by (metis (full_types) Un_iff set_GConsD toSet.simps(2) wfG_cons2) + qed +qed + +lemma lookup_iff: + fixes \::\ and \::\ + assumes "\; \ \\<^sub>w\<^sub>f \" + shows "Some (b,c) = lookup \ x \ (x,b,c) \ toSet \ \ (\b' c'. (x,b',c') \ toSet \ \ b'=b \ c'=c)" + using assms lookup_if1 lookup_if2 by meson + +lemma wfG_lookup_wf: + fixes \::\ and \::\ and b::b and \::\ + assumes "\; \ \\<^sub>w\<^sub>f \" and "Some (b,c) = lookup \ x" + shows "\; \ \\<^sub>w\<^sub>f b" +using assms proof(induct \ rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + then show ?case proof(cases "x=x'") + case True + then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce + next + case False + then show ?thesis using lookup.simps wfG_elims(2) GCons by fastforce + qed +qed + +lemma wfG_unique: + fixes \::\ + assumes "wfG B \ ((x, b, c) #\<^sub>\ \)" and "(x1,b1,c1) \ toSet ((x, b, c) #\<^sub>\ \)" and "x1=x" + shows "b1 = b \ c1 = c" +proof - + have "(x, b, c) \ toSet ((x, b, c) #\<^sub>\ \)" by simp + thus ?thesis using wf_g_unique assms by blast +qed + +lemma wfG_unique_full: + fixes \::\ + assumes "wfG \ B (\'@(x, b, c) #\<^sub>\ \)" and "(x1,b1,c1) \ toSet (\'@(x, b, c) #\<^sub>\ \)" and "x1=x" + shows "b1 = b \ c1 = c" +proof - + have "(x, b, c) \ toSet (\'@(x, b, c) #\<^sub>\ \)" by simp + thus ?thesis using wf_g_unique assms by blast +qed + +section \Converting between wb forms\ + +text \ We cannot prove wfB properties here for expressions and statements as need some more facts about @{term \} + context which we can prove without this lemma. Trying to cram everything into a single large + mutually recursive lemma is not a good idea \ + +lemma wfX_wfY1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s + and css::branch_list + shows wfV_wf: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and + wfC_wf: "\; \; \ \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and + wfG_wf :"\; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and + wfT_wf: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f b_of \" and + wfTs_wf:"\; \; \ \\<^sub>w\<^sub>f ts \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and + "\\<^sub>w\<^sub>f \ \ True" and + wfB_wf: "\; \ \\<^sub>w\<^sub>f b \ \\<^sub>w\<^sub>f \" and + wfCE_wf: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and + wfTD_wf: "\ \\<^sub>w\<^sub>f td \ \\<^sub>w\<^sub>f \" +proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts) + + case (wfV_varI \ \ \ b c x) + hence "(x,b,c) \ toSet \" using lookup_iff lookup_in_g by presburger + hence "b \ fst`snd`toSet \" by force + hence "wfB \ \ b" using wfV_varI using wfG_lookup_wf by auto + then show ?case using wfV_varI wfV_elims wf_intros by metis +next + case (wfV_litI \ \ \ l) + moreover have "wfTh \" using wfV_litI by metis + ultimately show ?case using wf_intros base_for_lit.simps l.exhaust by metis +next + case (wfV_pairI \ \ \ v1 b1 v2 b2) + then show ?case using wfB_pairI by simp +next + case (wfV_consI s dclist \ dc x b c \ \ v) + then show ?case using wf_intros by metis +next + case (wfTI z \ \ \ b c) + then show ?case using wf_intros b_of.simps wfG_cons2 by metis +qed(auto) + +lemma wfX_wfY2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s + and css::branch_list + shows + wfE_wf: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and + wfS_wf: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and + "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and + "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " and + wfPhi_wf: "\ \\<^sub>w\<^sub>f (\::\) \ \\<^sub>w\<^sub>f \" and + wfD_wf: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \ " and + wfFTQ_wf: "\ ; \ \\<^sub>w\<^sub>f ftq \ \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" and + wfFT_wf: "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \" +proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts) + case (wfS_varI \ \ \ \ v u \ \ s b) + then show ?case using wfD_elims by auto +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wf_intros by metis +next + case (wfD_emptyI \ \ \) + then show ?case using wfX_wfY1 by auto +next + case (wfS_assertI \ \ \ x c \ \ s b) + then have "\; \ \\<^sub>w\<^sub>f \" using wfX_wfY1 by auto + moreover have "\; \; \ \\<^sub>w\<^sub>f \" using wfS_assertI by auto + moreover have "\\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ " using wfS_assertI by auto + ultimately show ?case by auto +qed(auto) + +lemmas wfX_wfY=wfX_wfY1 wfX_wfY2 + +lemma setD_ConsD: + "ut \ setD (ut' #\<^sub>\ D) = (ut = ut' \ ut \ setD D)" +proof(induct D rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' x2) + then show ?case using setD.simps by auto +qed + +lemma wfD_wfT: + fixes \::\ and \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" + shows "\(u,\) \ setD \. \; \; \ \\<^sub>w\<^sub>f \" +using assms proof(induct \ rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' x2) + then show ?case using wfD_elims DCons setD_ConsD + by (metis case_prodI2 set_ConsD) +qed + +lemma subst_b_lookup_d: + assumes "u \ fst ` setD \" + shows "u \ fst ` setD \[bv::=b]\<^sub>\\<^sub>b" +using assms proof(induct \ rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' x2) + hence "u\u'" using DCons by simp + show ?case using DCons subst_db.simps by simp +qed + +lemma wfG_cons_splitI: + fixes \::\ and \::\ + assumes "\; \ \\<^sub>w\<^sub>f \" and "atom x \ \" and "wfB \ \ b" and + "c \ { TRUE, FALSE } \ \; \ \\<^sub>w\<^sub>f \ " and + "c \ { TRUE, FALSE } \ \ ;\ ; (x,b,C_true) #\<^sub>\\ \\<^sub>w\<^sub>f c" + shows "\; \ \\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\\)" + using wfG_cons1I wfG_cons2I assms by metis + +lemma wfG_consI: + fixes \::\ and \::\ and c::c + assumes "\; \ \\<^sub>w\<^sub>f \" and "atom x \ \" and "wfB \ \ b" and + "\ ; \ ; (x,b,C_true) #\<^sub>\\ \\<^sub>w\<^sub>f c" + shows "\ ; \ \\<^sub>w\<^sub>f ((x,b,c) #\<^sub>\\)" + using wfG_cons1I wfG_cons2I wfG_cons_splitI wfC_trueI assms by metis + +lemma wfG_elim2: + fixes c::c + assumes "wfG P \ ((x,b,c) #\<^sub>\\)" + shows "P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c \ wfB P \ b" +proof(cases "c \ {TRUE,FALSE}") + case True + have "P; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ wfB P \ b" using wfG_elims(2)[OF assms] by auto + hence "P; \ \\<^sub>w\<^sub>f ((x,b,TRUE) #\<^sub>\\) \ wfB P \ b" using wfG_cons2I by auto + thus ?thesis using wfC_trueI wfC_falseI True by auto +next + case False + then show ?thesis using wfG_elims(2)[OF assms] by auto +qed + +lemma wfG_cons_wfC: + fixes \::\ and c::c + assumes "\ ; B \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" + shows "\ ; B ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f c" + using assms wfG_elim2 by auto + +lemma wfG_wfB: + assumes "wfG P \ \" and "b \ fst`snd`toSet \" + shows "wfB P \ b" +using assms proof(induct \ rule:\_induct) +case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + show ?case proof(cases "b=b'") + case True + then show ?thesis using wfG_elim2 GCons by auto + next + case False + hence "b \ fst`snd`toSet \'" using GCons by auto + moreover have "wfG P \ \'" using wfG_cons GCons by auto + ultimately show ?thesis using GCons by auto + qed +qed + +lemma wfG_cons_TRUE: + fixes \::\ and b::b + assumes "P; \ \\<^sub>w\<^sub>f \" and "atom z \ \" and "P; \ \\<^sub>w\<^sub>f b" + shows "P ; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \" + using wfG_cons2I wfG_wfB assms by simp + +lemma wfG_cons_TRUE2: + assumes "P; \ \\<^sub>w\<^sub>f (z,b,c) #\<^sub>\\" and "atom z \ \" + shows "P; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \" + using wfG_cons wfG_cons2I assms by simp + +lemma wfG_suffix: + fixes \::\ + assumes "wfG P \ (\'@\)" + shows "wfG P \ \" +using assms proof(induct \' rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x b c \') + hence " P; \ \\<^sub>w\<^sub>f \' @ \" using wfG_elims by auto + then show ?case using GCons wfG_elims by auto +qed + +lemma wfV_wfCE: + fixes v::v + assumes "\; \; \ \\<^sub>w\<^sub>f v : b" + shows " \ ; \ ; \ \\<^sub>w\<^sub>f CE_val v : b" +proof - + have "\ \\<^sub>w\<^sub>f ([]::\) " using wfPhi_emptyI wfV_wf wfG_wf assms by metis + moreover have "\; \; \ \\<^sub>w\<^sub>f []\<^sub>\" using wfD_emptyI wfV_wf wfG_wf assms by metis + ultimately show ?thesis using wfCE_valI assms by auto +qed + +section \Support\ + +lemma wf_supp1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list + + shows wfV_supp: "\; \; \ \\<^sub>w\<^sub>f v : b \ supp v \ atom_dom \ \ supp \" and + wfC_supp: "\; \; \ \\<^sub>w\<^sub>f c \ supp c \ atom_dom \ \ supp \" and + wfG_supp: "\; \ \\<^sub>w\<^sub>f \ \ atom_dom \ \ supp \" and + wfT_supp: "\; \; \ \\<^sub>w\<^sub>f \ \ supp \ \ atom_dom \ \ supp \ " and + wfTs_supp: "\; \; \ \\<^sub>w\<^sub>f ts \ supp ts \ atom_dom \ \ supp \" and + wfTh_supp: "\\<^sub>w\<^sub>f \ \ supp \ = {}" and + wfB_supp: "\; \ \\<^sub>w\<^sub>f b \ supp b \ supp \" and + wfCE_supp: "\; \; \ \\<^sub>w\<^sub>f ce : b \ supp ce \ atom_dom \ \ supp \" and + wfTD_supp: "\ \\<^sub>w\<^sub>f td \ supp td \ {}" +proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts) + case (wfB_consI \ s dclist \) + then show ?case by(auto simp add: b.supp pure_supp) +next + case (wfB_appI \ \ b s bv dclist) + then show ?case by(auto simp add: b.supp pure_supp) +next + case (wfV_varI \ \ \ b c x) + then show ?case using v.supp wfV_elims + empty_subsetI insert_subset supp_at_base + fresh_dom_free2 lookup_if1 + by (metis sup.coboundedI1) +next + case (wfV_litI \ \ \ l) + then show ?case using supp_l_empty v.supp by simp +next + case (wfV_pairI \ \ \ v1 b1 v2 b2) + then show ?case using v.supp wfV_elims by (metis Un_subset_iff) +next + case (wfV_consI s dclist \ dc x b c \ \ v) + then show ?case using v.supp wfV_elims + Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp by metis +next + case (wfV_conspI typid bv dclist \ dc x b' c \ \ v b) + then show ?case unfolding v.supp + using wfV_elims + Un_commute b.supp sup_bot.right_neutral supp_b_empty pure_supp + by (simp add: Un_commute pure_supp sup.coboundedI1) +next + case (wfC_eqI \ \ \ e1 b e2) + hence "supp e1 \ atom_dom \ \ supp \" using c.supp wfC_elims + image_empty list.set(1) sup_bot.right_neutral by (metis IntI UnE empty_iff subsetCE subsetI) + moreover have "supp e2 \ atom_dom \ \ supp \" using c.supp wfC_elims + image_empty list.set(1) sup_bot.right_neutral IntI UnE empty_iff subsetCE subsetI + by (metis wfC_eqI.hyps(4)) + ultimately show ?case using c.supp by auto +next + case (wfG_cons1I c \ \ \ x b) + then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis +next + case (wfG_cons2I c \ \ \ x b) + then show ?case using atom_dom.simps dom_supp_g supp_GCons by metis +next + case wfTh_emptyI + then show ?case by (simp add: supp_Nil) +next + case (wfTh_consI \ lst) + then show ?case using supp_Cons by fast +next + case (wfTD_simpleI \ lst s) + then have "supp (AF_typedef s lst ) = supp lst \ supp s" using type_def.supp by auto + then show ?case using wfTD_simpleI pure_supp + by (simp add: pure_supp supp_Cons supp_at_base) +next + case (wfTD_poly \ bv lst s) + then have "supp (AF_typedef_poly s bv lst ) = supp lst - { atom bv } \ supp s" using type_def.supp by auto + then show ?case using wfTD_poly pure_supp + by (simp add: pure_supp supp_Cons supp_at_base) +next + case (wfTs_nil \ \ \) + then show ?case using supp_Nil by auto +next + case (wfTs_cons \ \ \ \ dc ts) + then show ?case using supp_Cons supp_Pair pure_supp[of dc] by blast +next + case (wfCE_valI \ \ \ v b) + thus ?case using ce.supp wfCE_elims by simp +next + case (wfCE_plusI \ \ \ v1 v2) + hence "supp (CE_op Plus v1 v2) \ atom_dom \ \ supp \" using ce.supp pure_supp + by (simp add: wfCE_plusI opp.supp) + then show ?case using ce.supp wfCE_elims UnCI subsetCE subsetI x_not_in_b_set by auto +next + case (wfCE_leqI \ \ \ v1 v2) + hence "supp (CE_op LEq v1 v2) \ atom_dom \ \ supp \" using ce.supp pure_supp + by (simp add: wfCE_plusI opp.supp) + then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto +next + case (wfCE_eqI \ \ \ v1 b v2 ) + hence "supp (CE_op Eq v1 v2) \ atom_dom \ \ supp \" using ce.supp pure_supp + by (simp add: wfCE_eqI opp.supp) + then show ?case using ce.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by auto +next + case (wfCE_fstI \ \ \ v1 b1 b2) + thus ?case using ce.supp wfCE_elims by simp +next + case (wfCE_sndI \ \ \ v1 b1 b2) + thus ?case using ce.supp wfCE_elims by simp +next + case (wfCE_concatI \ \ \ v1 v2) + thus ?case using ce.supp wfCE_elims by simp +next + case (wfCE_lenI \ \ \ v1) + thus ?case using ce.supp wfCE_elims by simp +next + case (wfTI z \ \ \ b c) + hence "supp c \ supp z \ atom_dom \ \ supp \" using supp_at_base dom_cons by metis + moreover have "supp b \ supp \" using wfTI by auto + ultimately have " supp \ z : b | c \ \ atom_dom \ \ supp \" using \.supp supp_at_base by force + thus ?case by auto +qed(auto) + +lemma wf_supp2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and + ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and + ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css ::branch_list + shows + wfE_supp: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ (supp e \ atom_dom \ \ supp \ \ atom ` fst ` setD \)" and (* \ ( \ = [] \ supp e \ supp \ = {})" and*) + wfS_supp: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ supp s \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and + "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ supp cs \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and + "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ supp css \ atom_dom \ \ atom ` fst ` setD \ \ supp \" and + wfPhi_supp: "\ \\<^sub>w\<^sub>f (\::\) \ supp \ = {}" and + wfD_supp: "\; \; \ \\<^sub>w\<^sub>f \ \ supp \ \ atom`fst`(setD \) \ atom_dom \ \ supp \ " and + "\ ; \ \\<^sub>w\<^sub>f ftq \ supp ftq = {}" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ supp ft \ supp \" +proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts) + case (wfE_valI \ \ \ \ \ v b) + hence "supp (AE_val v) \ atom_dom \ \ supp \" using e.supp wf_supp1 by simp + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + hence "supp (AE_op Plus v1 v2) \ atom_dom \ \ supp \" + using wfE_plusI opp.supp wf_supp1 e.supp pure_supp Un_least + by (metis sup_bot.left_neutral) + + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + hence "supp (AE_op LEq v1 v2) \ atom_dom \ \ supp \" using e.supp pure_supp Un_least + sup_bot.left_neutral using opp.supp wf_supp1 by auto + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + hence "supp (AE_op Eq v1 v2) \ atom_dom \ \ supp \" using e.supp pure_supp Un_least + sup_bot.left_neutral using opp.supp wf_supp1 by auto + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + hence "supp (AE_fst v1 ) \ atom_dom \ \ supp \" using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + hence "supp (AE_snd v1 ) \ atom_dom \ \ supp \" using e.supp pure_supp wfE_plusI opp.supp wf_supp1 by (metis Un_least) + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + hence "supp (AE_concat v1 v2) \ atom_dom \ \ supp \" using e.supp pure_supp + wfE_plusI opp.supp wf_supp1 by (metis Un_least) + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + hence "supp (AE_split v1 v2) \ atom_dom \ \ supp \" using e.supp pure_supp + wfE_plusI opp.supp wf_supp1 by (metis Un_least) + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_lenI \ \ \ \ \ v1) + hence "supp (AE_len v1 ) \ atom_dom \ \ supp \" using e.supp pure_supp + using e.supp pure_supp sup_bot.left_neutral using opp.supp wf_supp1 by auto + then show ?case using e.supp wfE_elims UnCI subsetCE subsetI x_not_in_b_set by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then obtain b where "\; \; \ \\<^sub>w\<^sub>f v : b" using wfE_elims by metis + hence "supp v \ atom_dom \ \ supp \" using wfE_appI wf_supp1 by metis + hence "supp (AE_app f v) \ atom_dom \ \ supp \" using e.supp pure_supp by fast + then show ?case using e.supp(2) UnCI subsetCE subsetI wfE_appI using b.supp(3) pure_supp x_not_in_b_set by metis +next + case (wfE_appPI \ \ \ \ \ b' bv v \ f xa ba ca s) + then obtain b where "\; \; \ \\<^sub>w\<^sub>f v : ( b[bv::=b']\<^sub>b)" using wfE_elims by metis + hence "supp v \ atom_dom \ \ supp \ " using wfE_appPI wf_supp1 by auto + moreover have "supp b' \ supp \" using wf_supp1(7) wfE_appPI by simp + ultimately show ?case unfolding e.supp using wfE_appPI pure_supp by fast +next + case (wfE_mvarI \ \ \ \ \ u \) + then obtain \ where "(u,\) \ setD \" using wfE_elims(10) by metis + hence "atom u \ atom`fst`setD \" by force + hence "supp (AE_mvar u ) \ atom`fst`setD \" using e.supp + by (simp add: supp_at_base) + thus ?case using UnCI subsetCE subsetI e.supp wfE_mvarI supp_at_base subsetCE supp_at_base u_not_in_b_set + by (simp add: supp_at_base) +next + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_supp1 + by (metis s_branch_s_branch_list.supp(1) sup.coboundedI2 sup_assoc sup_commute) +next + case (wfS_letI \ \ \ \ \ e b' x s b) + then show ?case by auto +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + then show ?case unfolding s_branch_s_branch_list.supp (3) using wf_supp1(4)[OF wfS_let2I(3)] by auto +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wf_supp1(1)[OF wfS_ifI(1)] by auto +next + case (wfS_varI \ \ \ \ v u \ \ s b) + then show ?case using wf_supp1(1)[OF wfS_varI(2)] wf_supp1(4)[OF wfS_varI(1)] by auto +next +next + case (wfS_assignI u \ \ \ \ \ \ v) + hence "supp u \ atom ` fst ` setD \" proof(induct \ rule:\_induct) + case DNil + then show ?case by auto + next + case (DCons u' t' \') + show ?case proof(cases "u=u'") + case True + then show ?thesis using toSet.simps DCons supp_at_base by fastforce + next + case False + then show ?thesis using toSet.simps DCons supp_at_base wfS_assignI + by (metis empty_subsetI fstI image_eqI insert_subset) + qed + qed + then show ?case using s_branch_s_branch_list.supp(8) wfS_assignI wf_supp1(1)[OF wfS_assignI(6)] by auto +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + then show ?case using wf_supp1(1)[OF wfS_matchI(1)] by auto +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + moreover have "supp s \ supp x \ atom_dom \ \ atom ` fst ` setD \ \ supp \" + using dom_cons supp_at_base wfS_branchI by auto + moreover hence "supp s - set [atom x] \ atom_dom \ \ atom ` fst ` setD \ \ supp \" using supp_at_base by force + ultimately have + "(supp s - set [atom x]) \ (supp dc ) \ atom_dom \ \ atom ` fst ` setD \ \ supp \" + by (simp add: pure_supp) + thus ?case using s_branch_s_branch_list.supp(2) by auto +next + case (wfD_emptyI \ \ \) + then show ?case using supp_DNil by auto +next + case (wfD_cons \ \ \ \ \ u) + have "supp ((u, \) #\<^sub>\ \) = supp u \ supp \ \ supp \" using supp_DCons supp_Pair by metis + also have "... \ supp u \ atom ` fst ` setD \ \ atom_dom \ \ supp \" + using wfD_cons wf_supp1(4)[OF wfD_cons(3)] by auto + also have "... \ atom ` fst ` setD ((u, \) #\<^sub>\ \) \ atom_dom \ \ supp \" using supp_at_base by auto + finally show ?case by auto +next + case (wfPhi_emptyI \) + then show ?case using supp_Nil by auto +next + case (wfPhi_consI f \ \ ft) + then show ?case using fun_def.supp + by (simp add: pure_supp supp_Cons) +next + case (wfFTI \ B' b s x c \ \) + have " supp (AF_fun_typ x b c \ s) = supp c \ (supp \ \ supp s) - set [atom x] \ supp b" using fun_typ.supp by auto + thus ?case using wfFTI wf_supp1 + proof - + have f1: "supp \ \ {atom x} \ atom_dom GNil \ supp B'" + using dom_cons wfFTI.hyps wf_supp1(4) by blast (* 0.0 ms *) + have "supp b \ supp B'" + using wfFTI.hyps(1) wf_supp1(7) by blast (* 0.0 ms *) + then show ?thesis + using f1 \supp (AF_fun_typ x b c \ s) = supp c \ (supp \ \ supp s) - set [atom x] \ supp b\ + wfFTI.hyps(4) wfFTI.hyps by auto (* 234 ms *) + qed +next + case (wfFTNone \ \ ft) + then show ?case by (simp add: fun_typ_q.supp(2)) +next + case (wfFTSome \ \ bv ft) + then show ?case using fun_typ_q.supp + by (simp add: supp_at_base) +next + case (wfS_assertI \ \ \ x c \ \ s b) + then have "supp c \ atom_dom \ \ atom ` fst ` setD \ \ supp \" using wf_supp1 + by (metis Un_assoc Un_commute le_supI2) + moreover have "supp s \ atom_dom \ \ atom ` fst ` setD \ \ supp \" proof + fix z + assume *:"z \ supp s" + have **:"atom x \ supp s" using wfS_assertI fresh_prodN fresh_def by metis + have "z \ atom_dom ((x, B_bool, c) #\<^sub>\ \) \ atom ` fst ` setD \ \ supp \" using wfS_assertI * by blast + have "z \ atom_dom ((x, B_bool, c) #\<^sub>\ \) \ z \ atom_dom \" using * ** by auto + thus "z \ atom_dom \ \ atom ` fst ` setD \ \ supp \" using * ** + using \z \ atom_dom ((x, B_bool, c) #\<^sub>\ \) \ atom ` fst ` setD \ \ supp \\ by blast + qed + ultimately show ?case by auto +qed(auto) + +lemmas wf_supp = wf_supp1 wf_supp2 + +lemma wfV_supp_nil: + fixes v::v + assumes "P ; {||} ; GNil \\<^sub>w\<^sub>f v : b" + shows "supp v = {}" + using wfV_supp[of P " {||}" GNil v b] dom.simps toSet.simps + using assms by auto + +lemma wfT_TRUE_aux: + assumes "wfG P \ \" and "atom z \ (P, \, \)" and "wfB P \ b" + shows "wfT P \ \ (\ z : b | TRUE \)" +proof (rule) + show \ atom z \ (P, \, \)\ using assms by auto + show \ P; \ \\<^sub>w\<^sub>f b \ using assms by auto + show \ P ;\ ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f TRUE \ using wfG_cons2I wfC_trueI assms by auto +qed + +lemma wfT_TRUE: + assumes "wfG P \ \" and "wfB P \ b" + shows "wfT P \ \ (\ z : b | TRUE \)" +proof - + obtain z'::x where *:"atom z' \ (P, \, \)" using obtain_fresh by metis + hence "\ z : b | TRUE \ = \ z' : b | TRUE \" by auto + thus ?thesis using wfT_TRUE_aux assms * by metis +qed + +lemma phi_flip_eq: + assumes "wfPhi T P" + shows "(x \ xa) \ P = P" + using wfPhi_supp[OF assms] flip_fresh_fresh fresh_def by blast + +lemma wfC_supp_cons: + fixes c'::c and G::\ + assumes "P; \ ; (x', b' , TRUE) #\<^sub>\G \\<^sub>w\<^sub>f c'" + shows "supp c' \ atom_dom G \ supp x' \ supp \" and "supp c' \ supp G \ supp x' \ supp \" +proof - + show "supp c' \ atom_dom G \ supp x' \ supp \" + using wfC_supp[OF assms] dom_cons supp_at_base by blast + moreover have "atom_dom G \ supp G" + by (meson assms wfC_wf wfG_cons wfG_supp) + ultimately show "supp c' \ supp G \ supp x' \ supp \" using wfG_supp assms wfG_cons wfC_wf by fast +qed + +lemma wfG_dom_supp: + fixes x::x + assumes "wfG P \ G" + shows "atom x \ atom_dom G \ atom x \ supp G" +using assms proof(induct G rule: \_induct) + case GNil + then show ?case using dom.simps supp_of_atom_list + using supp_GNil by auto +next + case (GCons x' b' c' G) + + show ?case proof(cases "x' = x") + case True + then show ?thesis using dom.simps supp_of_atom_list supp_at_base + using supp_GCons by auto + next + case False + have "(atom x \ atom_dom ((x', b', c') #\<^sub>\ G)) = (atom x \ atom_dom G)" using atom_dom.simps False by simp + also have "... = (atom x \ supp G)" using GCons wfG_elims by metis + also have "... = (atom x \ (supp (x', b', c') \ supp G))" proof + show "atom x \ supp G \ atom x \ supp (x', b', c') \ supp G" by auto + assume "atom x \ supp (x', b', c') \ supp G" + then consider "atom x \ supp (x', b', c')" | "atom x \ supp G" by auto + then show "atom x \ supp G" proof(cases) + case 1 + assume " atom x \ supp (x', b', c') " + hence "atom x \ supp c'" using supp_triple False supp_b_empty supp_at_base by force + + moreover have "P; \ ; (x', b' , TRUE) #\<^sub>\G \\<^sub>w\<^sub>f c'" using wfG_elim2 GCons by simp + moreover hence "supp c' \ supp G \ supp x' \ supp \" using wfC_supp_cons by auto + ultimately have "atom x \ supp G \ supp x' " using x_not_in_b_set by auto + then show ?thesis using False supp_at_base by (simp add: supp_at_base) + next + case 2 + then show ?thesis by simp + qed + qed + also have "... = (atom x \ supp ((x', b', c') #\<^sub>\ G))" using supp_at_base False supp_GCons by simp + finally show ?thesis by simp + qed +qed + +lemma wfG_atoms_supp_eq : + fixes x::x + assumes "wfG P \ G" + shows "atom x \ atom_dom G \ atom x \ supp G" + using wfG_dom_supp assms by auto + +lemma beta_flip_eq: + fixes x::x and xa::x and \::\ + shows "(x \ xa) \ \ = \" +proof - + have "atom x \ \ \ atom xa \ \" using x_not_in_b_set fresh_def supp_set by metis + thus ?thesis by (simp add: flip_fresh_fresh fresh_def) +qed + +lemma theta_flip_eq2: + assumes "\\<^sub>w\<^sub>f \" + shows " (z \ za ) \ \ = \" +proof - + have "supp \ = {}" using wfTh_supp assms by simp + thus ?thesis + by (simp add: flip_fresh_fresh fresh_def) + qed + +lemma theta_flip_eq: + assumes "wfTh \" + shows "(x \ xa) \ \ = \" + using wfTh_supp flip_fresh_fresh fresh_def + by (simp add: assms theta_flip_eq2) + +lemma wfT_wfC: + fixes c::c + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom z \ \" + shows "\; \; (z,b,TRUE) #\<^sub>\\ \\<^sub>w\<^sub>f c" +proof - + obtain za ba ca where *:"\ z : b | c \ = \ za : ba | ca \ \ atom za \ (\,\,\) \ \; \; (za, ba, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f ca" + using wfT_elims[OF assms(1)] by metis + hence c1: "[[atom z]]lst. c = [[atom za]]lst. ca" using \.eq_iff by meson + show ?thesis proof(cases "z=za") + case True + hence "ca = c" using c1 by (simp add: Abs1_eq_iff(3)) + then show ?thesis using * True by simp + next + case False + have " \\<^sub>w\<^sub>f \" using wfT_wf wfG_wf assms by metis + moreover have "atom za \ \" using * fresh_prodN by auto + ultimately have "\; \; (z \ za ) \ (za, ba, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f (z \ za ) \ ca" + using wfC.eqvt theta_flip_eq2 beta_flip_eq * GCons_eqvt assms flip_fresh_fresh by metis + moreover have "atom z \ ca" + proof - + have "supp ca \ atom_dom \ \ { atom za } \ supp \" using * wfC_supp atom_dom.simps toSet.simps by fastforce + moreover have "atom z \ atom_dom \ " using assms fresh_def wfT_wf wfG_dom_supp wfC_supp by metis + moreover hence "atom z \ atom_dom \ \ { atom za }" using False by simp + moreover have "atom z \ supp \" using x_not_in_b_set by simp + ultimately show ?thesis using fresh_def False by fast + qed + moreover hence "(z \ za ) \ ca = c" using type_eq_subst_eq1(3) * by metis + ultimately show ?thesis using assms G_cons_flip_fresh * by auto + qed +qed + +lemma u_not_in_dom_g: + fixes u::u + shows "atom u \ atom_dom G" + using toSet.simps atom_dom.simps u_not_in_x_atoms by auto + +lemma bv_not_in_dom_g: + fixes bv::bv + shows "atom bv \ atom_dom G" + using toSet.simps atom_dom.simps u_not_in_x_atoms by auto + +text \An important lemma that confirms that @{term \} does not rely on mutable variables\ +lemma u_not_in_g: + fixes u::u + assumes "wfG \ B G" + shows "atom u \ supp G" +using assms proof(induct G rule: \_induct) +case GNil + then show ?case using supp_GNil fresh_def + using fresh_set_empty by fastforce +next + case (GCons x b c \') + moreover hence "atom u \ supp b" using + wfB_supp wfC_supp u_not_in_x_atoms wfG_elims wfX_wfY by auto + moreover hence "atom u \ supp x" using u_not_in_x_atoms supp_at_base by blast + moreover hence "atom u \ supp c" proof - + have "\ ; B ; (x, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c" using wfG_cons_wfC GCons by simp + hence "supp c \ atom_dom ((x, b, TRUE) #\<^sub>\ \') \ supp B" using wfC_supp by blast + thus ?thesis using u_not_in_dom_g u_not_in_b_atoms + using u_not_in_b_set by auto + qed + ultimately have "atom u \ supp (x,b,c)" using supp_Pair by simp + thus ?case using supp_GCons GCons wfG_elims by blast +qed + +text \An important lemma that confirms that types only depend on immutable variables\ +lemma u_not_in_t: + fixes u::u + assumes "wfT \ B G \" + shows "atom u \ supp \" +proof - + have "supp \ \ atom_dom G \ supp B" using wfT_supp assms by auto + thus ?thesis using u_not_in_dom_g u_not_in_b_set by blast +qed + +lemma wfT_supp_c: + fixes \::\ and z::x + assumes "wfT P \ \ (\ z : b | c \)" + shows "supp c - { atom z } \ atom_dom \ \ supp \" + using wf_supp \.supp assms + by (metis Un_subset_iff empty_set list.simps(15)) + +lemma wfG_wfC[ms_wb]: + assumes "wfG P \ ((x,b,c) #\<^sub>\\)" + shows "wfC P \ ((x,b,TRUE) #\<^sub>\\) c" +using assms proof(cases "c \ {TRUE,FALSE}") + case True + have "atom x \ \ \ wfG P \ \ \ wfB P \ b" using wfG_cons assms by auto + hence "wfG P \ ((x,b,TRUE) #\<^sub>\\)" using wfG_cons2I by auto + then show ?thesis using wfC_trueI wfC_falseI True by auto +next + case False + then show ?thesis using wfG_elims assms by blast +qed + +lemma wfT_wf_cons: + assumes "wfT P \ \ \ z : b | c \" and "atom z \ \" + shows "wfG P \ ((z,b,c) #\<^sub>\\)" +using assms proof(cases "c \ { TRUE,FALSE }") + case True + then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons2I assms wfT_wf by fastforce +next + case False + then show ?thesis using wfT_wfC wfC_wf wfG_wfB wfG_cons1I wfT_wf wfT_wfC assms by fastforce +qed + +lemma wfV_b_fresh: + fixes b::b and v::v and bv::bv + assumes "\; \; \ \\<^sub>w\<^sub>f v: b" and "bv |\| \" + shows "atom bv \ v" +using wfV_supp bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp by blast + +lemma wfCE_b_fresh: + fixes b::b and ce::ce and bv::bv + assumes "\; \; \ \\<^sub>w\<^sub>f ce: b" and "bv |\| \" + shows "atom bv \ ce" +using bv_not_in_dom_g fresh_def assms bv_not_in_bset_supp wf_supp1(8) by fast + +section \Freshness\ + +lemma wfG_fresh_x: + fixes \::\ and z::x + assumes "\; \ \\<^sub>w\<^sub>f \" and "atom z \ \" + shows "atom z \ (\,\, \)" +unfolding fresh_prodN apply(intro conjI) + using wf_supp1 wfX_wfY assms fresh_def x_not_in_b_set by(metis empty_iff)+ + +lemma wfG_wfT: + assumes "wfG P \ ((x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ G)" and "atom x \ c" + shows "P; \ ; G \\<^sub>w\<^sub>f \ z : b | c \" +proof - + have " P; \ ; (x, b, TRUE) #\<^sub>\ G \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \ wfB P \ b" using assms + using wfG_elim2 by auto + moreover have "atom x \ (P ,\,G)" using wfG_elims assms wfG_fresh_x by metis + ultimately have "wfT P \ G \ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \" using wfTI assms by metis + moreover have "\ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \ = \ z : b | c \" using type_eq_subst \atom x \ c\ by auto + ultimately show ?thesis by auto +qed + +lemma wfT_wfT_if: + assumes "wfT \ \ \ (\ z2 : b | CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v \)" and "atom z2 \ (c,\)" + shows "wfT \ \ \ \ z : b | c \" +proof - + have *: "atom z2 \ (\, \, \)" using wfG_fresh_x wfX_wfY assms fresh_Pair by metis + have "wfB \ \ b" using assms wfT_elims by metis + have "\; \; (GCons (z2,b,TRUE) \) \\<^sub>w\<^sub>f (CE_val v == CE_val (V_lit L_false) IMP c[z::=V_var z2]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_Pair by auto + hence "\; \; ((z2,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c[z::=V_var z2]\<^sub>c\<^sub>v" using wfC_elims by metis + hence "wfT \ \ \ (\ z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v\)" using assms fresh_Pair wfTI \wfB \ \ b\ * by auto + moreover have "\ z : b | c \ = \ z2 : b | c[z::=V_var z2]\<^sub>c\<^sub>v \" using type_eq_subst assms fresh_Pair by auto + ultimately show ?thesis using wfTI assms by argo +qed + +lemma wfT_fresh_c: + fixes x::x + assumes "wfT P \ \ \ z : b | c \" and "atom x \ \" and "x \ z" + shows "atom x \ c" +proof(rule ccontr) + assume "\ atom x \ c" + hence *:"atom x \ supp c" using fresh_def by auto + moreover have "supp c - set [atom z] \ supp b \ atom_dom \ \ supp \" + using assms wfT_supp \.supp by blast + moreover hence "atom x \ supp c - set [atom z]" using assms * by auto + ultimately have "atom x \ atom_dom \" using x_not_in_b_set by auto + thus False using assms wfG_atoms_supp_eq wfT_wf fresh_def by metis +qed + +lemma wfG_x_fresh [simp]: + fixes x::x + assumes "wfG P \ G" + shows "atom x \ atom_dom G \ atom x \ G" + using wfG_atoms_supp_eq assms fresh_def by metis + +lemma wfD_x_fresh: + fixes x::x + assumes "atom x \ \" and "wfD P B \ \" + shows "atom x \ \" +using assms proof(induct \ rule: \_induct) + case DNil + then show ?case using supp_DNil fresh_def by auto +next + case (DCons u' t' \') + have wfg: "wfG P B \" using wfD_wf DCons by blast + hence wfd: "wfD P B \ \'" using wfD_elims DCons by blast + have "supp t' \ atom_dom \ \ supp B" using wfT_supp DCons wfD_elims by metis + moreover have "atom x \ atom_dom \" using DCons(2) fresh_def wfG_supp wfg by blast + ultimately have "atom x \ t'" using fresh_def DCons wfG_supp wfg x_not_in_b_set by blast + moreover have "atom x \ u'" using supp_at_base fresh_def by fastforce + ultimately have "atom x \ (u',t')" using supp_Pair by fastforce + thus ?case using DCons fresh_DCons wfd by fast +qed + +lemma wfG_fresh_x2: + fixes \::\ and z::x and \::\ and \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" and "\ \\<^sub>w\<^sub>f \" and "atom z \ \" + shows "atom z \ (\,\,\, \,\)" + unfolding fresh_prodN apply(intro conjI) + using wfG_fresh_x assms fresh_prod3 wfX_wfY apply metis + using wf_supp2(5) assms fresh_def apply blast + using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis + using assms wfG_fresh_x wfX_wfY fresh_prod3 apply metis + using wf_supp2(6) assms fresh_def wfD_x_fresh by metis + +lemma wfV_x_fresh: + fixes v::v and b::b and \::\ and x::x + assumes "\; \; \ \\<^sub>w\<^sub>f v : b" and "atom x \ \" + shows "atom x \ v" +proof - + have "supp v \ atom_dom \ \ supp \ " using assms wfV_supp by auto + moreover have "atom x \ atom_dom \" using fresh_def assms + dom.simps subsetCE wfG_elims wfG_supp by (metis dom_supp_g) + moreover have "atom x \ supp \" using x_not_in_b_set by auto + ultimately show ?thesis using fresh_def by fast +qed + +lemma wfE_x_fresh: + fixes e::e and b::b and \::\ and \::\ and \::\ and x::x + assumes "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b" and "atom x \ \" + shows "atom x \ e" +proof - + have "wfG \ \ \" using assms wfE_wf by auto + hence "supp e \ atom_dom \ \ supp \ \ atom`fst`setD \" using wfE_supp dom.simps assms by auto + moreover have "atom x \ atom_dom \" using fresh_def assms + dom.simps subsetCE \wfG \ \ \\ wfG_supp by (metis dom_supp_g) + moreover have "atom x \ atom`fst`setD \" by auto + ultimately show ?thesis using fresh_def x_not_in_b_set by fast +qed + +lemma wfT_x_fresh: + fixes \::\ and \::\ and x::x + assumes "\; \; \ \\<^sub>w\<^sub>f \" and "atom x \ \" + shows "atom x \ \" +proof - + have "wfG \ \ \" using assms wfX_wfY by auto + hence "supp \ \ atom_dom \ \ supp \" using wfT_supp dom.simps assms by auto + moreover have "atom x \ atom_dom \" using fresh_def assms + dom.simps subsetCE \wfG \ \ \\ wfG_supp by (metis dom_supp_g) + moreover have "atom x \ supp \" using x_not_in_b_set by simp + ultimately show ?thesis using fresh_def by fast +qed + +lemma wfS_x_fresh: + fixes s::s and \::\ and x::x + assumes "\; \; \; \; \ \\<^sub>w\<^sub>f s : b" and "atom x \ \" + shows "atom x \ s" +proof - + have "supp s \ atom_dom \ \ atom ` fst ` setD \ \ supp \" using wf_supp assms by metis + moreover have "atom x \ atom ` fst ` setD \" by auto + moreover have "atom x \ atom_dom \" using assms fresh_def wfG_dom_supp wfX_wfY by metis + moreover have "atom x \ supp \" using supp_b_empty supp_fset + by (simp add: x_not_in_b_set) + ultimately show ?thesis using fresh_def by fast +qed + +lemma wfTh_fresh: + fixes x + assumes "wfTh T" + shows "atom x \ T" + using wf_supp1 assms fresh_def by fastforce + +lemmas wfTh_x_fresh = wfTh_fresh + +lemma wfPhi_fresh: + fixes x + assumes "wfPhi T P" + shows "atom x \ P" + using wf_supp assms fresh_def by fastforce + +lemmas wfPhi_x_fresh = wfPhi_fresh +lemmas wb_x_fresh = wfTh_x_fresh wfPhi_x_fresh wfD_x_fresh wfT_x_fresh wfV_x_fresh + +lemma wfG_inside_fresh[ms_fresh]: + fixes \::\ and x::x + assumes "wfG P \ (\'@((x,b,c) #\<^sub>\\))" + shows "atom x \ atom_dom \'" +using assms proof(induct \' rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x1 b1 c1 \1) + moreover hence "atom x \ atom ` fst `({(x1,b1,c1)})" proof - + have *: "P; \ \\<^sub>w\<^sub>f (\1 @ (x, b, c) #\<^sub>\ \)" using wfG_elims append_g.simps GCons by metis + have "atom x1 \ (\1 @ (x, b, c) #\<^sub>\ \)" using GCons wfG_elims append_g.simps by metis + hence "atom x1 \ atom_dom (\1 @ (x, b, c) #\<^sub>\ \)" using wfG_dom_supp fresh_def * by metis + thus ?thesis by auto + qed + ultimately show ?case using append_g.simps atom_dom.simps toSet.simps wfG_elims dom.simps + by (metis image_insert insert_iff insert_is_Un) +qed + +lemma wfG_inside_x_in_atom_dom: + fixes c::c and x::x and \::\ + shows "atom x \ atom_dom ( \'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + by(induct \' rule: \_induct, (simp add: toSet.simps atom_dom.simps)+) + +lemma wfG_inside_x_neq: + fixes c::c and x::x and \::\ and G::\ and xa::x + assumes "G=( \'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and " \; \ \\<^sub>w\<^sub>f G" + shows "xa \ x" +proof - + have "atom xa \ atom_dom G" using fresh_def wfG_atoms_supp_eq assms by metis + moreover have "atom x \ atom_dom G" using wfG_inside_x_in_atom_dom assms by simp + ultimately show ?thesis by auto +qed + +lemma wfG_inside_x_fresh: + fixes c::c and x::x and \::\ and G::\ and xa::x + assumes "G=( \'@ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" and "atom xa \ G" and " \; \ \\<^sub>w\<^sub>f G" + shows "atom xa \ x" + using fresh_def supp_at_base wfG_inside_x_neq assms by auto + +lemma wfT_nil_supp: + fixes t::\ + assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f t" + shows "supp t = {}" + using wfT_supp atom_dom.simps assms toSet.simps by force + +section \Misc\ + +lemma wfG_cons_append: + fixes b'::b + assumes "\; \ \\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\ \') @ (x, b, c) #\<^sub>\ \" + shows "\; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) \ atom x' \ (\' @ (x, b, c) #\<^sub>\ \) \ \; \ \\<^sub>w\<^sub>f b' \ x' \ x" +proof - + have "((x', b', c') #\<^sub>\ \') @ (x, b, c) #\<^sub>\ \ = (x', b', c') #\<^sub>\ (\' @ (x, b, c) #\<^sub>\ \)" using append_g.simps by auto + hence *:"\; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) \ atom x' \ (\' @ (x, b, c) #\<^sub>\ \) \ \; \ \\<^sub>w\<^sub>f b'" using assms wfG_cons by metis + moreover have "atom x' \ x" proof(rule wfG_inside_x_fresh[of "(\' @ (x, b, c) #\<^sub>\ \)"]) + show "\' @ (x, b, c) #\<^sub>\ \ = \' @ (x, b, c[x::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \" by simp + show " atom x' \ \' @ (x, b, c) #\<^sub>\ \" using * by auto + show "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c) #\<^sub>\ \ " using * by auto + qed + ultimately show ?thesis by auto +qed + +lemma flip_u_eq: + fixes u::u and u'::u and \::\ and \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" + shows "(u \ u') \ \ = \" and "(u \ u') \ \ = \" and "(u \ u') \ \ = \" and "(u \ u') \ \ = \" +proof - + show "(u \ u') \ \ = \" using wfT_supp flip_fresh_fresh + by (metis assms(1) fresh_def u_not_in_t) + show "(u \ u') \ \ = \" using u_not_in_g wfX_wfY assms flip_fresh_fresh fresh_def by metis + show "(u \ u') \ \ = \" using theta_flip_eq assms wfX_wfY by metis + show "(u \ u') \ \ = \" using u_not_in_b_set flip_fresh_fresh fresh_def by metis +qed + +lemma wfT_wf_cons_flip: + fixes c::c and x::x + assumes "wfT P \ \ \ z : b | c \" and "atom x \ (c,\)" + shows "wfG P \ ((x,b,c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\\)" +proof - + have "\ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \ = \ z : b | c \" using assms freshers type_eq_subst by metis + hence *:"wfT P \ \ \ x : b | c[z::=V_var x]\<^sub>c\<^sub>v \" using assms by metis + show ?thesis proof(rule wfG_consI) + show \ P; \ \\<^sub>w\<^sub>f \ \ using assms wfT_wf by auto + show \atom x \ \\ using assms by auto + show \ P; \ \\<^sub>w\<^sub>f b \ using assms wfX_wfY b_of.simps by metis + show \ P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=V_var x]\<^sub>c\<^sub>v \ using wfT_wfC * assms fresh_Pair by metis + qed +qed + +section \Context Strengthening\ + +text \We can remove an entry for a variable from the context if the variable doesn't appear in the +term and the variable is not used later in the context or any other context\ + +lemma fresh_restrict: + fixes y::"'a::at_base" and \::\ + assumes "atom y \ (\' @ (x, b, c) #\<^sub>\ \)" + shows "atom y \ (\'@\)" +using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using fresh_GCons fresh_GNil by auto +next + case (GCons x' b' c' \'') + then show ?case using fresh_GCons fresh_GNil by auto +qed + +lemma wf_restrict1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows "\; \; \ \\<^sub>w\<^sub>f v : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ v \ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f v : b" and + + "\; \; \ \\<^sub>w\<^sub>f c \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ c\ atom x \ \\<^sub>1 \ \ ; \ ; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f c" and + "\; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\<^sub>1 \ \; \ \\<^sub>w\<^sub>f \\<^sub>1@\\<^sub>2" and + "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \True" and + + "\; \ \\<^sub>w\<^sub>f b \ True" and + + "\; \; \ \\<^sub>w\<^sub>f ce : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ ce \ atom x \ \\<^sub>1 \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f ce : b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(induct arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts) + case (wfV_varI \ \ \ b c y) + hence "y\x" using v.fresh by auto + hence "Some (b, c) = lookup (\\<^sub>1@\\<^sub>2) y" using lookup_restrict wfV_varI by metis + then show ?case using wfV_varI wf_intros by metis +next + case (wfV_litI \ \ l) + then show ?case using e.fresh wf_intros by metis +next + case (wfV_pairI \ \ \ v1 b1 v2 b2) + show ?case proof + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v1 : b1" using wfV_pairI by auto + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v2 : b2" using wfV_pairI by auto + qed +next + case (wfV_consI s dclist \ dc x b c \ \ v) + show ?case proof + show "AF_typedef s dclist \ set \" using wfV_consI by auto + show "(dc, \ x : b | c \) \ set dclist" using wfV_consI by auto + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v : b" using wfV_consI by auto + qed +next + case (wfV_conspI s bv dclist \ dc x b' c \ b \ v) + show ?case proof + show "AF_typedef_poly s bv dclist \ set \" using wfV_conspI by auto + show "(dc, \ x : b' | c \) \ set dclist" using wfV_conspI by auto + show "\; \ \\<^sub>w\<^sub>f b" using wfV_conspI by auto + show " \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_conspI by auto + show "atom bv \ (\, \, \\<^sub>1 @ \\<^sub>2, b, v)" unfolding fresh_prodN fresh_append_g using wfV_conspI fresh_prodN fresh_GCons fresh_append_g by metis + qed +next + case (wfCE_valI \ \ \ v b) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_eqI \ \ \ v1 v2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using ce.fresh wf_intros by metis +next + case (wfCE_lenI \ \ \ v1) + then show ?case using ce.fresh wf_intros by metis +next + case (wfTI z \ \ \ b c) + hence "x \ z" using wfTI + fresh_GCons fresh_prodN fresh_PairD(1) fresh_gamma_append not_self_fresh by metis + show ?case proof + show \atom z \ (\, \, \\<^sub>1 @ \\<^sub>2)\ using wfTI fresh_restrict[of z] using wfG_fresh_x wfX_wfY wfTI fresh_prodN by metis + show \ \; \ \\<^sub>w\<^sub>f b \ using wfTI by auto + have "\; \; ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ \\<^sub>2 \\<^sub>w\<^sub>f c " proof(rule wfTI(5)[of "(z, b, TRUE) #\<^sub>\ \\<^sub>1" ]) + show \(z, b, TRUE) #\<^sub>\ \ = ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ (x, b', c') #\<^sub>\ \\<^sub>2\ using wfTI by auto + show \atom x \ c\ using wfTI \.fresh \x \ z\ by auto + show \atom x \ (z, b, TRUE) #\<^sub>\ \\<^sub>1\ using wfTI \x \ z\ fresh_GCons by simp + qed + thus \ \; \; (z, b, TRUE) #\<^sub>\ \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f c \ by auto + qed +next + case (wfC_eqI \ \ \ e1 b e2) + show ?case proof + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f e1 : b " using wfC_eqI c.fresh fresh_Nil by auto + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f e2 : b " using wfC_eqI c.fresh fresh_Nil by auto + qed +next + case (wfC_trueI \ \) + then show ?case using c.fresh wf_intros by metis +next + case (wfC_falseI \ \) + then show ?case using c.fresh wf_intros by metis +next + case (wfC_conjI \ \ c1 c2) + then show ?case using c.fresh wf_intros by metis +next + case (wfC_disjI \ \ c1 c2) + then show ?case using c.fresh wf_intros by metis +next +case (wfC_notI \ \ c1) + then show ?case using c.fresh wf_intros by metis +next + case (wfC_impI \ \ c1 c2) + then show ?case using c.fresh wf_intros by metis +next + case (wfG_nilI \) + then show ?case using wfV_varI wf_intros + by (meson GNil_append \.simps(3)) +next + case (wfG_cons1I c1 \ \ G x1 b1) + show ?case proof(cases "\\<^sub>1=GNil") + case True + then show ?thesis using wfG_cons1I wfG_consI by auto + next + case False + then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \\<^sub>1" using GCons_eq_append_conv wfG_cons1I by auto + hence **:"G=G' @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons1I by auto + + have " \; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ (G' @ \\<^sub>2)" proof(rule Wellformed.wfG_cons1I) + show \c1 \ {TRUE, FALSE}\ using wfG_cons1I by auto + show \atom x1 \ G' @ \\<^sub>2\ using wfG_cons1I(4) ** fresh_restrict by metis + have " atom x \ G'" using wfG_cons1I * using fresh_GCons by blast + thus \ \; \ \\<^sub>w\<^sub>f G' @ \\<^sub>2 \ using wfG_cons1I(3)[of G'] ** by auto + have "atom x \ c1 \ atom x \ (x1, b1, TRUE) #\<^sub>\ G'" using fresh_GCons \atom x \ \\<^sub>1\ * by auto + thus \ \; \; (x1, b1, TRUE) #\<^sub>\ G' @ \\<^sub>2 \\<^sub>w\<^sub>f c1 \ using wfG_cons1I(6)[of "(x1, b1, TRUE) #\<^sub>\ G'"] ** * wfG_cons1I by auto + show \ \; \ \\<^sub>w\<^sub>f b1 \ using wfG_cons1I by auto + qed + thus ?thesis using * by auto + qed +next + case (wfG_cons2I c1 \ \ G x1 b1) + show ?case proof(cases "\\<^sub>1=GNil") + case True + then show ?thesis using wfG_cons2I wfG_consI by auto + next + case False + then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \\<^sub>1" using GCons_eq_append_conv wfG_cons2I by auto + hence **:"G=G' @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons2I by auto + + have " \; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ (G' @ \\<^sub>2)" proof(rule Wellformed.wfG_cons2I) + show \c1 \ {TRUE, FALSE}\ using wfG_cons2I by auto + show \atom x1 \ G' @ \\<^sub>2\ using wfG_cons2I ** fresh_restrict by metis + have " atom x \ G'" using wfG_cons2I * using fresh_GCons by blast + thus \ \; \ \\<^sub>w\<^sub>f G' @ \\<^sub>2 \ using wfG_cons2I ** by auto + show \ \; \ \\<^sub>w\<^sub>f b1 \ using wfG_cons2I by auto + qed + thus ?thesis using * by auto + qed +qed(auto)+ + +lemma wf_restrict2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ e \ atom x \ \\<^sub>1 \ atom x \ \ \ \; \; \; \\<^sub>1@\\<^sub>2 ; \ \\<^sub>w\<^sub>f e : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and + "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and + "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and + "\ \\<^sub>w\<^sub>f (\::\) \ True " and + "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ atom x \ \\<^sub>1 \ atom x \ \ \ \; \; \\<^sub>1@\\<^sub>2 \\<^sub>w\<^sub>f \" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" + +proof(induct arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts) + case (wfE_valI \ \ \ \ v b) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_plusI \ \ \ \ v1 v2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_leqI \ \ \ \ v1 v2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_eqI \ \ \ \ v1 b v2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_fstI \ \ \ \ v1 b1 b2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_sndI \ \ \ \ v1 b1 b2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_concatI \ \ \ \ v1 v2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_splitI \ \ \ \ v1 v2) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_lenI \ \ \ \ v1) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_appI \ \ \ \ f x b c \ s' v) + then show ?case using e.fresh wf_intros wf_restrict1 by metis +next + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto + + have "atom bv \ \\<^sub>1 @ \\<^sub>2" using wfE_appPI fresh_prodN fresh_restrict by metis + thus \atom bv \ (\, \, \, \\<^sub>1 @ \\<^sub>2, \, b', v, (b_of \)[bv::=b']\<^sub>b)\ + using wfE_appPI fresh_prodN by auto + + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f\ using wfE_appPI by auto + show \ \; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI wf_restrict1 by auto + qed +next + case (wfE_mvarI \ \ \ \ u \) + then show ?case using e.fresh wf_intros by metis +next + case (wfD_emptyI \ \) + then show ?case using c.fresh wf_intros wf_restrict1 by metis +next + case (wfD_cons \ \ \ \ \ u) + show ?case proof + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f \" using wfD_cons fresh_DCons by metis + show "\; \; \\<^sub>1 @ \\<^sub>2 \\<^sub>w\<^sub>f \ " using wfD_cons fresh_DCons fresh_Pair wf_restrict1 by metis + show "u \ fst ` setD \" using wfD_cons by auto + qed +next + case (wfFTNone \ ft) + then show ?case by auto +next + case (wfFTSome \ bv ft) + then show ?case by auto +next + case (wfFTI \ B b \ x c s \) + then show ?case by auto +qed(auto)+ + +lemmas wf_restrict=wf_restrict1 wf_restrict2 + +lemma wfT_restrict2: + fixes \::\ + assumes "wfT \ \ ((x, b, c) #\<^sub>\ \) \" and "atom x \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \" + using wf_restrict1(4)[of \ \ "((x, b, c) #\<^sub>\ \)" \ GNil x "b" "c" \] assms fresh_GNil append_g.simps by auto + +lemma wfG_intros2: + assumes "wfC P \ ((x,b,TRUE) #\<^sub>\\) c" + shows "wfG P \ ((x,b,c) #\<^sub>\\)" +proof - + have "wfG P \ ((x,b,TRUE) #\<^sub>\\)" using wfC_wf assms by auto + hence *:"wfG P \ \ \ atom x \ \ \ wfB P \ b" using wfG_elims by metis + show ?thesis using assms proof(cases "c \ {TRUE,FALSE}") + case True + then show ?thesis using wfG_cons2I * by auto + next + case False + then show ?thesis using wfG_cons1I * assms by auto + qed +qed + +section \Type Definitions\ + +lemma wf_theta_weakening1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and \ :: \ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list and t::\ + + shows "\; \; \ \\<^sub>w\<^sub>f v : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b" and + "\; \; \ \\<^sub>w\<^sub>f c \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f c" and + "\; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f ts" and + "\\<^sub>w\<^sub>f P \ True " and + "\; \ \\<^sub>w\<^sub>f b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f b" and + "\; \; \ \\<^sub>w\<^sub>f ce : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f ce : b" and + "\ \\<^sub>w\<^sub>f td \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' \\<^sub>w\<^sub>f td" +proof(nominal_induct b and c and \ and \ and ts and P and b and b and td + avoiding: \' + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_consI s dclist \ dc x b c \ \ v) + show ?case proof + show \AF_typedef s dclist \ set \'\ using wfV_consI by auto + show \(dc, \ x : b | c \) \ set dclist\ using wfV_consI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b \ using wfV_consI by auto + qed +next + case (wfV_conspI s bv dclist \ dc x b' c \ b \ v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \'\ using wfV_conspI by auto + show \(dc, \ x : b' | c \) \ set dclist\ using wfV_conspI by auto + show \\' ; \ ; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by auto + show "\' ; \ \\<^sub>w\<^sub>f b " using wfV_conspI by auto + show "atom bv \ (\', \, \, b, v)" using wfV_conspI fresh_prodN by auto + qed +next + case (wfTI z \ \ \ b c) + thus ?case using Wellformed.wfTI by auto +next + case (wfB_consI \ s dclist) + show ?case proof + show \ \\<^sub>w\<^sub>f \' \ using wfB_consI by auto + show \AF_typedef s dclist \ set \'\ using wfB_consI by auto + qed +next + case (wfB_appI \ \ b s bv dclist) + show ?case proof + show \ \\<^sub>w\<^sub>f \' \ using wfB_appI by auto + show \AF_typedef_poly s bv dclist \ set \'\ using wfB_appI by auto + show "\' ; \ \\<^sub>w\<^sub>f b" using wfB_appI by simp + qed +qed(metis wf_intros)+ + +lemma wf_theta_weakening2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and \ :: \ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list and t::\ + shows + "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b" and + "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and + "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' \\<^sub>w\<^sub>f (\::\)" and + "\; \; \ \\<^sub>w\<^sub>f \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f \" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ \\<^sub>w\<^sub>f ftq" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \' ; \ ; \ \\<^sub>w\<^sub>f ft" + +proof(nominal_induct b and b and b and b and \ and \ and ftq and ft + avoiding: \' +rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + show ?case proof + show \ \' \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \' ; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI wf_theta_weakening1 by auto + show \atom bv \ (\, \', \, \, \, b', v, (b_of \)[bv::=b']\<^sub>b)\ using wfE_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f\ using wfE_appPI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI wf_theta_weakening1 by auto + qed +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + show ?case proof + show \ \' ; \ ; \ \\<^sub>w\<^sub>f v : B_id tid \ using wfS_matchI wf_theta_weakening1 by auto + show \AF_typedef tid dclist \ set \'\ using wfS_matchI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \ \' \\<^sub>w\<^sub>f \ \ using wfS_matchI by auto + show \\'; \; \; \; \; tid; dclist \\<^sub>w\<^sub>f cs : b \ using wfS_matchI by auto + qed +next + case (wfS_varI \ \ \ \ v u \ \ b s) + show ?case proof + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_varI wf_theta_weakening1 by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI wf_theta_weakening1 by auto + show \atom u \ (\, \', \, \, \, \, v, b)\ using wfS_varI by auto + show \ \' ; \ ; \ ; \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b \ using wfS_varI by auto + qed +next + case (wfS_letI \ \ \ \ \ e b' x s b) + show ?case proof + show \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto + show \ \' ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_letI by auto + show \atom x \ (\, \', \, \, \, e, b)\ using wfS_letI by auto + qed +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + show ?case proof + show \ \' ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_theta_weakening1 by auto + show \ \' ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : b \ using wfS_let2I by auto + show \atom x \ (\, \', \, \, \, s1, b, \)\ using wfS_let2I by auto + qed +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + show ?case proof + show \ \' ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_branchI by auto + show \atom x \ (\, \', \, \, \, \, \)\ using wfS_branchI by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_branchI by auto + qed +next + case (wfPhi_consI f \ \ ft) + show ?case proof + show "f \ name_of_fun ` set \" using wfPhi_consI by auto + show "\' ; \ \\<^sub>w\<^sub>f ft" using wfPhi_consI by auto + show "\' \\<^sub>w\<^sub>f \" using wfPhi_consI by auto + qed +next + case (wfFTNone \ ft) + then show ?case using wf_intros by metis +next + case (wfFTSome \ bv ft) + then show ?case using wf_intros by metis +next + case (wfFTI \ B b \ x c s \) + thus ?case using Wellformed.wfFTI wf_theta_weakening1 by simp +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case proof + show \ \' ; \ ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI wf_theta_weakening1 by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f c \ using wfS_assertI wf_theta_weakening1 by auto + show \ \' ; \ ; \ \\<^sub>w\<^sub>f \ \ using wfS_assertI wf_theta_weakening1 by auto + have "atom x \ \'" using wf_supp(6)[OF \\\<^sub>w\<^sub>f \' \] fresh_def by auto + thus \atom x \ (\, \', \, \, \, c, b, s)\ using wfS_assertI fresh_prodN fresh_def by simp + qed +qed(metis wf_intros wf_theta_weakening1 )+ + +lemmas wf_theta_weakening = wf_theta_weakening1 wf_theta_weakening2 + +lemma lookup_wfTD: + fixes td::type_def + assumes "td \ set \" and "\\<^sub>w\<^sub>f \" + shows "\ \\<^sub>w\<^sub>f td" + using assms proof(induct \ ) + case Nil + then show ?case by auto +next + case (Cons td' \') + then consider "td = td'" | "td \ set \'" by auto + then have "\' \\<^sub>w\<^sub>f td" proof(cases) + case 1 + then show ?thesis using Cons using wfTh_elims by auto + next + case 2 + then show ?thesis using Cons using wfTh_elims by auto + qed + then show ?case using wf_theta_weakening Cons by (meson set_subset_Cons) +qed + +subsection \Simple\ + +lemma wfTh_dclist_unique: + assumes "wfTh \" and "AF_typedef tid dclist1 \ set \" and "AF_typedef tid dclist2 \ set \" + shows "dclist1 = dclist2" +using assms proof(induct \ rule: \_induct) + case TNil + then show ?case by auto +next + case (AF_typedef tid' dclist' \') + then show ?case using wfTh_elims + by (metis image_eqI name_of_type.simps(1) set_ConsD type_def.eq_iff(1)) +next + case (AF_typedef_poly tid bv dclist \') + then show ?case using wfTh_elims by auto +qed + +lemma wfTs_ctor_unique: + fixes dclist::"(string*\) list" + assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f dclist" and "(c, t1) \ set dclist" and "(c,t2) \ set dclist" + shows "t1 = t2" + using assms proof(induct dclist rule: list.inducts) + case Nil + then show ?case by auto +next + case (Cons x1 x2) + consider "x1 = (c,t1)" | "x1 = (c,t2)" | "x1 \ (c,t1) \ x1 \ (c,t2)" by auto + thus ?case proof(cases) + case 1 + then show ?thesis using Cons wfTs_elims set_ConsD + by (metis fst_conv image_eqI prod.inject) + next + case 2 + then show ?thesis using Cons wfTs_elims set_ConsD + by (metis fst_conv image_eqI prod.inject) + next + case 3 + then show ?thesis using Cons wfTs_elims by (metis set_ConsD) + qed +qed + +lemma wfTD_ctor_unique: + assumes "\ \\<^sub>w\<^sub>f (AF_typedef tid dclist)" and "(c, t1) \ set dclist" and "(c,t2) \ set dclist" + shows "t1 = t2" + using wfTD_elims wfTs_elims assms wfTs_ctor_unique by metis + +lemma wfTh_ctor_unique: + assumes "wfTh \" and "AF_typedef tid dclist \ set \" and "(c, t1) \ set dclist" and "(c,t2) \ set dclist" + shows "t1 = t2" + using lookup_wfTD wfTD_ctor_unique assms by metis + +lemma wfTs_supp_t: + fixes dclist::"(string*\) list" + assumes "(c,t) \ set dclist" and "\ ; B ; GNil \\<^sub>w\<^sub>f dclist" + shows "supp t \ supp B" +using assms proof(induct dclist arbitrary: c t rule:list.induct) + case Nil + then show ?case by auto +next + case (Cons ct dclist') + then consider "ct = (c,t)" | "(c,t) \ set dclist'" by auto + then show ?case proof(cases) + case 1 + then have "\ ; B ; GNil \\<^sub>w\<^sub>f t" using Cons wfTs_elims by blast + thus ?thesis using wfT_supp atom_dom.simps by force + next + case 2 + then show ?thesis using Cons wfTs_elims by metis + qed +qed + +lemma wfTh_lookup_supp_empty: + fixes t::\ + assumes "AF_typedef tid dclist \ set \" and "(c,t) \ set dclist" and "\\<^sub>w\<^sub>f \" + shows "supp t = {}" +proof - + have "\ ; {||} ; GNil \\<^sub>w\<^sub>f dclist" using assms lookup_wfTD wfTD_elims by metis + thus ?thesis using wfTs_supp_t assms by force +qed + +lemma wfTh_supp_b: + assumes "AF_typedef tid dclist \ set \" and "(dc,\ z : b | c \ ) \ set dclist" and "\\<^sub>w\<^sub>f \" + shows "supp b = {}" + using assms wfTh_lookup_supp_empty \.supp by blast + +lemma wfTh_b_eq_iff: + fixes bva1::bv and bva2::bv and dc::string + assumes "(dc, \ x1 : b1 | c1 \) \ set dclist1" and "(dc, \ x2 : b2 | c2 \) \ set dclist2" and + "wfTs P {|bva1|} GNil dclist1" and "wfTs P {|bva2|} GNil dclist2" + "[[atom bva1]]lst.dclist1 = [[atom bva2]]lst.dclist2" + shows "[[atom bva1]]lst. (dc,\ x1 : b1 | c1 \) = [[atom bva2]]lst. (dc,\ x2 : b2 | c2 \)" +using assms proof(induct dclist1 arbitrary: dclist2) + case Nil + then show ?case by auto +next + case (Cons dct1' dclist1') + show ?case proof(cases "dclist2 = []") + case True + then show ?thesis using Cons by auto + next + case False + then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using list.exhaust by metis + hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'" + using Cons lst_head_cons Cons cons by metis + hence **: "fst dct1' = fst dct2'" using lst_fst[THEN lst_pure] + by (metis (no_types) \[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'\ + \\x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \ t1 = t2\ fst_conv surj_pair) + show ?thesis proof(cases "fst dct1' = dc") + case True + have "dc \ fst ` set dclist1'" using wfTs_elims Cons by (metis True fstI) + hence 1:"(dc, \ x1 : b1 | c1 \) = dct1'" using Cons by (metis fstI image_iff set_ConsD) + have "dc \ fst ` set dclist2'" using wfTs_elims Cons cons + by (metis "**" True fstI) + hence 2:"(dc, \ x2 : b2 | c2 \) = dct2' " using Cons cons by (metis fst_conv image_eqI set_ConsD) + then show ?thesis using Cons * 1 2 by blast + next + case False + hence "fst dct2' \ dc" using ** by auto + hence "(dc, \ x1 : b1 | c1 \) \ set dclist1' \ (dc, \ x2 : b2 | c2 \) \ set dclist2' " using Cons cons False + by (metis fstI set_ConsD) + moreover have "[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2'" using * False by metis + ultimately show ?thesis using Cons ** * + using cons wfTs_elims(2) by blast + qed + qed +qed + + +subsection \Polymorphic\ + +lemma wfTh_wfTs_poly: + fixes dclist::"(string * \) list" + assumes "AF_typedef_poly tyid bva dclist \ set P" and "\\<^sub>w\<^sub>f P" + shows "P ; {|bva|} ; GNil \\<^sub>w\<^sub>f dclist" +proof - + have *:"P \\<^sub>w\<^sub>f AF_typedef_poly tyid bva dclist" using lookup_wfTD assms by simp + + obtain bv lst where *:"P ; {|bv|} ; GNil \\<^sub>w\<^sub>f lst \ + (\c. atom c \ (dclist, lst) \ atom c \ (bva, bv, dclist, lst) \ (bva \ c) \ dclist = (bv \ c) \ lst)" + using wfTD_elims(2)[OF *] by metis + + obtain c::bv where **:"atom c \ ((dclist, lst),(bva, bv, dclist, lst))" using obtain_fresh by metis + have "P ; {|bv|} ; GNil \\<^sub>w\<^sub>f lst" using * by metis + hence "wfTs ((bv \ c) \ P) ((bv \ c) \ {|bv|}) ((bv \ c) \ GNil) ((bv \ c) \ lst)" using ** wfTs.eqvt by metis + hence "wfTs P {|c|} GNil ((bva \ c) \ dclist)" using * theta_flip_eq fresh_GNil assms + proof - + have "\b ba. (ba::bv \ b) \ P = P" by (metis \\\<^sub>w\<^sub>f P\ theta_flip_eq) + then show ?thesis + using "*" "**" \(bv \ c) \ P ; (bv \ c) \ {|bv|} ; (bv \ c) \ GNil \\<^sub>w\<^sub>f (bv \ c) \ lst\ by fastforce + qed + hence "wfTs ((bva \ c) \ P) ((bva \ c) \ {|bva|}) ((bva \ c) \ GNil) ((bva \ c) \ dclist)" + using wfTs.eqvt fresh_GNil + by (simp add: assms(2) theta_flip_eq2) + + thus ?thesis using wfTs.eqvt permute_flip_cancel by metis +qed + +lemma wfTh_dclist_poly_unique: + assumes "wfTh \" and "AF_typedef_poly tid bva dclist1 \ set \" and "AF_typedef_poly tid bva2 dclist2 \ set \" + shows "[[atom bva]]lst. dclist1 = [[atom bva2]]lst.dclist2" +using assms proof(induct \ rule: \_induct) + case TNil + then show ?case by auto +next + case (AF_typedef tid' dclist' \') + then show ?case using wfTh_elims by auto +next + case (AF_typedef_poly tid bv dclist \') + then show ?case using wfTh_elims image_eqI name_of_type.simps set_ConsD type_def.eq_iff + by (metis Abs1_eq(3)) +qed + +lemma wfTh_poly_lookup_supp: + fixes t::\ + assumes "AF_typedef_poly tid bv dclist \ set \" and "(c,t) \ set dclist" and "\\<^sub>w\<^sub>f \" + shows "supp t \ {atom bv}" +proof - + have "supp dclist \ {atom bv}" using assms lookup_wfTD wf_supp1 type_def.supp + by (metis Diff_single_insert Un_subset_iff list.simps(15) supp_Nil supp_of_atom_list) + then show ?thesis using assms(2) proof(induct dclist) + case Nil + then show ?case by auto + next + case (Cons a dclist) + then show ?case using supp_Pair supp_Cons + by (metis (mono_tags, hide_lams) Un_empty_left Un_empty_right pure_supp subset_Un_eq subset_singletonD supp_list_member) + qed +qed + +lemma wfTh_poly_supp_b: + assumes "AF_typedef_poly tid bv dclist \ set \" and "(dc,\ z : b | c \ ) \ set dclist" and "\\<^sub>w\<^sub>f \" + shows "supp b \ {atom bv}" + using assms wfTh_poly_lookup_supp \.supp by force + +lemma subst_g_inside: + fixes x::x and c::c and \::\ and \'::\ + assumes "wfG P \ (\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" + shows "(\' @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)[x::=v]\<^sub>\\<^sub>v = (\'[x::=v]\<^sub>\\<^sub>v@\)" +using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using subst_gb.simps by simp +next + case (GCons x' b' c' G) + hence wfg:"wfG P \ (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \) \ atom x' \ (G @ (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \)" using wfG_elims(2) + using GCons.prems append_g.simps by metis + hence "atom x \ atom_dom ((x', b', c') #\<^sub>\ G)" using GCons wfG_inside_fresh by fast + hence "x\x'" + using GCons append_Cons wfG_inside_fresh atom_dom.simps toSet.simps by simp + hence "((GCons (x', b', c') G) @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \))[x::=v]\<^sub>\\<^sub>v = + (GCons (x', b', c') (G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \)))[x::=v]\<^sub>\\<^sub>v" by auto + also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c[z::=V_var x]\<^sub>c\<^sub>v) \))[x::=v]\<^sub>\\<^sub>v)" + using subst_gv.simps \x\x'\ by simp + also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (G[x::=v]\<^sub>\\<^sub>v @ \)" using GCons wfg by blast + also have "... = ((x', b', c') #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>v @ \" using subst_gv.simps \x\x'\ by simp + finally show ?case by auto +qed + +lemma wfTh_td_eq: + assumes "td1 \ set (td2 # P)" and "wfTh (td2 # P)" and "name_of_type td1 = name_of_type td2" + shows "td1 = td2" +proof(rule ccontr) + assume as: "td1 \ td2" + have "name_of_type td2 \ name_of_type ` set P" using wfTh_elims(2)[OF assms(2)] by metis + moreover have "td1 \ set P" using assms as by simp + ultimately have "name_of_type td1 \ name_of_type td2" + by (metis rev_image_eqI) + thus False using assms by auto +qed + +lemma wfTh_td_unique: + assumes "td1 \ set P" and "td2 \ set P" and "wfTh P" and "name_of_type td1 = name_of_type td2" + shows "td1 = td2" +using assms proof(induct P rule: list.induct) + case Nil + then show ?case by auto +next + case (Cons td \') + consider "td = td1" | "td = td2" | "td \ td1 \ td \ td2" by auto + then show ?case proof(cases) + case 1 + then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis + next + case 2 + then show ?thesis using Cons wfTh_elims wfTh_td_eq by metis + next + case 3 + then show ?thesis using Cons wfTh_elims by auto + qed +qed + +lemma wfTs_distinct: + fixes dclist::"(string * \) list" + assumes "\ ; B ; GNil \\<^sub>w\<^sub>f dclist" + shows "distinct (map fst dclist)" +using assms proof(induct dclist rule: list.induct) + case Nil + then show ?case by auto +next + case (Cons x1 x2) + then show ?case + by (metis Cons.hyps Cons.prems distinct.simps(2) fst_conv list.set_map list.simps(9) wfTs_elims(2)) +qed + +lemma wfTh_dclist_distinct: + assumes "AF_typedef s dclist \ set P" and "wfTh P" + shows "distinct (map fst dclist)" +proof - + have "wfTD P (AF_typedef s dclist)" using assms lookup_wfTD by auto + hence "wfTs P {||} GNil dclist" using wfTD_elims by metis + thus ?thesis using wfTs_distinct by metis +qed + +lemma wfTh_dc_t_unique2: + assumes "AF_typedef s dclist' \ set P" and "(dc,tc' ) \ set dclist'" and "AF_typedef s dclist \ set P" and "wfTh P" and + "(dc, tc) \ set dclist" + shows "tc= tc'" +proof - + have "dclist = dclist'" using assms wfTh_td_unique name_of_type.simps by force + moreover have "distinct (map fst dclist)" using wfTh_dclist_distinct assms by auto + ultimately show ?thesis using assms + by (meson eq_key_imp_eq_value) +qed + +lemma wfTh_dc_t_unique: + assumes "AF_typedef s dclist' \ set P" and "(dc, \ x' : b' | c' \ ) \ set dclist'" and "AF_typedef s dclist \ set P" and "wfTh P" and + "(dc, \ x : b | c \) \ set dclist" + shows "\ x' : b' | c' \= \ x : b | c \" + using assms wfTh_dc_t_unique2 by metis + +lemma wfTs_wfT: + fixes dclist::"(string *\) list" and t::\ + assumes "\; \; GNil \\<^sub>w\<^sub>f dclist" and "(dc,t) \ set dclist" + shows "\; \; GNil \\<^sub>w\<^sub>f t" +using assms proof(induct dclist rule:list.induct) + case Nil + then show ?case by auto +next + case (Cons x1 x2) + thus ?case using wfTs_elims(2)[OF Cons(2)] by auto +qed + +lemma wfTh_wfT: + fixes t::\ + assumes "wfTh P" and "AF_typedef tid dclist \ set P" and "(dc,t) \ set dclist" + shows "P ; {||} ; GNil \\<^sub>w\<^sub>f t" +proof - + have "P \\<^sub>w\<^sub>f AF_typedef tid dclist" using lookup_wfTD assms by auto + hence "P ; {||} ; GNil \\<^sub>w\<^sub>f dclist" using wfTD_elims by auto + thus ?thesis using wfTs_wfT assms by auto +qed + +lemma td_lookup_eq_iff: + fixes dc :: string and bva1::bv and bva2::bv + assumes "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst. dclist2" and "(dc, \ x : b | c \) \ set dclist1" + shows "\x2 b2 c2. (dc, \ x2 : b2 | c2 \) \ set dclist2" +using assms proof(induct dclist1 arbitrary: dclist2) + case Nil + then show ?case by auto +next + case (Cons dct1' dclist1') + then obtain dct2' and dclist2' where cons:"dct2' # dclist2' = dclist2" using lst_head_cons_neq_nil[OF Cons(2)] list.exhaust by metis + hence *:"[[atom bva1]]lst. dclist1' = [[atom bva2]]lst. dclist2' \ [[atom bva1]]lst. dct1' = [[atom bva2]]lst. dct2'" + using Cons lst_head_cons Cons cons by metis + show ?case proof(cases "dc=fst dct1'") + case True + hence "dc = fst dct2'" using * lst_fst[ THEN lst_pure ] + proof - + show ?thesis + by (metis (no_types) "local.*" True \\x2 x1 t2' t2a t2 t1. [[atom x1]]lst. (t1, t2a) = [[atom x2]]lst. (t2, t2') \ t1 = t2\ prod.exhaust_sel) (* 31 ms *) + qed + obtain x2 b2 and c2 where "snd dct2' = \ x2 : b2 | c2 \" using obtain_fresh_z by metis + hence "(dc, \ x2 : b2 | c2 \) = dct2'" using \dc = fst dct2'\ + by (metis prod.exhaust_sel) + then show ?thesis using cons by force + next + case False + hence "(dc, \ x : b | c \) \ set dclist1'" using Cons by auto + then show ?thesis using Cons + by (metis "local.*" cons list.set_intros(2)) + qed +qed + +lemma lst_t_b_eq_iff: + fixes bva1::bv and bva2::bv + assumes "[[atom bva1]]lst. \ x1 : b1 | c1 \ = [[atom bva2]]lst. \ x2 : b2 | c2 \" + shows "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2" +proof(subst Abs1_eq_iff_all(3)[of bva1 b1 bva2 b2],rule,rule,rule) + fix c::bv + assume "atom c \ ( \ x1 : b1 | c1 \ , \ x2 : b2 | c2 \)" and "atom c \ (bva1, bva2, b1, b2)" + + show "(bva1 \ c) \ b1 = (bva2 \ c) \ b2" using assms Abs1_eq_iff(3) assms + by (metis Abs1_eq_iff_fresh(3) \atom c \ (bva1, bva2, b1, b2)\ \.fresh \.perm_simps type_eq_subst_eq2(2)) +qed + +lemma wfTh_typedef_poly_b_eq_iff: + assumes "AF_typedef_poly tyid bva1 dclist1 \ set P" and "(dc, \ x1 : b1 | c1 \) \ set dclist1" + and "AF_typedef_poly tyid bva2 dclist2 \ set P" and "(dc, \ x2 : b2 | c2 \) \ set dclist2" and "\\<^sub>w\<^sub>f P" +shows "b1[bva1::=b]\<^sub>b\<^sub>b = b2[bva2::=b]\<^sub>b\<^sub>b" +proof - + have "[[atom bva1]]lst. dclist1 = [[atom bva2]]lst.dclist2" using assms wfTh_dclist_poly_unique by metis + hence "[[atom bva1]]lst. (dc,\ x1 : b1 | c1 \) = [[atom bva2]]lst. (dc,\ x2 : b2 | c2 \)" using wfTh_b_eq_iff assms wfTh_wfTs_poly by metis + hence "[[atom bva1]]lst. \ x1 : b1 | c1 \ = [[atom bva2]]lst. \ x2 : b2 | c2 \" using lst_snd by metis + hence "[[atom bva1]]lst. b1 = [[atom bva2]]lst.b2" using lst_t_b_eq_iff by metis + thus ?thesis using subst_b_flip_eq_two subst_b_b_def by metis +qed + +section \Equivariance Lemmas\ + +lemma x_not_in_u_set[simp]: + fixes x::x and us::"u fset" + shows "atom x \ supp us" + by(induct us,auto, simp add: supp_finsert supp_at_base) + +lemma wfS_flip_eq: + fixes s1::s and x1::x and s2::s and x2::x and \::\ + assumes "[[atom x1]]lst. s1 = [[atom x2]]lst. s2" and "[[atom x1]]lst. t1 = [[atom x2]]lst. t2" and "[[atom x1]]lst. c1 = [[atom x2]]lst. c2" and "atom x2 \ \" and + " \; \; \ \\<^sub>w\<^sub>f \" and + "\ ; \ ; \ ; (x1, b, c1) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of t1" + shows "\ ; \ ; \ ; (x2, b, c2) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s2 : b_of t2" +proof(cases "x1=x2") + case True + hence "s1 = s2 \ t1 = t2 \ c1 = c2" using assms Abs1_eq_iff by metis + then show ?thesis using assms True by simp +next + case False + have "\\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ \; \; \ \\<^sub>w\<^sub>f \" using wfX_wfY assms by metis + moreover have "atom x1 \ \" using wfX_wfY wfG_elims assms by metis + moreover hence "atom x1 \ \ \ atom x2 \ \ " using wfD_x_fresh assms by auto + ultimately have " \ ; \ ; \ ; (x2 \ x1) \ ((x1, b, c1) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>f (x2 \ x1) \ s1 : (x2 \ x1) \ b_of t1" + using wfS.eqvt theta_flip_eq phi_flip_eq assms flip_base_eq beta_flip_eq flip_fresh_fresh supp_b_empty by metis + hence " \ ; \ ; \ ; ((x2, b, (x2 \ x1) \ c1) #\<^sub>\ ((x2 \ x1) \ \)) ; \ \\<^sub>w\<^sub>f (x2 \ x1) \ s1 : b_of ((x2 \ x2) \ t1)" by fastforce + thus ?thesis using assms Abs1_eq_iff + proof - + have f1: "x2 = x1 \ t2 = t1 \ x2 \ x1 \ t2 = (x2 \ x1) \ t1 \ atom x2 \ t1" + by (metis (full_types) Abs1_eq_iff(3) \[[atom x1]]lst. t1 = [[atom x2]]lst. t2\) (* 125 ms *) + then have "x2 \ x1 \ s2 = (x2 \ x1) \ s1 \ atom x2 \ s1 \ b_of t2 = (x2 \ x1) \ b_of t1" + by (metis b_of.eqvt) (* 0.0 ms *) + then show ?thesis + using f1 by (metis (no_types) Abs1_eq_iff(3) G_cons_flip_fresh3 \[[atom x1]]lst. c1 = [[atom x2]]lst. c2\ \[[atom x1]]lst. s1 = [[atom x2]]lst. s2\ \\ ; \ ; \ ; (x1, b, c1) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s1 : b_of t1\ \\ ; \ ; \ ; (x2 \ x1) \ ((x1, b, c1) #\<^sub>\ \) ; \ \\<^sub>w\<^sub>f (x2 \ x1) \ s1 : (x2 \ x1) \ b_of t1\ \atom x1 \ \\ \atom x2 \ \\) (* 593 ms *) + qed +qed + +section \Lookup\ + +lemma wf_not_in_prefix: + assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b1,c1) #\<^sub>\\)" + shows "x \ fst ` toSet \'" +using assms proof(induct \' rule: \.induct) + case GNil + then show ?case by simp +next + case (GCons xbc \') + then obtain x' and b' and c'::c where xbc: "xbc=(x',b',c')" + using prod_cases3 by blast + hence *:"(xbc #\<^sub>\ \') @ (x, b1, c1) #\<^sub>\ \ = ((x',b',c') #\<^sub>\(\'@ ((x, b1, c1) #\<^sub>\ \)))" by simp + hence "atom x' \ (\'@(x,b1,c1) #\<^sub>\\)" using wfG_elims(2) GCons by metis + + moreover have "\ ; B \\<^sub>w\<^sub>f (\' @ (x, b1, c1) #\<^sub>\ \)" using GCons wfG_elims * by metis + ultimately have "atom x' \ atom_dom (\'@(x,b1,c1) #\<^sub>\\)" using wfG_dom_supp GCons append_g.simps xbc fresh_def by fast + hence "x' \ x" using GCons fresh_GCons xbc by fastforce + then show ?case using GCons xbc toSet.simps + using Un_commute \\ ; B \\<^sub>w\<^sub>f \' @ (x, b1, c1) #\<^sub>\ \\ atom_dom.simps by auto +qed + +lemma lookup_inside_wf[simp]: + assumes "\ ; B \\<^sub>w\<^sub>f (\'@(x,b1,c1) #\<^sub>\\)" + shows "Some (b1,c1) = lookup (\'@(x,b1,c1) #\<^sub>\\) x" + using wf_not_in_prefix lookup_inside assms by fast + +lemma lookup_weakening: + fixes \::\ and \::\ and \'::\ + assumes "Some (b,c) = lookup \ x" and "toSet \ \ toSet \'" and "\; \ \\<^sub>w\<^sub>f \'" and "\; \ \\<^sub>w\<^sub>f \" + shows "Some (b,c) = lookup \' x" +proof - + have "(x,b,c) \ toSet \ \ (\b' c'. (x,b',c') \ toSet \ \ b'=b \ c'=c)" using assms lookup_iff toSet.simps by force + hence "(x,b,c) \ toSet \'" using assms by auto + moreover have "(\b' c'. (x,b',c') \ toSet \' \ b'=b \ c'=c)" using assms wf_g_unique + using calculation by auto + ultimately show ?thesis using lookup_iff + using assms(3) by blast +qed + +lemma wfPhi_lookup_fun_unique: + fixes \::\ + assumes "\ \\<^sub>w\<^sub>f \" and "AF_fundef f fd \ set \" + shows "Some (AF_fundef f fd) = lookup_fun \ f" +using assms proof(induct \ rule: list.induct ) + case Nil + then show ?case using lookup_fun.simps by simp +next + case (Cons a \') + then obtain f' and fd' where a:"a = AF_fundef f' fd'" using fun_def.exhaust by auto + have wf: "\ \\<^sub>w\<^sub>f \' \ f' \ name_of_fun ` set \' " using wfPhi_elims Cons a by metis + then show ?case using Cons lookup_fun.simps using Cons lookup_fun.simps wf a + by (metis image_eqI name_of_fun.simps set_ConsD) +qed + +lemma lookup_fun_weakening: + fixes \'::\ + assumes "Some fd = lookup_fun \ f" and "set \ \ set \'" and "\ \\<^sub>w\<^sub>f \'" + shows "Some fd = lookup_fun \' f" +using assms proof(induct \ ) + case Nil + then show ?case using lookup_fun.simps by simp +next + case (Cons a \'') + then obtain f' and fd' where a: "a = AF_fundef f' fd'" using fun_def.exhaust by auto + then show ?case proof(cases "f=f'") + case True + then show ?thesis using lookup_fun.simps Cons wfPhi_lookup_fun_unique a + by (metis lookup_fun_member subset_iff) + next + case False + then show ?thesis using lookup_fun.simps Cons + using \a = AF_fundef f' fd'\ by auto + qed +qed + +lemma fundef_poly_fresh_bv: + assumes "atom bv2 \ (bv1,b1,c1,\1,s1)" + shows * : "(AF_fun_typ_some bv2 (AF_fun_typ x1 ((bv1\bv2) \ b1) ((bv1\bv2) \c1) ((bv1\bv2) \ \1) ((bv1\bv2) \ s1)) = (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s1)))" + (is "(AF_fun_typ_some ?bv ?fun_typ = AF_fun_typ_some ?bva ?fun_typa)") +proof - + have 1:"atom bv2 \ set [atom x1]" using bv_not_in_x_atoms by simp + have 2:"bv1 \ bv2" using assms by auto + have 3:"(bv2 \ bv1) \ x1 = x1" using pure_fresh flip_fresh_fresh + by (simp add: flip_fresh_fresh) + have " AF_fun_typ x1 ((bv1 \ bv2) \ b1) ((bv1 \ bv2) \ c1) ((bv1 \ bv2) \ \1) ((bv1 \ bv2) \ s1) = (bv2 \ bv1) \ AF_fun_typ x1 b1 c1 \1 s1" + using 1 2 3 assms by (simp add: flip_commute) + moreover have "(atom bv2 \ c1 \ atom bv2 \ \1 \ atom bv2 \ s1 \ atom bv2 \ set [atom x1]) \ atom bv2 \ b1" + using 1 2 3 assms fresh_prod5 by metis + ultimately show ?thesis unfolding fun_typ_q.eq_iff Abs1_eq_iff(3) fun_typ.fresh 1 2 by fastforce +qed + + +text \It is possible to collapse some of the easy to prove inductive cases into a single proof at the qed line + but this makes it fragile under change. For example, changing the lemma statement might make one of the previously + trivial cases non-trivial and so the collapsing needs to be unpacked. Is there a way to find which case + has failed in the qed line?\ + +lemma wb_b_weakening1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + + shows "\; \; \ \\<^sub>w\<^sub>f v : b \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f v : b" and + "\; \; \ \\<^sub>w\<^sub>f c \\ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f c" and + "\; \ \\<^sub>w\<^sub>f \ \\ |\| \' \ \; \' \\<^sub>w\<^sub>f \ " and + "\; \; \ \\<^sub>w\<^sub>f \ \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f ts" and + "\\<^sub>w\<^sub>f P \ True " and + "wfB \ \ b \ \ |\| \' \ wfB \ \' b" and + "\; \; \ \\<^sub>w\<^sub>f ce : b \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f ce : b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct b and c and \ and \ and ts and P and b and b and td + avoiding: \' +rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_conspI s bv dclist \ dc x b' c \ b \ v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by metis + show \(dc, \ x : b' | c \) \ set dclist\ using wfV_conspI by auto + show \ \ ; \' \\<^sub>w\<^sub>f b \ using wfV_conspI by auto + show \atom bv \ (\, \', \, b, v)\ using fresh_prodN wfV_conspI by auto + thus \ \; \' ; \ \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by simp + qed +next + case (wfTI z \ \ \ b c) + show ?case proof + show "atom z \ (\, \', \)" using wfTI by auto + show "\; \' \\<^sub>w\<^sub>f b " using wfTI by auto + show "\; \' ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c " using wfTI by auto + qed +qed( (auto simp add: wf_intros | metis wf_intros)+ ) + +lemma wb_b_weakening2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + + shows + "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f e : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f s : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \ |\| \' \ \ ; \ ; \' ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ \ |\| \' \ \; \' ; \ \\<^sub>w\<^sub>f \" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ |\| \' \ \ ; \ ; \' \\<^sub>w\<^sub>f ft" +proof(nominal_induct b and b and b and b and \ and \ and ftq and ft + avoiding: \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_intros wb_b_weakening1 + by meson +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using Wellformed.wfE_fstI wb_b_weakening1 by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfE_appI \ \ \ \ \ f ft v) + then show ?case using wf_intros using wb_b_weakening1 by meson +next + case (wfE_appPI \ \ \1 \ \ b' bv1 v1 \1 f1 x1 b1 c1 s1) + + have "\ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f AE_appP f1 b' v1 : (b_of \1)[bv1::=b']\<^sub>b" + proof + show "\ \\<^sub>w\<^sub>f \" using wfE_appPI by auto + show "\; \' ; \ \\<^sub>w\<^sub>f \ " using wfE_appPI by auto + show "\; \' \\<^sub>w\<^sub>f b' " using wfE_appPI wb_b_weakening1 by auto + thus " atom bv1 \ (\, \, \', \, \, b', v1, (b_of \1)[bv1::=b']\<^sub>b)" + using wfE_appPI fresh_prodN by auto + + show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s1))) = lookup_fun \ f1" using wfE_appPI by auto + show "\; \' ; \ \\<^sub>w\<^sub>f v1 : b1[bv1::=b']\<^sub>b " using wfE_appPI wb_b_weakening1 by auto + qed + then show ?case by auto +next + case (wfE_mvarI \ \ \ \ \ u \) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfS_letI \ \ \ \ \ e b' x s b) + show ?case proof + show \ \ ; \ ; \' ; \ ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto + show \ \ ; \ ; \' ; (x, b', TRUE) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by auto + show \ \; \' ; \ \\<^sub>w\<^sub>f \ \ using wfS_letI by auto + show \atom x \ (\, \, \', \, \, e, b)\ using wfS_letI by auto + qed +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + then show ?case using wb_b_weakening1 Wellformed.wfS_let2I by simp +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wb_b_weakening1 Wellformed.wfS_ifI by simp +next + case (wfS_varI \ \ \ \ v u \ \ s b) + then show ?case using wb_b_weakening1 Wellformed.wfS_varI by simp +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wb_b_weakening1 Wellformed.wfS_assignI by simp +next +case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wb_b_weakening1 Wellformed.wfS_whileI by simp +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using Wellformed.wfS_seqI by metis +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + then show ?case using wb_b_weakening1 Wellformed.wfS_matchI by metis +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + then show ?case using Wellformed.wfS_branchI by auto +next + case (wfS_finalI \ \ \ \ \ tid dclist' cs b dclist) + then show ?case using wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dclist' cs b css dclist) + then show ?case using wf_intros by metis +next + case (wfD_emptyI \ \ \) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfD_cons \ \ \ \ \ u) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfPhi_emptyI \) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfPhi_consI f \ \ ft) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfFTSome \ bv ft) + then show ?case using wf_intros wb_b_weakening1 by metis +next + case (wfFTI \ B b s x c \ \) + show ?case proof + show "\; \' \\<^sub>w\<^sub>f b" using wfFTI wb_b_weakening1 by auto + + show "supp c \ {atom x}" using wfFTI wb_b_weakening1 by auto + show "\; \' ; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ " using wfFTI wb_b_weakening1 by auto + show "\ \\<^sub>w\<^sub>f \ " using wfFTI wb_b_weakening1 by auto + from \ B |\| \'\ have "supp B \ supp \'" proof(induct B) + case empty + then show ?case by auto + next + case (insert x B) + then show ?case + by (metis fsubset_funion_eq subset_Un_eq supp_union_fset) + qed + thus "supp s \ {atom x} \ supp \'" using wfFTI by auto + qed +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case proof + show \ \ ; \ ; \' ; (x, B_bool, c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wb_b_weakening1 wfS_assertI by simp + show \ \; \' ; \ \\<^sub>w\<^sub>f c \ using wb_b_weakening1 wfS_assertI by simp + show \ \; \' ; \ \\<^sub>w\<^sub>f \ \ using wb_b_weakening1 wfS_assertI by simp + have "atom x \ \'" using x_not_in_b_set fresh_def by metis + thus \atom x \ (\, \, \', \, \, c, b, s)\ using wfS_assertI fresh_prodN by simp + qed +qed(auto) + +lemmas wb_b_weakening = wb_b_weakening1 wb_b_weakening2 + +lemma wfG_b_weakening: + fixes \::\ + assumes "\ |\| \'" and "\; \ \\<^sub>w\<^sub>f \" + shows "\; \' \\<^sub>w\<^sub>f \ " + using wb_b_weakening assms by auto + +lemma wfT_b_weakening: + fixes \::\ and \::\ and \::\ + assumes "\ |\| \'" and "\; \; \ \\<^sub>w\<^sub>f \" + shows "\; \' ; \ \\<^sub>w\<^sub>f \ " + using wb_b_weakening assms by auto + +lemma wfB_subst_wfB: + fixes \::\ and b'::b and b::b + assumes "\ ; {|bv|} \\<^sub>w\<^sub>f b" and "\; \ \\<^sub>w\<^sub>f b'" + shows "\; \ \\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b " +using assms proof(nominal_induct b rule:b.strong_induct) + case B_int + hence "\ ; {||} \\<^sub>w\<^sub>f B_int" using wfB_intI wfX_wfY by fast + then show ?case using subst_bb.simps wb_b_weakening by fastforce +next + case B_bool + hence "\ ; {||} \\<^sub>w\<^sub>f B_bool" using wfB_boolI wfX_wfY by fast + then show ?case using subst_bb.simps wb_b_weakening by fastforce +next + case (B_id x ) + hence " \; \ \\<^sub>w\<^sub>f (B_id x)" using wfB_consI wfB_elims wfX_wfY by metis + then show ?case using subst_bb.simps(4) by auto +next + case (B_pair x1 x2) + then show ?case using subst_bb.simps + by (metis wfB_elims(1) wfB_pairI) +next + case B_unit + hence "\ ; {||} \\<^sub>w\<^sub>f B_unit" using wfB_unitI wfX_wfY by fast + then show ?case using subst_bb.simps wb_b_weakening by fastforce +next + case B_bitvec + hence "\ ; {||} \\<^sub>w\<^sub>f B_bitvec" using wfB_bitvecI wfX_wfY by fast + then show ?case using subst_bb.simps wb_b_weakening by fastforce +next + case (B_var x) + then show ?case + proof - + have False + using B_var.prems(1) wfB.cases by fastforce (* 781 ms *) + then show ?thesis by metis + qed +next + case (B_app s b) + then obtain bv' dclist where *:"AF_typedef_poly s bv' dclist \ set \ \ \ ; {|bv|} \\<^sub>w\<^sub>f b" using wfB_elims by metis + show ?case unfolding subst_b_simps proof + show "\\<^sub>w\<^sub>f \ " using B_app wfX_wfY by metis + show "\ ; \ \\<^sub>w\<^sub>f b[bv::=b']\<^sub>b\<^sub>b " using * B_app forget_subst wfB_supp fresh_def + by (metis ex_in_conv subset_empty subst_b_b_def supp_empty_fset) + show "AF_typedef_poly s bv' dclist \ set \" using * by auto + qed +qed + +lemma wfT_subst_wfB: + fixes \::\ and b'::b + assumes "\ ; {|bv|} ; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \" and "\; \ \\<^sub>w\<^sub>f b'" + shows "\; \ \\<^sub>w\<^sub>f (b_of \)[bv::=b']\<^sub>b\<^sub>b " +proof - + obtain b where "\ ; {|bv|} \\<^sub>w\<^sub>f b \ b_of \ = b" using wfT_elims b_of.simps assms by metis + thus ?thesis using wfB_subst_wfB assms by auto +qed + +lemma wfG_cons_unique: + assumes "(x1,b1,c1) \ toSet (((x,b,c) #\<^sub>\\))" and "wfG \ \ (((x,b,c) #\<^sub>\\))" and "x = x1" + shows "b1 = b \ c1 = c" +proof - + have "x1 \ fst ` toSet \" + proof - + have "atom x1 \ \" using assms wfG_cons by metis + then show ?thesis + using fresh_gamma_elem + by (metis assms(2) atom_dom.simps dom.simps rev_image_eqI wfG_cons2 wfG_x_fresh) + qed + thus ?thesis using assms by force +qed + +lemma wfG_member_unique: + assumes "(x1,b1,c1) \ toSet (\'@((x,b,c) #\<^sub>\\))" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" and "x = x1" + shows "b1 = b \ c1 = c" + using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using wfG_suffix wfG_cons_unique append_g.simps by metis +next + case (GCons x' b' c' \') + moreover hence "(x1, b1, c1) \ toSet (\' @ (x, b, c) #\<^sub>\ \)" using wf_not_in_prefix by fastforce + ultimately show ?case using wfG_cons by fastforce +qed + +section \Function Definitions\ + +lemma wb_phi_weakening: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list and \::\ + shows + "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ \\<^sub>w\<^sub>f e : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ True" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' \\<^sub>w\<^sub>f ftq" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ \\<^sub>w\<^sub>f \' \ set \ \ set \' \ \ ; \' ; \ \\<^sub>w\<^sub>f ft" +proof(nominal_induct + b and b and b and b and \ and \ and ftq and ft + avoiding: \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_intros by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_intros by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_intros by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then show ?case using wf_intros lookup_fun_weakening by metis +next + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \' \ using wfE_appPI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto + show \atom bv \ (\', \, \, \, \, b', v, (b_of \)[bv::=b']\<^sub>b)\ using wfE_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \' f\ + using wfE_appPI lookup_fun_weakening by metis + show \ \; \; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI by auto + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + then show ?case using wf_intros by metis +next + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_intros by metis +next + case (wfS_letI \ \ \ \ \ e b' x s b) + then show ?case using Wellformed.wfS_letI by fastforce +next + case (wfS_let2I \ \ \ \ \ s1 b' x s2 b) + then show ?case using Wellformed.wfS_let2I by fastforce +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wf_intros by metis +next + case (wfS_varI \ \ \ \ v u \ \ b s) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_varI by simp + show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI by simp + show \atom u \ (\', \, \, \, \, \, v, b)\ using wfS_varI by simp + show \ \ ; \' ; \ ; \ ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b \ using wfS_varI by simp + qed +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wf_intros by metis +next + case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros by metis +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros by metis +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + then show ?case using wf_intros by metis +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + then show ?case using Wellformed.wfS_branchI by fastforce +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case proof + show \ \ ; \' ; \ ; (x, B_bool, c) #\<^sub>\ \ ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f c \ using wfS_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_assertI by auto + have "atom x \ \'" using wfS_assertI wfPhi_supp fresh_def by blast + thus \atom x \ (\', \, \, \, \, c, b, s)\ using fresh_prodN wfS_assertI wfPhi_supp fresh_def by auto + qed +next + case (wfFTI \ B b s x c \ \) + show ?case proof + show \ \ ; B \\<^sub>w\<^sub>f b \ using wfFTI by auto + next + show \supp c \ {atom x}\ using wfFTI by auto + next + show \ \ ; B ; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ using wfFTI by auto + next + show \ \ \\<^sub>w\<^sub>f \' \ using wfFTI by auto + next + show \supp s \ {atom x} \ supp B\ using wfFTI by auto + qed +qed(auto|metis wf_intros)+ + + +lemma wfT_fun_return_t: + fixes \a'::\ and \'::\ + assumes "\; \; (xa, b, ca) #\<^sub>\ GNil \\<^sub>w\<^sub>f \a'" and "(AF_fun_typ x b c \' s') = (AF_fun_typ xa b ca \a' sa')" + shows "\; \; (x, b, c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \'" +proof - + obtain cb::x where xf: "atom cb \ (c, \', s', sa', \a', ca, x , xa)" using obtain_fresh by blast + hence "atom cb \ (c, \', s', sa', \a', ca) \ atom cb \ (x, xa, ((c, \'), s'), (ca, \a'), sa')" using fresh_prod6 fresh_prod4 fresh_prod8 by auto + hence *:"c[x::=V_var cb]\<^sub>c\<^sub>v = ca[xa::=V_var cb]\<^sub>c\<^sub>v \ \'[x::=V_var cb]\<^sub>\\<^sub>v = \a'[xa::=V_var cb]\<^sub>\\<^sub>v" using assms \.eq_iff Abs1_eq_iff_all by auto + + have **: "\; \; (xa \ cb ) \ ((xa, b, ca) #\<^sub>\ GNil) \\<^sub>w\<^sub>f (xa \ cb ) \ \a'" using assms True_eqvt beta_flip_eq theta_flip_eq wfG_wf + by (metis GCons_eqvt GNil_eqvt wfT.eqvt wfT_wf) + + have "\; \; (x \ cb ) \ ((x, b, c) #\<^sub>\ GNil) \\<^sub>w\<^sub>f (x \ cb ) \ \'" proof - + have "(xa \ cb ) \ xa = (x \ cb ) \ x" using xf by auto + hence "(x \ cb ) \ ((x, b, c) #\<^sub>\ GNil) = (xa \ cb ) \ ((xa, b, ca) #\<^sub>\ GNil)" using * ** xf G_cons_flip fresh_GNil by simp + thus ?thesis using ** * xf by simp + qed + thus ?thesis using beta_flip_eq theta_flip_eq wfT_wf wfG_wf * ** True_eqvt wfT.eqvt permute_flip_cancel by metis +qed + +lemma wfFT_wf_aux: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q and s::s and \::\ + assumes "\ ; \ ; B \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" + shows "\ ; B ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ supp s \ { atom x } \ supp B" +proof - + obtain xa and ca and sa and \' where *:"\ ; B \\<^sub>w\<^sub>f b \ (\ \\<^sub>w\<^sub>f \ ) \ + supp sa \ {atom xa} \ supp B \ (\ ; B ; (xa, b, ca) #\<^sub>\ GNil \\<^sub>w\<^sub>f \') \ + AF_fun_typ x b c \ s = AF_fun_typ xa b ca \' sa " + using wfFT.simps[of \ \ B "AF_fun_typ x b c \ s"] assms by auto + + moreover hence **: "(AF_fun_typ x b c \ s) = (AF_fun_typ xa b ca \' sa)" by simp + ultimately have "\ ; B ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfT_fun_return_t by metis + moreover have " (\ \\<^sub>w\<^sub>f \ ) " using * by auto + moreover have "supp s \ { atom x } \ supp B" proof - + have "[[atom x]]lst.s = [[atom xa]]lst.sa" using ** fun_typ.eq_iff lst_fst lst_snd by metis + thus ?thesis using lst_supp_subset * by metis + qed + ultimately show ?thesis by auto +qed + +lemma wfFT_simple_wf: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q and s::s and \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_none (AF_fun_typ x b c \ s))" + shows "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ supp s \ { atom x } " +proof - + have *:"\ ; \ ; {||} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" using wfFTQ_elims assms by auto + thus ?thesis using wfFT_wf_aux by force +qed + +lemma wfFT_poly_wf: + fixes \::\ and \::\ and \::\ and ftq :: fun_typ_q and s::s and \::\ + assumes "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" + shows "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ \ ; \ ; {|bv|} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" +proof - + obtain bv1 ft1 where *:"\ ; \ ; {|bv1|} \\<^sub>w\<^sub>f ft1 \ [[atom bv1]]lst. ft1 = [[atom bv]]lst. AF_fun_typ x b c \ s" + using wfFTQ_elims(3)[OF assms] by metis + + show ?thesis proof(cases "bv1 = bv") + case True + then show ?thesis using * fun_typ_q.eq_iff Abs1_eq_iff by (metis (no_types, hide_lams) wfFT_wf_aux) + next + case False + obtain x1 b1 c1 t1 s1 where **: "ft1 = AF_fun_typ x1 b1 c1 t1 s1" using fun_typ.eq_iff + by (meson fun_typ.exhaust) + + hence eqv: "(bv \ bv1) \ AF_fun_typ x1 b1 c1 t1 s1 = AF_fun_typ x b c \ s \ atom bv1 \ AF_fun_typ x b c \ s" using + Abs1_eq_iff(3) * False by metis + + have "(bv \ bv1) \ \ ; (bv \ bv1) \ \ ; (bv \ bv1) \ {|bv1|} \\<^sub>w\<^sub>f (bv \ bv1) \ ft1" using wfFT.eqvt * by metis + moreover have "(bv \ bv1) \ \ = \" using phi_flip_eq wfX_wfY * by metis + moreover have "(bv \ bv1) \ \ =\" using wfX_wfY * theta_flip_eq2 by metis + moreover have "(bv \ bv1) \ ft1 = AF_fun_typ x b c \ s" using eqv ** by metis + ultimately have "\ ; \ ; {|bv|} \\<^sub>w\<^sub>f AF_fun_typ x b c \ s" by auto + thus ?thesis using wfFT_wf_aux by auto + qed +qed + +lemma wfFT_poly_wfT: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "\ ; \ \\<^sub>w\<^sub>f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))" + shows "\ ; {| bv |} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" + using wfFT_poly_wf assms by simp + +lemma b_of_supp: + "supp (b_of t) \ supp t" +proof(nominal_induct t rule:\.strong_induct) + case (T_refined_type x b c) + then show ?case by auto +qed + +lemma wfPhi_f_simple_wf: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q and s::s and \'::\ + assumes "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) \ set \ " and "\ \\<^sub>w\<^sub>f \" and "set \ \ set \'" and "\ \\<^sub>w\<^sub>f \'" + shows "\ ; {||} ; (x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \ \ supp s \ { atom x }" +using assms proof(induct \ rule: \_induct) + case PNil + then show ?case by auto +next + case (PConsSome f1 bv x1 b1 c1 \1 s' \'') + hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) \ set \''" by auto + moreover have " \ \\<^sub>w\<^sub>f \'' \ set \'' \ set \'" using wfPhi_elims(3) PConsSome by auto + ultimately show ?case using PConsSome wfPhi_elims wfFT_simple_wf by auto +next + case (PConsNone f' x' b' c' \' s' \'') + show ?case proof(cases "f=f'") + case True + have "AF_fun_typ_none (AF_fun_typ x' b' c' \' s') = AF_fun_typ_none (AF_fun_typ x b c \ s)" + by (metis PConsNone.prems(1) PConsNone.prems(2) True fun_def.eq_iff image_eqI name_of_fun.simps set_ConsD wfPhi_elims(2)) + hence *:"\ ; \'' \\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x b c \ s) " using wfPhi_elims(2)[OF PConsNone(3)] by metis + hence "\ ; \'' ; {||} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" using wfFTQ_elims(1) by metis + thus ?thesis using wfFT_simple_wf[OF *] wb_phi_weakening PConsNone by force + next + case False + hence "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) \ set \''" using PConsNone by simp + moreover have " \ \\<^sub>w\<^sub>f \'' \ set \'' \ set \'" using wfPhi_elims(3) PConsNone by auto + ultimately show ?thesis using PConsNone wfPhi_elims wfFT_simple_wf by auto + qed +qed + +lemma wfPhi_f_simple_wfT: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" + using wfPhi_f_simple_wf assms using lookup_fun_member by blast + +lemma wfPhi_f_simple_supp_b: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp b = {}" +proof - + have "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfPhi_f_simple_wfT assms by auto + thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce +qed + +lemma wfPhi_f_simple_supp_t: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp \ \ { atom x }" + using wfPhi_f_simple_wfT wfT_supp assms by fastforce + +lemma wfPhi_f_simple_supp_c: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp c \ { atom x }" +proof - + have "\ ; {||} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfPhi_f_simple_wfT assms by auto + thus ?thesis using wfG_wfC wfC_supp wfT_wf by fastforce +qed + +lemma wfPhi_f_simple_supp_s: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp s \ {atom x}" +proof - + have "AF_fundef f (AF_fun_typ_none (AF_fun_typ x b c \ s)) \ set \" using lookup_fun_member assms by auto + hence "supp s \ { atom x }" using wfPhi_f_simple_wf assms by blast + thus ?thesis using wf_supp(3) atom_dom.simps toSet.simps x_not_in_u_set x_not_in_b_set setD.simps + using wf_supp2(2) by fastforce +qed + +lemma wfPhi_f_poly_wf: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q and s::s and \'::\ + assumes "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)) \ set \ " and "\ \\<^sub>w\<^sub>f \" and "set \ \ set \'" and "\ \\<^sub>w\<^sub>f \'" + shows "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \ \ \ \\<^sub>w\<^sub>f \' \ \ ; \' ; {|bv|} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" +using assms proof(induct \ rule: \_induct) + case PNil + then show ?case by auto +next + case (PConsNone f x b c \ s' \'') + moreover have " \ \\<^sub>w\<^sub>f \'' \ set \'' \ set \'" using wfPhi_elims(3) PConsNone by auto + ultimately show ?case using PConsNone wfPhi_elims wfFT_poly_wf by auto +next + case (PConsSome f1 bv1 x1 b1 c1 \1 s1 \'') + show ?case proof(cases "f=f1") + case True + have "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s1) = AF_fun_typ_some bv (AF_fun_typ x b c \ s)" + by (metis PConsSome.prems(1) PConsSome.prems(2) True fun_def.eq_iff list.set_intros(1) option.inject wfPhi_lookup_fun_unique) + hence *:"\ ; \'' \\<^sub>w\<^sub>f AF_fun_typ_some bv (AF_fun_typ x b c \ s) " using wfPhi_elims PConsSome by metis + thus ?thesis using wfFT_poly_wf * wb_phi_weakening PConsSome + by (meson set_subset_Cons) + next + case False + hence "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)) \ set \''" using PConsSome + by (meson fun_def.eq_iff set_ConsD) + moreover have " \ \\<^sub>w\<^sub>f \'' \ set \'' \ set \'" using wfPhi_elims(3) PConsSome + by (meson dual_order.trans set_subset_Cons) + ultimately show ?thesis using PConsSome wfPhi_elims wfFT_poly_wf + by blast + qed +qed + +lemma wfPhi_f_poly_wfT: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "\ ; {| bv |} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" +using assms proof(induct \ rule: \_induct) + case PNil + then show ?case by auto +next + case (PConsSome f1 bv1 x1 b1 c1 \1 s' \') + then show ?case proof(cases "f1=f") + case True + hence "lookup_fun (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s')) # \') f = Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s')))" using + lookup_fun.simps using PConsSome.prems by simp + then show ?thesis using PConsSome.prems wfPhi_elims wfFT_poly_wfT + by (metis option.inject) + next + case False + then show ?thesis using PConsSome using lookup_fun.simps + using wfPhi_elims(3) by auto + qed +next + case (PConsNone f' x' b' c' \' s' \') + then show ?case proof(cases "f'=f") + case True + then have *:"\ ; \' \\<^sub>w\<^sub>f AF_fun_typ_none (AF_fun_typ x' b' c' \' s') " using lookup_fun.simps PConsNone wfPhi_elims by metis + thus ?thesis using PConsNone wfFT_poly_wfT wfPhi_elims lookup_fun.simps + by (metis fun_def.eq_iff fun_typ_q.distinct(1) option.inject) + next + case False + thus ?thesis using PConsNone wfPhi_elims + by (metis False lookup_fun.simps(2)) + qed +qed + +lemma wfPhi_f_poly_supp_b: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp b \ supp bv" +proof - + have "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfPhi_f_poly_wfT assms by auto + thus ?thesis using wfT_wf wfG_cons wfB_supp by fastforce +qed + +lemma wfPhi_f_poly_supp_t: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp \ \ { atom x , atom bv }" + using wfPhi_f_poly_wfT[OF assms, THEN wfT_supp] atom_dom.simps supp_at_base by auto + + +lemma wfPhi_f_poly_supp_b_of_t: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp (b_of \) \ { atom bv }" +proof - + have "atom x \ supp (b_of \)" using x_fresh_b by auto + moreover have "supp (b_of \) \ { atom x , atom bv }" using wfPhi_f_poly_supp_t + using supp_at_base b_of.simps wfPhi_f_poly_supp_t \.supp b_of_supp assms by fast + ultimately show ?thesis by blast +qed + +lemma wfPhi_f_poly_supp_c: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp c \ { atom x, atom bv }" +proof - + have "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" using wfPhi_f_poly_wfT assms by auto + thus ?thesis using wfG_wfC wfC_supp wfT_wf + using supp_at_base by fastforce +qed + +lemma wfPhi_f_poly_supp_s: + fixes \::\ and \::\ and \::\ and ft :: fun_typ_q + assumes "Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f" and "\ \\<^sub>w\<^sub>f \" + shows "supp s \ {atom x, atom bv}" +proof - + + have "AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s)) \ set \" using lookup_fun_member assms by auto + hence *:"\ ; \ ; {|bv|} \\<^sub>w\<^sub>f (AF_fun_typ x b c \ s)" using assms wfPhi_f_poly_wf by simp + + thus ?thesis using wfFT_wf_aux[OF *] using supp_at_base by auto +qed + +lemmas wfPhi_f_supp = wfPhi_f_poly_supp_b wfPhi_f_simple_supp_b wfPhi_f_poly_supp_c + wfPhi_f_simple_supp_t wfPhi_f_poly_supp_t wfPhi_f_simple_supp_t wfPhi_f_poly_wfT wfPhi_f_simple_wfT + wfPhi_f_poly_supp_s wfPhi_f_simple_supp_s + +lemma fun_typ_eq_ret_unique: + assumes "(AF_fun_typ x1 b1 c1 \1' s1') = (AF_fun_typ x2 b2 c2 \2' s2')" + shows "\1'[x1::=v]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>v" +proof - + have "[[atom x1]]lst. \1' = [[atom x2]]lst. \2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis + thus ?thesis using subst_v_flip_eq_two[of x1 \1' x2 \2' v] subst_v_\_def by metis +qed + +lemma fun_typ_eq_body_unique: + fixes v::v and x1::x and x2::x and s1'::s and s2'::s + assumes "(AF_fun_typ x1 b1 c1 \1' s1') = (AF_fun_typ x2 b2 c2 \2' s2')" + shows "s1'[x1::=v]\<^sub>s\<^sub>v = s2'[x2::=v]\<^sub>s\<^sub>v" +proof - + have "[[atom x1]]lst. s1' = [[atom x2]]lst. s2'" using assms lst_fst fun_typ.eq_iff lst_snd by metis + thus ?thesis using subst_v_flip_eq_two[of x1 s1' x2 s2' v] subst_v_s_def by metis +qed + +lemma fun_ret_unique: + assumes "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f" + shows "\1'[x1::=v]\<^sub>\\<^sub>v = \2'[x2::=v]\<^sub>\\<^sub>v" +proof - + have *: " (AF_fundef f (AF_fun_typ_none (AF_fun_typ x1 b1 c1 \1' s1'))) = (AF_fundef f (AF_fun_typ_none (AF_fun_typ x2 b2 c2 \2' s2')))" using option.inject assms by metis + thus ?thesis using fun_typ_eq_ret_unique fun_def.eq_iff fun_typ_q.eq_iff by metis +qed + +lemma fun_poly_arg_unique: + fixes bv1::bv and bv2::bv and b::b and \1::\ and \2::\ + assumes "[[atom bv1]]lst. (AF_fun_typ x1 b1 c1 \1 s1) = [[atom bv2]]lst. (AF_fun_typ x2 b2 c2 \2 s2)" (is "[[atom ?x]]lst. ?a = [[atom ?y]]lst. ?b") + shows "\ x1 : b1[bv1::=b]\<^sub>b\<^sub>b | c1[bv1::=b]\<^sub>c\<^sub>b \ = \ x2 : b2[bv2::=b]\<^sub>b\<^sub>b | c2[bv2::=b]\<^sub>c\<^sub>b \" +proof - + obtain c::bv where *:"atom c \ (b,b1,b2,c1,c2) \ atom c \ (bv1, bv2, AF_fun_typ x1 b1 c1 \1 s1, AF_fun_typ x2 b2 c2 \2 s2)" using obtain_fresh fresh_Pair by metis + hence "(bv1 \ c) \ AF_fun_typ x1 b1 c1 \1 s1 = (bv2 \ c) \ AF_fun_typ x2 b2 c2 \2 s2" using Abs1_eq_iff_all(3)[of ?x ?a ?y ?b] assms by metis + hence "AF_fun_typ x1 ((bv1 \ c) \ b1) ((bv1 \ c) \ c1) ((bv1 \ c) \ \1) ((bv1 \ c) \ s1) = AF_fun_typ x2 ((bv2 \ c) \ b2) ((bv2 \ c) \ c2) ((bv2 \ c) \ \2) ((bv2 \ c) \ s2)" + using fun_typ_flip by metis + hence **:"\ x1 :((bv1 \ c) \ b1) | ((bv1 \ c) \ c1) \ = \ x2 : ((bv2 \ c) \ b2) | ((bv2 \ c) \ c2) \" (is "\ x1 : ?b1 | ?c1 \ = \ x2 : ?b2 | ?c2 \") using fun_arg_unique_aux by metis + hence "\ x1 :((bv1 \ c) \ b1) | ((bv1 \ c) \ c1) \[c::=b]\<^sub>\\<^sub>b = \ x2 : ((bv2 \ c) \ b2) | ((bv2 \ c) \ c2) \[c::=b]\<^sub>\\<^sub>b" by metis + hence "\ x1 :((bv1 \ c) \ b1)[c::=b]\<^sub>b\<^sub>b | ((bv1 \ c) \ c1)[c::=b]\<^sub>c\<^sub>b \ = \ x2 : ((bv2 \ c) \ b2)[c::=b]\<^sub>b\<^sub>b | ((bv2 \ c) \ c2)[c::=b]\<^sub>c\<^sub>b \" using subst_tb.simps by metis + thus ?thesis using * flip_subst_subst subst_b_c_def subst_b_b_def fresh_prodN flip_commute by metis +qed + +lemma fun_poly_ret_unique: + assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f" + shows "\1'[bv1::=b]\<^sub>\\<^sub>b[x1::=v]\<^sub>\\<^sub>v = \2'[bv2::=b]\<^sub>\\<^sub>b[x2::=v]\<^sub>\\<^sub>v" +proof - + have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2')))" using option.inject assms by metis + hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2')" + (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis + hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis + + hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis + have "[[atom x1]]lst. \1'[bv1::=b]\<^sub>\\<^sub>b = [[atom x2]]lst. \2'[bv2::=b]\<^sub>\\<^sub>b" + apply(rule lst_snd[of _ "c1[bv1::=b]\<^sub>c\<^sub>b" _ _ "c2[bv2::=b]\<^sub>c\<^sub>b"]) + apply(rule lst_fst[of _ _ "s1'[bv1::=b]\<^sub>s\<^sub>b" _ _ "s2'[bv2::=b]\<^sub>s\<^sub>b"]) + using * subst_ft_b.simps fun_typ.eq_iff by metis + thus ?thesis using subst_v_flip_eq_two subst_v_\_def by metis +qed + +lemma fun_poly_body_unique: + assumes "Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = lookup_fun \ f" and "Some (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2'))) = lookup_fun \ f" + shows "s1'[bv1::=b]\<^sub>s\<^sub>b[x1::=v]\<^sub>s\<^sub>v = s2'[bv2::=b]\<^sub>s\<^sub>b[x2::=v]\<^sub>s\<^sub>v" +proof - + have *: " (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1'))) = (AF_fundef f (AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2')))" + using option.inject assms by metis + hence "AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1' s1') = AF_fun_typ_some bv2 (AF_fun_typ x2 b2 c2 \2' s2')" + (is "AF_fun_typ_some bv1 ?ft1 = AF_fun_typ_some bv2 ?ft2") using fun_def.eq_iff by metis + hence **:"[[atom bv1]]lst. ?ft1 = [[atom bv2]]lst. ?ft2" using fun_typ_q.eq_iff(1) by metis + + hence *:"subst_ft_b ?ft1 bv1 b = subst_ft_b ?ft2 bv2 b" using subst_b_flip_eq_two subst_b_fun_typ_def by metis + have "[[atom x1]]lst. s1'[bv1::=b]\<^sub>s\<^sub>b = [[atom x2]]lst. s2'[bv2::=b]\<^sub>s\<^sub>b" + using lst_snd lst_fst subst_ft_b.simps fun_typ.eq_iff + by (metis "local.*") + + thus ?thesis using subst_v_flip_eq_two subst_v_s_def by metis +qed + +lemma funtyp_eq_iff_equalities: + fixes s'::s and s::s + assumes " [[atom x']]lst. ((c', \'), s') = [[atom x]]lst. ((c, \), s)" + shows "\ x' : b | c' \ = \ x : b | c \ \ s'[x'::=v]\<^sub>s\<^sub>v = s[x::=v]\<^sub>s\<^sub>v \ \'[x'::=v]\<^sub>\\<^sub>v = \[x::=v]\<^sub>\\<^sub>v" +proof - + have "[[atom x']]lst. s' = [[atom x]]lst. s" and "[[atom x']]lst. \' = [[atom x]]lst. \" and + " [[atom x']]lst. c' = [[atom x]]lst. c" using lst_snd lst_fst assms by metis+ + thus ?thesis using subst_v_flip_eq_two \.eq_iff + by (metis assms fun_typ.eq_iff fun_typ_eq_body_unique fun_typ_eq_ret_unique) +qed + +section \Weakening\ + +lemma wfX_wfB1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and \::\ and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows wfV_wfB: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f b" and + "\; \; \ \\<^sub>w\<^sub>f c \ True" and + "\; \ \\<^sub>w\<^sub>f \ \ True" and + wfT_wfB: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f b_of \ " and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \ True" and + "\; \ \\<^sub>w\<^sub>f b \ True" and + wfCE_wfB: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(induct rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.inducts) + case (wfV_varI \ \ \ b c x) + hence "(x,b,c) \ toSet \" using lookup_iff wfV_wf using lookup_in_g by presburger + hence "b \ fst`snd`toSet \" by force + hence "wfB \ \ b" using wfG_wfB wfV_varI by metis + then show ?case using wfV_elims wfG_wf wf_intros by metis +next + case (wfV_litI \ \ l) + moreover have "wfTh \" using wfV_wf wfG_wf wfV_litI by metis + ultimately show ?case using wfV_wf wfG_wf wf_intros base_for_lit.simps l.exhaust by metis +next + case (wfV_pairI \ \ v1 b1 v2 b2) + then show ?case using wfG_wf wf_intros by metis +next + case (wfV_consI s dclist \ dc x b c B \ v) + then show ?case + using wfV_wf wfG_wf wfB_consI by metis +next + case (wfV_conspI s bv dclist \ dc x b' c \ b \ v) + then show ?case + using wfV_wf wfG_wf using wfB_appI by metis +next + case (wfCE_valI \ \ \ v b) + then show ?case using wfB_elims by auto +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using wfB_elims by auto +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis +next + case (wfCE_eqI \ \ \ v1 b v2) + then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case using wfB_elims by metis +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case using wfB_elims by metis +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using wfB_elims by auto +next + case (wfCE_lenI \ \ \ v1) + then show ?case using wfV_wf wfG_wf wf_intros wfX_wfY by metis +qed(auto | metis wfV_wf wfG_wf wf_intros )+ + +lemma wfX_wfB2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and b::b and \::\ and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows + wfE_wfB: "\; \; \; \; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f b" and + wfS_wfB: "\; \; \; \; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f b" and + wfCS_wfB: "\; \; \; \; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f b" and + wfCSS_wfB: "\; \; \; \; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f b" and + "\ \\<^sub>w\<^sub>f \ \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ True" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ \ |\| \' \ \ ; \ ; \' \\<^sub>w\<^sub>f ft" +proof(induct rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.inducts) + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wfB_elims wfX_wfB1 by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wfB_elims wfX_wfB1 by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wfB_boolI wfX_wfY by metis +next + case (wfE_fstI \ \ \ \ v1 b1 b2) + then show ?case using wfB_elims wfX_wfB1 by metis +next + case (wfE_sndI \ \ \ \ v1 b1 b2) + then show ?case using wfB_elims wfX_wfB1 by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wfB_elims wfX_wfB1 by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wfB_elims wfX_wfB1 + using wfB_pairI by auto +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wfB_elims wfX_wfB1 + using wfB_intI wfX_wfY1(1) by auto +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + hence "\; \;(x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \" using wfPhi_f_simple_wfT wfT_b_weakening by fast + then show ?case using b_of.simps using wfT_b_weakening + by (metis b_of.cases bot.extremum wfT_elims(2)) +next + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + hence "\ ; {| bv |} ;(x,b,c) #\<^sub>\ GNil \\<^sub>w\<^sub>f \" using wfPhi_f_poly_wfT wfX_wfY by blast + then show ?case using wfE_appPI b_of.simps using wfT_b_weakening wfT_elims wfT_subst_wfB subst_b_b_def by metis +next + case (wfE_mvarI \ \ \ \ \ u \) + hence "\; \; \ \\<^sub>w\<^sub>f \" using wfD_wfT by fast + then show ?case using wfT_elims b_of.simps by metis +next + case (wfFTNone \ ft) + then show ?case by auto +next + case (wfFTSome \ bv ft) + then show ?case by auto +next + case (wfS_valI \ \ \ \ v b \) + then show ?case using wfX_wfB1 by auto +next + case (wfS_letI \ \ \ \ \ e b' x s b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wfX_wfB1 by auto +next + case (wfS_varI \ \ \ \ v u \ \ b s) + then show ?case using wfX_wfB1 by auto +next + case (wfS_assignI u \ \ \ \ \ \ v) + then show ?case using wfX_wfB1 + using wfB_unitI wfX_wfY2(5) by auto +next + case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + then show ?case using wfX_wfB1 by auto +next + case (wfS_finalI \ \ \ \ \ tid dc t cs b) + then show ?case using wfX_wfB1 by auto +next + case (wfS_cons \ \ \ \ \ tid dc t cs b dclist css) + then show ?case using wfX_wfB1 by auto +next + case (wfD_emptyI \ \ \) + then show ?case using wfX_wfB1 by auto +next + case (wfD_cons \ \ \ \ \ u) + then show ?case using wfX_wfB1 by auto +next + case (wfPhi_emptyI \) + then show ?case using wfX_wfB1 by auto +next + case (wfPhi_consI f \ \ ft) + then show ?case using wfX_wfB1 by auto +next + case (wfFTI \ B b \ x c s \) + then show ?case using wfX_wfB1 + by (meson Wellformed.wfFTI wb_b_weakening2(8)) +qed(metis wfV_wf wfG_wf wf_intros wfX_wfB1) + +lemmas wfX_wfB = wfX_wfB1 wfX_wfB2 + +lemma wf_weakening1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + + shows wfV_weakening: "\; \; \ \\<^sub>w\<^sub>f v : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f v : b" and + wfC_weakening: "\; \; \ \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f c" and + "\; \ \\<^sub>w\<^sub>f \ \ True" and + wfT_weakening: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f P \ True " and + wfB_weakening: "wfB \ \ b \ \ |\| \' \ wfB \ \ b" and + wfCE_weakening: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f ce : b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b and c and \ and \ and ts and P and b and b and td + avoiding: \' + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_varI \ \ \ b c x) + hence "Some (b, c) = lookup \' x" using lookup_weakening by metis + then show ?case using Wellformed.wfV_varI wfV_varI by metis +next + case (wfTI z \ \ \ b c) (* This proof pattern is used elsewhere when proving weakening for typing predicates *) + show ?case proof + show \atom z \ (\, \, \')\ using wfTI by auto + show \ \; \ \\<^sub>w\<^sub>f b \ using wfTI by auto + have *:"toSet ((z, b, TRUE) #\<^sub>\ \) \ toSet ((z, b, TRUE) #\<^sub>\ \')" using toSet.simps wfTI by auto + thus \ \; \; (z, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c \ using wfTI(8)[OF _ *] wfTI wfX_wfY + by (simp add: wfG_cons_TRUE) + qed +next + case (wfV_conspI s bv dclist \ dc x b' c \ b \ v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by auto + show \(dc, \ x : b' | c \) \ set dclist\ using wfV_conspI by auto + show \ \; \ \\<^sub>w\<^sub>f b \ using wfV_conspI by auto + show \atom bv \ (\, \, \', b, v)\ using wfV_conspI by simp + show \ \; \; \' \\<^sub>w\<^sub>f v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by auto + qed +qed(metis wf_intros)+ + +lemma wf_weakening2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows + wfE_weakening: "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ \\<^sub>w\<^sub>f e : b" and + wfS_weakening: "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ \\<^sub>w\<^sub>f s : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \; \' ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ True" and + wfD_weakning: "\; \; \ \\<^sub>w\<^sub>f \ \ \; \ \\<^sub>w\<^sub>f \' \ toSet \ \ toSet \' \ \; \; \' \\<^sub>w\<^sub>f \" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b and b and b and b and \ and \ and ftq and ft + avoiding: \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto + show \atom bv \ (\, \, \, \', \, b', v, (b_of \)[bv::=b']\<^sub>b)\ using wfE_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f\ using wfE_appPI by auto + show \ \; \; \' \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI wf_weakening1 by auto + qed +next + case (wfS_letI \ \ \ \ \ e b' x s b) + show ?case proof(rule) + show \ \ ; \ ; \ ; \' ; \ \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto + have "toSet ((x, b', TRUE) #\<^sub>\ \) \ toSet ((x, b', TRUE) #\<^sub>\ \')" using wfS_letI by auto + thus \ \ ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_letI by (meson wfG_cons wfG_cons_TRUE wfS_wf) + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfS_letI by auto + show \atom x \ (\, \, \, \', \, e, b)\ using wfS_letI by auto + qed +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + show ?case proof + show \ \ ; \ ; \ ; \' ; \ \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfS_let2I wf_weakening1 by auto + have "toSet ((x, b_of \, TRUE) #\<^sub>\ \) \ toSet ((x, b_of \, TRUE) #\<^sub>\ \')" using wfS_let2I by auto + thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s2 : b \ using wfS_let2I by (meson wfG_cons wfG_cons_TRUE wfS_wf) + show \atom x \ (\, \, \, \', \, s1, b, \)\ using wfS_let2I by auto + qed +next + case (wfS_varI \ \ \ \ v u \ \ b s) + show ?case proof + show "\; \; \' \\<^sub>w\<^sub>f \ " using wfS_varI wf_weakening1 by auto + show "\; \; \' \\<^sub>w\<^sub>f v : b_of \ " using wfS_varI wf_weakening1 by auto + show "atom u \ (\, \, \, \', \, \, v, b)" using wfS_varI by auto + show "\ ; \ ; \ ; \' ; (u, \) #\<^sub>\ \ \\<^sub>w\<^sub>f s : b " using wfS_varI by auto + qed +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + show ?case proof + have "toSet ((x, b_of \, TRUE) #\<^sub>\ \) \ toSet ((x, b_of \, TRUE) #\<^sub>\ \')" using wfS_branchI by auto + thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_branchI by (meson wfG_cons wfG_cons_TRUE wfS_wf) + show \atom x \ (\, \, \, \', \, \', \)\ using wfS_branchI by auto + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfS_branchI by auto + qed +next + case (wfS_finalI \ \ \ \ \ tid dclist' cs b dclist) + then show ?case using wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dclist' cs b css dclist) + then show ?case using wf_intros by metis +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case proof(rule) + show \ \; \; \' \\<^sub>w\<^sub>f c \ using wfS_assertI wf_weakening1 by auto + have "\; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \'" proof(rule wfG_consI) + show \ \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto + show \atom x \ \'\ using wfS_assertI by auto + show \ \; \ \\<^sub>w\<^sub>f B_bool \ using wfS_assertI wfB_boolI wfX_wfY by metis + have "\; \ \\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\ \'" proof + show "(TRUE) \ {TRUE, FALSE}" by auto + show \ \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto + show \atom x \ \'\ using wfS_assertI by auto + show \ \; \ \\<^sub>w\<^sub>f B_bool \ using wfS_assertI wfB_boolI wfX_wfY by metis + qed + thus \ \; \; (x, B_bool, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c \ + using wf_weakening1(2)[OF \ \; \; \' \\<^sub>w\<^sub>f c \ \ \; \ \\<^sub>w\<^sub>f (x, B_bool, TRUE) #\<^sub>\ \' \] by force + qed + thus \ \; \; \; (x, B_bool, c) #\<^sub>\ \' ; \ \\<^sub>w\<^sub>f s : b \ using wfS_assertI by fastforce + show \ \; \; \' \\<^sub>w\<^sub>f \ \ using wfS_assertI by auto + show \atom x \ (\, \, \, \', \, c, b, s)\ using wfS_assertI by auto + qed +qed(metis wf_intros wf_weakening1)+ + +lemmas wf_weakening = wf_weakening1 wf_weakening2 + +lemma wfV_weakening_cons: + fixes \::\ and \'::\ and v::v and c::c + assumes "\; \; \ \\<^sub>w\<^sub>f v : b" and "atom y \ \" and "\; \; ((y,b',TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f c" + shows "\; \; (y,b',c) #\<^sub>\\ \\<^sub>w\<^sub>f v : b" +proof - + have "wfG \ \ ((y,b',c) #\<^sub>\\)" using wfG_intros2 assms by auto + moreover have "toSet \ \ toSet ((y,b',c) #\<^sub>\\)" using toSet.simps by auto + ultimately show ?thesis using wf_weakening using assms(1) by blast +qed + +lemma wfG_cons_weakening: + fixes \'::\ + assumes "\; \ \\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\ \)" and "\; \ \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "atom x \ \'" + shows "\; \ \\<^sub>w\<^sub>f ((x, b, c) #\<^sub>\ \')" +proof(cases "c \ {TRUE,FALSE}") + case True + then show ?thesis using wfG_wfB wfG_cons2I assms by auto +next + case False + hence *:"\; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ \; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c" + using wfG_elims(2)[OF assms(1)] by auto + have a1:"\; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \'" using wfG_wfB wfG_cons2I assms by simp + moreover have a2:"toSet ((x, b, TRUE) #\<^sub>\ \ ) \ toSet ((x, b, TRUE) #\<^sub>\ \')" using toSet.simps assms by blast + moreover have " \; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \'" proof + show "(TRUE) \ {TRUE, FALSE}" by auto + show "\; \ \\<^sub>w\<^sub>f \'" using assms by auto + show "atom x \ \'" using assms by auto + show "\; \ \\<^sub>w\<^sub>f b" using assms wfG_elims by metis + qed + hence " \; \; (x, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c" using wf_weakening a1 a2 * by auto + then show ?thesis using wfG_cons1I[of c \ \ \' x b, OF False ] wfG_wfB assms by simp +qed + +lemma wfT_weakening_aux: + fixes \::\ and \'::\ and c::c + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "\; \ \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "atom z \ \'" + shows "\; \; \' \\<^sub>w\<^sub>f \ z : b | c \" +proof + show \atom z \ (\, \, \')\ + using wf_supp wfX_wfY assms fresh_prodN fresh_def x_not_in_b_set wfG_fresh_x by metis + show \ \; \ \\<^sub>w\<^sub>f b \ using assms wfT_elims by metis + show \ \; \; (z, b, TRUE) #\<^sub>\ \' \\<^sub>w\<^sub>f c \ proof - + have *:"\; \; (z,b,TRUE) #\<^sub>\\ \\<^sub>w\<^sub>f c" using wfT_wfC fresh_weakening assms by auto + moreover have a1:"\; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \'" using wfG_cons2I assms \\; \ \\<^sub>w\<^sub>f b\ by simp + moreover have a2:"toSet ((z, b, TRUE) #\<^sub>\ \ ) \ toSet ((z, b, TRUE) #\<^sub>\ \')" using toSet.simps assms by blast + moreover have " \; \ \\<^sub>w\<^sub>f (z, b, TRUE) #\<^sub>\ \' " proof + show "(TRUE) \ {TRUE, FALSE}" by auto + show "\; \ \\<^sub>w\<^sub>f \'" using assms by auto + show "atom z \ \'" using assms by auto + show "\; \ \\<^sub>w\<^sub>f b" using assms wfT_elims by metis + qed + thus ?thesis using wf_weakening a1 a2 * by auto + qed +qed + +lemma wfT_weakening_all: + fixes \::\ and \'::\ and \::\ + assumes "\; \; \ \\<^sub>w\<^sub>f \" and "\; \' \\<^sub>w\<^sub>f \'" and "toSet \ \ toSet \'" and "\ |\| \'" + shows "\; \' ; \' \\<^sub>w\<^sub>f \" + using wb_b_weakening assms wfT_weakening by metis + +lemma wfT_weakening_nil: + fixes \::\ and \'::\ and \::\ + assumes "\ ; {||} ; GNil \\<^sub>w\<^sub>f \" and "\; \' \\<^sub>w\<^sub>f \'" + shows "\; \' ; \' \\<^sub>w\<^sub>f \" + using wfT_weakening_all + using assms(1) assms(2) toSet.simps(1) by blast + +lemma wfTh_wfT2: + fixes x::x and v::v and \::\ and G::\ + assumes "wfTh \" and "AF_typedef s dclist \ set \" and + "(dc, \) \ set dclist" and "\ ; B \\<^sub>w\<^sub>f G" + shows "supp \ = {}" and "\[x::=v]\<^sub>\\<^sub>v = \" and "wfT \ B G \" +proof - + show "supp \ = {}" proof(rule ccontr) + assume a1: "supp \ \ {}" + have "supp \ \ {}" proof - + obtain dclist where dc: "AF_typedef s dclist \ set \ \ (dc, \) \ set dclist" + using assms by auto + hence "supp (dc,\) \ {}" + using a1 by (simp add: supp_Pair) + hence "supp dclist \ {}" using dc supp_list_member by auto + hence "supp (AF_typedef s dclist) \ {}" using type_def.supp by auto + thus ?thesis using supp_list_member dc by auto + qed + thus False using assms wfTh_supp by simp + qed + thus "\[x::=v]\<^sub>\\<^sub>v = \" by (simp add: fresh_def) + have "wfT \ {||} GNil \" using assms wfTh_wfT by auto + thus "wfT \ B G \" using assms wfT_weakening_nil by simp +qed + +lemma wf_d_weakening: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and s::s and \::\ and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows + "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' \\<^sub>w\<^sub>f e : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' \\<^sub>w\<^sub>f s : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b" and + "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \; \; \ \\<^sub>w\<^sub>f \' \ setD \ \ setD \' \ \; \; \; \ ; \' ; tid ; dclist \\<^sub>w\<^sub>f css : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ True" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b and b and b and b and \ and \ and ftq and ft + avoiding: \' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_intros by metis +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_intros by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_intros by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_intros by metis +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_intros by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then show ?case using wf_intros by metis +next + case (wfE_appPI \ \ \ \ \ b' bv v \ f x b c s) + show ?case proof(rule, (rule wfE_appPI)+) + show \atom bv \ (\, \, \, \, \', b', v, (b_of \)[bv::=b']\<^sub>b)\ using wfE_appPI by auto + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x b c \ s))) = lookup_fun \ f\ using wfE_appPI by auto + show \ \; \; \ \\<^sub>w\<^sub>f v : b[bv::=b']\<^sub>b \ using wfE_appPI by auto + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_mvarI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfE_mvarI by auto + show \(u, \) \ setD \'\ using wfE_mvarI by auto + qed +next + case (wfS_valI \ \ \ \ v b \) + then show ?case using wf_intros by metis +next + case (wfS_letI \ \ \ \ \ e b' x s b) + show ?case proof(rule) + show \ \; \; \; \; \' \\<^sub>w\<^sub>f e : b' \ using wfS_letI by auto + have "\; \ \\<^sub>w\<^sub>f (x, b', TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_letI by metis + hence "\; \; (x, b', TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_letI by force + thus \ \ ; \ ; \ ; (x, b', TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_letI by metis + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_letI by auto + show \atom x \ (\, \, \, \, \', e, b)\ using wfS_letI by auto + qed +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case proof + have "\; \; (x, B_bool, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" proof(rule wf_weakening2(6)) + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto + next + show \ \; \ \\<^sub>w\<^sub>f (x, B_bool, c) #\<^sub>\ \ \ using wfS_assertI wfX_wfY by metis + next + show \toSet \ \ toSet ((x, B_bool, c) #\<^sub>\ \)\ using wfS_assertI by auto + qed + thus \ \; \; \; (x, B_bool, c) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_assertI wfX_wfY by metis + next + show \ \; \; \ \\<^sub>w\<^sub>f c \ using wfS_assertI by auto + next + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_assertI by auto + next + show \atom x \ (\, \, \, \, \', c, b, s)\ using wfS_assertI by auto + qed +next + case (wfS_let2I \ \ \ \ \ s1 \ x s2 b) + show ?case proof + show \ \; \; \; \; \' \\<^sub>w\<^sub>f s1 : b_of \ \ using wfS_let2I by auto + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_let2I by auto + have "\; \ \\<^sub>w\<^sub>f (x, b_of \, TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_let2I by metis + hence "\; \; (x, b_of \, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_let2I by force + thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s2 : b \ using wfS_let2I by metis + show \atom x \ (\, \, \, \, \', s1, b,\)\ using wfS_let2I by auto + qed +next + case (wfS_ifI \ \ \ v \ \ s1 b s2) + then show ?case using wf_intros by metis +next + case (wfS_varI \ \ \ \ v u \ \ b s) + show ?case proof + show \ \; \; \ \\<^sub>w\<^sub>f \ \ using wfS_varI by auto + show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_varI by auto + show \atom u \ (\, \, \, \, \', \, v, b)\ using wfS_varI setD.simps by auto + have "\; \; \ \\<^sub>w\<^sub>f (u, \) #\<^sub>\ \'" using wfS_varI wfD_cons setD.simps u_fresh_d by metis + thus \ \ ; \ ; \ ; \ ; (u, \) #\<^sub>\ \' \\<^sub>w\<^sub>f s : b \ using wfS_varI setD.simps by blast + qed +next + case (wfS_assignI u \ \ \ \ \ \ v) + show ?case proof + show \(u, \) \ setD \'\ using wfS_assignI setD.simps by auto + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_assignI by auto + show \ \ \\<^sub>w\<^sub>f \ \ using wfS_assignI by auto + show \ \; \; \ \\<^sub>w\<^sub>f v : b_of \ \ using wfS_assignI by auto + qed +next + case (wfS_whileI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros by metis +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using wf_intros by metis +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + then show ?case using wf_intros by metis +next + case (wfS_branchI \ \ \ x \ \ \ s b tid dc) + show ?case proof + have "\; \ \\<^sub>w\<^sub>f (x, b_of \, TRUE) #\<^sub>\ \" using wfG_cons2I wfX_wfY wfS_branchI by metis + hence "\; \; (x, b_of \, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f \'" using wf_weakening2(6) wfS_branchI by force + thus \ \ ; \ ; \ ; (x, b_of \, TRUE) #\<^sub>\ \ ; \' \\<^sub>w\<^sub>f s : b \ using wfS_branchI by simp + show \ atom x \ (\, \, \, \, \', \, \)\ using wfS_branchI by auto + show \ \; \; \ \\<^sub>w\<^sub>f \' \ using wfS_branchI by auto + qed +next + case (wfS_finalI \ \ \ \ \ tid dclist' cs b dclist) + then show ?case using wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dclist' cs b css dclist) + then show ?case using wf_intros by metis +qed(auto+) + +section \Useful well-formedness instances\ + +text \Well-formedness for particular constructs that we will need later\ + +lemma wfC_e_eq: + fixes ce::ce and \::\ + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" and "atom x \ \ " + shows "\ ; \ ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f (CE_val (V_var x) == ce )" +proof - + have "\; \ \\<^sub>w\<^sub>f b" using assms wfX_wfB by auto + hence wbg: "\; \ \\<^sub>w\<^sub>f \" using wfX_wfY assms by auto + show ?thesis proof + show *:"\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) : b" + proof(rule) + show "\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f V_var x : b " proof + show "\ ; \ \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \ " using wfG_cons2I wfX_wfY assms \\; \ \\<^sub>w\<^sub>f b\ by auto + show "Some (b, TRUE) = lookup ((x, b, TRUE) #\<^sub>\ \) x" using lookup.simps by auto + qed + qed + show "\ ; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f ce : b" + using assms wf_weakening1(8)[OF assms(1), of "(x, b, TRUE) #\<^sub>\ \ "] * toSet.simps wfX_wfY + by (metis Un_subset_iff equalityE) + qed +qed + +lemma wfC_e_eq2: + fixes e1::ce and e2::ce + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f e1 : b" and "\ ; \ ; \ \\<^sub>w\<^sub>f e2 : b" and " \\<^sub>w\<^sub>f \" and "atom x \ \" + shows "\; \; (x, b, (CE_val (V_var x)) == e1) #\<^sub>\ \ \\<^sub>w\<^sub>f (CE_val (V_var x)) == e2 " +proof(rule wfC_eqI) + have *: "\; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \" proof(rule wfG_cons1I ) + show "(CE_val (V_var x) == e1 ) \ {TRUE, FALSE}" by auto + show "\; \ \\<^sub>w\<^sub>f \" using assms wfX_wfY by metis + show *:"atom x \ \" using assms by auto + show "\; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) == e1" using wfC_e_eq assms * by auto + show "\; \ \\<^sub>w\<^sub>f b" using assms wfX_wfB by auto + qed + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var x) : b" using assms * wfCE_valI wfV_varI by auto + show "\; \; (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \ \\<^sub>w\<^sub>f e2 : b" proof(rule wf_weakening1(8)) + show "\; \; \ \\<^sub>w\<^sub>f e2 : b " using assms by auto + show "\; \ \\<^sub>w\<^sub>f (x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \" using * by auto + show "toSet \ \ toSet ((x, b, CE_val (V_var x) == e1 ) #\<^sub>\ \)" by auto + qed +qed + +lemma wfT_wfT_if_rev: + assumes "wfV P \ \ v (base_for_lit l)" and "wfT P \ \ t" and \atom z1 \ \\ + shows "wfT P \ \ (\ z1 : b_of t | CE_val v == CE_val (V_lit l) IMP (c_of t z1) \)" +proof + show \ P; \ \\<^sub>w\<^sub>f b_of t \ using wfX_wfY assms by meson + have wfg: " P; \ \\<^sub>w\<^sub>f (z1, b_of t, TRUE) #\<^sub>\ \" using assms wfV_wf wfG_cons2I wfX_wfY + by (meson wfG_cons_TRUE) + show \ P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of t z1 \ proof + show *: \ P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e == [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e \ + proof(rule wfC_eqI[where b="base_for_lit l"]) + show "P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ v ]\<^sup>c\<^sup>e : base_for_lit l" + using assms wf_intros wf_weakening wfg by (meson wfV_weakening_cons) + show "P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f [ [ l ]\<^sup>v ]\<^sup>c\<^sup>e : base_for_lit l" using wfg assms wf_intros wf_weakening wfV_weakening_cons by meson + qed + have " t = \ z1 : b_of t | c_of t z1 \" using c_of_eq + using assms(2) assms(3) b_of_c_of_eq wfT_x_fresh by auto + thus \ P; \ ; (z1, b_of t, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c_of t z1 \ using wfT_wfC assms wfG_elims * by simp + qed + show \atom z1 \ (P, \, \)\ using assms wfG_fresh_x wfX_wfY by metis +qed + +lemma wfT_eq_imp: + fixes zz::x and ll::l and \'::\ + assumes "base_for_lit ll = B_bool" and "\ ; {||} ; GNil \\<^sub>w\<^sub>f \'" and + "\ ; {||} \\<^sub>w\<^sub>f (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil" and "atom zz \ x" + shows "\ ; {||} ; (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ + GNil \\<^sub>w\<^sub>f \ zz : b_of \' | [ [ x ]\<^sup>v ]\<^sup>c\<^sup>e == [ [ ll ]\<^sup>v ]\<^sup>c\<^sup>e IMP c_of \' zz \" +proof(rule wfT_wfT_if_rev) + show \ \ ; {||} ; (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil \\<^sub>w\<^sub>f [ x ]\<^sup>v : base_for_lit ll \ + using wfV_varI lookup.simps base_for_lit.simps assms by simp + show \ \ ; {||} ; (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil \\<^sub>w\<^sub>f \' \ + using wf_weakening assms toSet.simps by auto + show \atom zz \ (x, b_of \ z' : B_bool | TRUE \, c_of \ z' : B_bool | TRUE \ x) #\<^sub>\ GNil\ + unfolding fresh_GCons fresh_prod3 b_of.simps c_of_true + using x_fresh_b fresh_GNil c_of_true c.fresh assms by metis +qed + +lemma wfC_v_eq: + fixes ce::ce and \::\ and v::v + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f v : b" and "atom x \ \ " + shows "\ ; \ ; ((x, b, TRUE) #\<^sub>\ \) \\<^sub>w\<^sub>f (CE_val (V_var x) == CE_val v )" + using wfC_e_eq wfCE_valI assms wfX_wfY by auto + +lemma wfT_e_eq: + fixes ce::ce + assumes "\ ; \ ; \ \\<^sub>w\<^sub>f ce : b" and "atom z \ \" + shows "\; \; \ \\<^sub>w\<^sub>f \ z : b | CE_val (V_var z) == ce \" +proof + show "\; \ \\<^sub>w\<^sub>f b" using wfX_wfB assms by auto + show " atom z \ (\, \, \)" using assms wfG_fresh_x wfX_wfY by metis + show "\ ; \ ; (z, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f CE_val (V_var z) == ce " + using wfTI wfC_e_eq assms wfTI by auto +qed + +lemma wfT_v_eq: + assumes " wfB \ \ b" and "wfV \ \ \ v b" and "atom z \ \" + shows "wfT \ \ \ \ z : b | C_eq (CE_val (V_var z)) (CE_val v)\" + using wfT_e_eq wfE_valI assms wfX_wfY + by (simp add: wfCE_valI) + +lemma wfC_wfG: + fixes \::\ and c::c and b::b + assumes "\ ; B ; \ \\<^sub>w\<^sub>f c" and "\ ; B \\<^sub>w\<^sub>f b" and "atom x \ \" + shows "\ ; B \\<^sub>w\<^sub>f (x,b,c)#\<^sub>\ \" +proof - + have " \ ; B \\<^sub>w\<^sub>f (x, b, TRUE) #\<^sub>\ \" using wfG_cons2I assms wfX_wfY by fast + hence " \ ; B ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c " using wfC_weakening assms by force + thus ?thesis using wfG_consI assms wfX_wfY by metis +qed + +section \Replacing the constraint on a variable in a context\ + +lemma wfG_cons_fresh2: + fixes \'::\ + assumes "wfG P \ (( (x',b',c') #\<^sub>\ \' @ (x, b, c) #\<^sub>\ \))" + shows "x'\x" +proof - + have "atom x' \ (\' @ (x, b, c) #\<^sub>\ \)" + using assms wfG_elims(2) by blast + thus ?thesis + using fresh_gamma_append[of "atom x'" \' "(x, b, c) #\<^sub>\ \"] fresh_GCons fresh_prod3[of "atom x'" x b c] by auto +qed + +lemma replace_in_g_inside: + fixes \::\ + assumes "wfG P \ (\'@((x,b0,c0') #\<^sub>\\))" + shows "replace_in_g (\'@((x,b0,c0') #\<^sub>\\)) x c0 = (\'@((x,b0,c0) #\<^sub>\\))" +using assms proof(induct \' rule: \_induct) + case GNil + then show ?case using replace_in_g.simps by auto +next + case (GCons x' b' c' \'') + hence "P; \ \\<^sub>w\<^sub>f ((x', b', c') #\<^sub>\ (\''@ (x, b0, c0') #\<^sub>\ \ ))" by simp + hence "x \ x'" using wfG_cons_fresh2 by metis + then show ?case using replace_in_g.simps GCons by (simp add: wfG_cons) +qed + +lemma wfG_supp_rig_eq: + fixes \::\ + assumes "wfG P \ (\'' @ (x, b0, c0) #\<^sub>\ \)" and "wfG P \ (\'' @ (x, b0, c0') #\<^sub>\ \)" + shows "supp (\'' @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp (\'' @ (x, b0, c0) #\<^sub>\ \) \ supp \" +using assms proof(induct \'') + case GNil + have "supp (GNil @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp ((x, b0, c0') #\<^sub>\ \) \ supp \" using supp_Cons supp_GNil by auto + also have "... = supp x \ supp b0 \ supp c0' \ supp \ \ supp \ " using supp_GCons by auto + also have "... = supp x \ supp b0 \ supp c0 \ supp \ \ supp \ " using GNil wfG_wfC[THEN wfC_supp_cons(2) ] by fastforce + also have "... = (supp ((x, b0, c0) #\<^sub>\ \)) \ supp \ " using supp_GCons by auto + finally have "supp (GNil @ (x, b0, c0') #\<^sub>\ \) \ supp \ = supp (GNil @ (x, b0, c0) #\<^sub>\ \) \ supp \" using supp_Cons supp_GNil by auto + then show ?case using supp_GCons wfG_cons2 by auto +next + case (GCons xbc \1) + moreover have " (xbc #\<^sub>\ \1) @ (x, b0, c0) #\<^sub>\ \ = (xbc #\<^sub>\ (\1 @ (x, b0, c0) #\<^sub>\ \))" by simp + moreover have " (xbc #\<^sub>\ \1) @ (x, b0, c0') #\<^sub>\ \ = (xbc #\<^sub>\ (\1 @ (x, b0, c0') #\<^sub>\ \))" by simp + ultimately have "(P; \ \\<^sub>w\<^sub>f \1 @ ((x, b0, c0) #\<^sub>\ \)) \ P; \ \\<^sub>w\<^sub>f \1 @ ((x, b0, c0') #\<^sub>\ \)" using wfG_cons2 by metis + thus ?case using GCons supp_GCons by auto +qed + +lemma fresh_replace_inside[ms_fresh]: + fixes y::x and \::\ + assumes "wfG P \ (\'' @ (x, b, c) #\<^sub>\ \)" and "wfG P \ (\'' @ (x, b, c') #\<^sub>\ \)" + shows "atom y \ (\'' @ (x, b, c) #\<^sub>\ \) = atom y \ (\'' @ (x, b, c') #\<^sub>\ \)" + unfolding fresh_def using wfG_supp_rig_eq assms x_not_in_b_set by fast + +lemma wf_replace_inside1: + fixes \::\ and \::\ and \::\ and \'::\ and v::v and e::e and c::c and c''::c and c'::c and \::\ and ts::"(string*\) list" and \::\ and b'::b and b::b and s::s + and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list + +shows wfV_replace_inside: "\; \; G \\<^sub>w\<^sub>f v : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f v : b'" and + wfC_replace_inside: "\; \; G \\<^sub>w\<^sub>f c'' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f c''" and + wfG_replace_inside: "\; \ \\<^sub>w\<^sub>f G \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \; \ \\<^sub>w\<^sub>f (\' @ (x, b, c) #\<^sub>\ \) " and + wfT_replace_inside: "\; \; G \\<^sub>w\<^sub>f \ \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f P \ True" and + "\; \ \\<^sub>w\<^sub>f b \ True" and + wfCE_replace_inside: "\ ; \ ; G \\<^sub>w\<^sub>f ce : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f ce : b'" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b' and c'' and G and \ and ts and P and b and b' and td + avoiding: \' c' +rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_varI \ \ \2 b2 c2 x2) + then show ?case using wf_intros by (metis lookup_in_rig_eq lookup_in_rig_neq replace_in_g_inside) +next + case (wfV_conspI s bv dclist \ dc x1 b' c1 \ b1 \1 v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by auto + show \(dc, \ x1 : b' | c1 \) \ set dclist\ using wfV_conspI by auto + show \ \ ; \ \\<^sub>w\<^sub>f b1 \ using wfV_conspI by auto + show *: \ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b \ using wfV_conspI by auto + moreover have "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c') #\<^sub>\ \" using wfV_wf wfV_conspI by simp + ultimately have "atom bv \ \' @ (x, b, c) #\<^sub>\ \" unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfV_conspI + by (metis Un_iff fresh_def) + thus \atom bv \ (\, \, \' @ (x, b, c) #\<^sub>\ \, b1, v)\ + unfolding fresh_prodN using fresh_prodN wfV_conspI by metis + qed +next + case (wfTI z \ \ G b1 c1) + show ?case proof + show \ \; \ \\<^sub>w\<^sub>f b1 \ using wfTI by auto + + have "\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" using wfG_consI wfTI wfG_cons wfX_wfY by metis + moreover hence *:"wfG \ \ (\' @ (x, b, c) #\<^sub>\ \)" using wfX_wfY + by (metis append_g.simps(2) wfG_cons2 wfTI.hyps wfTI.prems(1) wfTI.prems(2)) + hence \atom z \ \' @ (x, b, c) #\<^sub>\ \\ + using fresh_replace_inside[of \ \ \' x b c \ c' z,OF *] wfTI wfX_wfY wfG_elims by metis + thus \atom z \ (\, \, \' @ (x, b, c) #\<^sub>\ \)\ using wfG_fresh_x[OF *] by auto + + have "(z, b1, TRUE) #\<^sub>\ G = ((z, b1, TRUE) #\<^sub>\ \') @ (x, b, c') #\<^sub>\ \" + using wfTI append_g.simps by metis + thus \ \; \; (z, b1, TRUE) #\<^sub>\ \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f c1 \ + using wfTI(9)[OF _ wfTI(11)] by fastforce + qed +next + case (wfG_nilI \) + hence "GNil = (x, b, c') #\<^sub>\ \" using append_g.simps \.distinct GNil_append by auto + hence "False" using \.distinct by auto + then show ?case by auto +next + case (wfG_cons1I c1 \ \ G x1 b1) + show ?case proof(cases "\'=GNil") + case True + then show ?thesis using wfG_cons1I wfG_consI by auto + next + case False + then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \'" using wfG_cons1I wfG_cons1I(7) GCons_eq_append_conv by auto + hence **:" G = G' @ (x, b, c') #\<^sub>\ \" using wfG_cons1I by auto + hence " \; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \" using wfG_cons1I by auto + have "\; \ \\<^sub>w\<^sub>f (x1, b1, c1) #\<^sub>\ G' @ (x, b, c) #\<^sub>\ \" proof(rule Wellformed.wfG_cons1I) + show "c1 \ {TRUE, FALSE}" using wfG_cons1I by auto + show "\; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \ " using wfG_cons1I(3)[of G',OF **] wfG_cons1I by auto + show "atom x1 \ G' @ (x, b, c) #\<^sub>\ \" using wfG_cons1I * ** fresh_replace_inside by metis + show "\; \; (x1, b1, TRUE) #\<^sub>\ G' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f c1" using wfG_cons1I(6)[of " (x1, b1, TRUE) #\<^sub>\ G'"] wfG_cons1I ** by auto + show "\; \ \\<^sub>w\<^sub>f b1" using wfG_cons1I by auto + qed + thus ?thesis using * by auto + qed +next + case (wfG_cons2I c1 \ \ G x1 b1) + show ?case proof(cases "\'=GNil") + case True + then show ?thesis using wfG_cons2I wfG_consI by auto + next + case False + then obtain G'::\ where *:"(x1, b1, c1) #\<^sub>\ G' = \'" using wfG_cons2I GCons_eq_append_conv by auto + hence **:" G = G' @ (x, b, c') #\<^sub>\ \" using wfG_cons2I by auto + moreover have " \; \ \\<^sub>w\<^sub>f G' @ (x, b, c) #\<^sub>\ \" using wfG_cons2I * ** by auto + moreover hence "atom x1 \ G' @ (x, b, c) #\<^sub>\ \" using wfG_cons2I * ** fresh_replace_inside by metis + ultimately show ?thesis using Wellformed.wfG_cons2I[OF wfG_cons2I(1), of \ \ "G'@ (x, b, c) #\<^sub>\ \" x1 b1] wfG_cons2I * ** by auto + qed +qed(metis wf_intros )+ + +lemma wf_replace_inside2: + fixes \::\ and \::\ and \::\ and \'::\ and v::v and e::e and c::c and c''::c and c'::c and \::\ and ts::"(string*\) list" and \::\ and b'::b and b::b and s::s + and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and cs::branch_s and css::branch_list +shows + "\ ; \ ; \ ; G ; D \\<^sub>w\<^sub>f e : b' \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \); D \\<^sub>w\<^sub>f e : b'" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and + "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and + "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and + "\ \\<^sub>w\<^sub>f \ \ True" and + "\; \; G \\<^sub>w\<^sub>f \ \ G = (\' @ (x, b, c') #\<^sub>\ \) \ \; \; ((x,b,TRUE) #\<^sub>\\) \\<^sub>w\<^sub>f c \ \ ; \ ; (\' @ (x, b, c) #\<^sub>\ \) \\<^sub>w\<^sub>f \" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b' and b and b and b and \ and \ and ftq and ft + avoiding: \' c' + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) +case (wfE_valI \ \ \ \ \ v b) + then show ?case using wf_replace_inside1 Wellformed.wfE_valI by auto +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case using wf_replace_inside1 Wellformed.wfE_plusI by auto +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case using wf_replace_inside1 Wellformed.wfE_leqI by auto +next + case (wfE_eqI \ \ \ \ \ v1 b v2) + then show ?case using wf_replace_inside1 Wellformed.wfE_eqI by metis +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_replace_inside1 Wellformed.wfE_fstI by metis +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case using wf_replace_inside1 Wellformed.wfE_sndI by metis +next + case (wfE_concatI \ \ \ \ \ v1 v2) + then show ?case using wf_replace_inside1 Wellformed.wfE_concatI by auto +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case using wf_replace_inside1 Wellformed.wfE_splitI by auto +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case using wf_replace_inside1 Wellformed.wfE_lenI by metis +next + case (wfE_appI \ \ \ \ \ f x b c \ s v) + then show ?case using wf_replace_inside1 Wellformed.wfE_appI by metis +next + case (wfE_appPI \ \ \ \'' \ b' bv v \ f x1 b1 c1 s) + show ?case proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \; \ \\<^sub>w\<^sub>f b' \ using wfE_appPI by auto + show *:\ \; \; \' @ (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f v : b1[bv::=b']\<^sub>b \ using wfE_appPI wf_replace_inside1 by auto + + moreover have "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c') #\<^sub>\ \" using wfV_wf wfE_appPI by metis + ultimately have "atom bv \ \' @ (x, b, c) #\<^sub>\ \" + unfolding fresh_def using wfV_wf wfG_supp_rig_eq wfE_appPI Un_iff fresh_def by metis + + thus \atom bv \ (\, \, \, \' @ (x, b, c) #\<^sub>\ \, \, b', v, (b_of \)[bv::=b']\<^sub>b)\ + using wfE_appPI fresh_prodN by metis + show \Some (AF_fundef f (AF_fun_typ_some bv (AF_fun_typ x1 b1 c1 \ s))) = lookup_fun \ f\ using wfE_appPI by auto + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + then show ?case using wf_replace_inside1 Wellformed.wfE_mvarI by metis +next + case (wfD_emptyI \ \ \) + then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis +next + case (wfD_cons \ \ \ \ \ u) + then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI + by (simp add: wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfD_cons) +next + case (wfFTNone \ \ ft) + then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis +next + case (wfFTSome \ \ bv ft) + then show ?case using wf_replace_inside1 Wellformed.wfD_emptyI by metis +qed(auto) + +lemmas wf_replace_inside = wf_replace_inside1 wf_replace_inside2 + +lemma wfC_replace_cons: + assumes "wfG P \ ((x,b,c1) #\<^sub>\\)" and "wfC P \ ((x,b,TRUE) #\<^sub>\\) c2" + shows "wfC P \ ((x,b,c1) #\<^sub>\\) c2" +proof - + have "wfC P \ (GNil@((x,b,c1) #\<^sub>\\)) c2" proof(rule wf_replace_inside1(2)) + show " P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c2 " using wfG_elim2 assms by auto + show \(x, b, TRUE) #\<^sub>\ \ = GNil @ (x, b, TRUE) #\<^sub>\ \\ using append_g.simps by auto + show \P; \ ; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c1 \ using wfG_elim2 assms by auto + qed + thus ?thesis using append_g.simps by auto +qed + +lemma wfC_refl: + assumes "wfG \ \ ((x, b', c') #\<^sub>\\)" + shows "wfC \ \ ((x, b', c') #\<^sub>\\) c'" + using wfG_wfC assms wfC_replace_cons by auto + +lemma wfG_wfC_inside: + assumes " (x, b, c) \ toSet G" and "wfG \ B G" + shows "wfC \ B G c" + using assms proof(induct G rule: \_induct) + case GNil + then show ?case by auto +next + case (GCons x' b' c' \') + then consider (hd) "(x, b, c) = (x',b',c')" | (tail) "(x, b, c) \ toSet \'" using toSet.simps by auto + then show ?case proof(cases) + case hd + then show ?thesis using GCons wf_weakening + by (metis wfC_replace_cons wfG_cons_wfC) + next + case tail + then show ?thesis using GCons wf_weakening + by (metis insert_iff insert_is_Un subsetI toSet.simps(2) wfG_cons2) + qed +qed + +lemma wfT_wf_cons3: + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom y \ (c,\)" + shows "\; \ \\<^sub>w\<^sub>f (y, b, c[z::=V_var y]\<^sub>c\<^sub>v) #\<^sub>\ \" +proof - + have "\ z : b | c \ = \ y : b | (y \ z) \ c \" using type_eq_flip assms by auto + moreover hence " (y \ z) \ c = c[z::=V_var y]\<^sub>c\<^sub>v" using assms subst_v_c_def by auto + ultimately have "\ z : b | c \ = \ y : b | c[z::=V_var y]\<^sub>c\<^sub>v \" by metis + thus ?thesis using assms wfT_wf_cons[of \ \ \ y b] fresh_Pair by metis +qed + +lemma wfT_wfC_cons: + assumes "wfT P \ \ \ z1 : b | c1 \" and "wfT P \ \ \ z2 : b | c2 \" and "atom x \ (c1,c2,\)" + shows "wfC P \ ((x,b,c1[z1::=V_var x]\<^sub>v) #\<^sub>\\) (c2[z2::=V_var x]\<^sub>v)" (is "wfC P \ ?G ?c") +proof - + have eq: "\ z2 : b | c2 \ = \ x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \" using type_eq_subst assms fresh_prod3 by simp + have eq2: "\ z1 : b | c1 \ = \ x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \" using type_eq_subst assms fresh_prod3 by simp + moreover have "wfT P \ \ \ x : b | c1[z1::=V_var x]\<^sub>c\<^sub>v \" using assms eq2 by auto + moreover hence "wfG P \ ((x,b,c1[z1::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\\)" using wfT_wf_cons fresh_prod3 assms by auto + moreover have "wfT P \ \ \ x : b | c2[z2::=V_var x]\<^sub>c\<^sub>v \" using assms eq by auto + moreover hence "wfC P \ ((x,b,TRUE) #\<^sub>\\) (c2[z2::=V_var x]\<^sub>c\<^sub>v)" using wfT_wfC assms fresh_prod3 by simp + ultimately show ?thesis using wfC_replace_cons subst_v_c_def by simp +qed + +lemma wfT_wfC2: + fixes c::c and x::x + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom x \ \" + shows "\; \; (x,b,TRUE)#\<^sub>\\ \\<^sub>w\<^sub>f c[z::=[x]\<^sup>v]\<^sub>v" +proof(cases "x=z") + case True + then show ?thesis using wfT_wfC assms by auto +next + case False + hence "atom x \ c" using wfT_fresh_c assms by metis + hence "\ x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \ = \ z : b | c \" + using \.eq_iff Abs1_eq_iff(3)[of x "c[z::=[ x ]\<^sup>v]\<^sub>v" z c] + by (metis flip_subst_v type_eq_flip) + hence " \; \; \ \\<^sub>w\<^sub>f \ x : b | c[z::=[ x ]\<^sup>v]\<^sub>v \" using assms by metis + thus ?thesis using wfT_wfC assms by auto +qed + +lemma wfT_wfG: + fixes x::x and \::\ and z::x and c::c and b::b + assumes "\; \; \ \\<^sub>w\<^sub>f \ z : b | c \" and "atom x \ \" + shows "\; \ \\<^sub>w\<^sub>f (x,b, c[z::=[ x ]\<^sup>v]\<^sub>v) #\<^sub>\ \" +proof - + have "\; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c[z::=[ x ]\<^sup>v]\<^sub>v" using wfT_wfC2 assms by metis + thus ?thesis using wfG_consI assms wfT_wfB b_of.simps wfX_wfY by metis +qed + +lemma wfG_replace_inside2: + fixes \::\ + assumes "wfG P \ (\' @ (x, b, c') #\<^sub>\ \)" and "wfG P \ ((x,b,c) #\<^sub>\\)" + shows "wfG P \ (\' @ (x, b, c) #\<^sub>\ \)" +proof - + have "wfC P \ ((x,b,TRUE) #\<^sub>\\) c" using wfG_wfC assms by auto + thus ?thesis using wf_replace_inside1(3)[OF assms(1)] by auto +qed + +lemma wfG_replace_inside_full: + fixes \::\ + assumes "wfG P \ (\' @ (x, b, c') #\<^sub>\ \)" and "wfG P \ (\'@((x,b,c) #\<^sub>\\))" + shows "wfG P \ (\' @ (x, b, c) #\<^sub>\ \)" +proof - + have "wfG P \ ((x,b,c) #\<^sub>\\)" using wfG_suffix assms by auto + thus ?thesis using wfG_replace_inside assms by auto +qed + +lemma wfT_replace_inside2: + assumes "wfT \ \ (\' @ (x, b, c') #\<^sub>\ \) t" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" + shows "wfT \ \ (\' @ (x, b, c) #\<^sub>\ \) t" +proof - + have "wfG \ \ (((x,b,c) #\<^sub>\\))" using wfG_suffix assms by auto + hence "wfC \ \ ((x,b,TRUE) #\<^sub>\\) c" using wfG_wfC by auto + thus ?thesis using wf_replace_inside assms by metis +qed + +lemma wfD_unique: + assumes "wfD P \ \ \" and " (u,\') \ setD \" and "(u,\) \ setD \" + shows "\'=\" +using assms proof(induct \ rule: \_induct) + case DNil + then show ?case by auto +next + case (DCons u' t' D) + hence *: "wfD P \ \ ((u',t') #\<^sub>\ D)" using Cons by auto + show ?case proof(cases "u=u'") + case True + then have "u \ fst ` setD D" using wfD_elims * by blast + then show ?thesis using DCons by force + next + case False + then show ?thesis using DCons wfD_elims * by (metis fst_conv setD_ConsD) + qed +qed + +lemma replace_in_g_forget: + fixes x::x + assumes "wfG P B G" + shows "atom x \ atom_dom G \ (G[x\c]) = G" and + "atom x \ G \ (G[x\c]) = G" +proof - + show "atom x \ atom_dom G \ G[x\c] = G" by (induct G rule: \_induct,auto) + thus "atom x \ G \ (G[x\c]) = G" using wfG_x_fresh assms by simp +qed + +lemma replace_in_g_fresh_single: + fixes G::\ and x::x + assumes \\; \ \\<^sub>w\<^sub>f G[x'\c'']\ and "atom x \ G" and \\; \ \\<^sub>w\<^sub>f G \ + shows "atom x \ G[x'\c'']" + using rig_dom_eq wfG_dom_supp assms fresh_def atom_dom.simps dom.simps by metis + +section \Preservation of well-formedness under substitution\ + +lemma wfC_cons_switch: + fixes c::c and c'::c + assumes "\; \; (x, b, c) #\<^sub>\ \ \\<^sub>w\<^sub>f c'" + shows "\; \; (x, b, c') #\<^sub>\ \ \\<^sub>w\<^sub>f c" +proof - + have *:"\; \ \\<^sub>w\<^sub>f (x, b, c) #\<^sub>\ \" using wfC_wf assms by auto + hence "atom x \ \ \ wfG \ \ \ \ \; \ \\<^sub>w\<^sub>f b" using wfG_cons by auto + hence " \; \; (x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f TRUE " using wfC_trueI wfG_cons2I by simp + hence "\; \;(x, b, TRUE) #\<^sub>\ \ \\<^sub>w\<^sub>f c'" + using wf_replace_inside1(2)[of \ \ "(x, b, c) #\<^sub>\ \" c' GNil x b c \ TRUE] assms by auto + hence "wfG \ \ ((x,b,c') #\<^sub>\\)" using wf_replace_inside1(3)[OF *, of GNil x b c \ c'] by auto + moreover have "wfC \ \ ((x,b,TRUE) #\<^sub>\\) c" proof(cases "c \ { TRUE, FALSE }") + case True + have "\; \ \\<^sub>w\<^sub>f \ \ atom x \ \ \ \; \ \\<^sub>w\<^sub>f b" using wfG_elims(2)[OF *] by auto + hence "\; \ \\<^sub>w\<^sub>f (x,b,TRUE) #\<^sub>\ \" using wfG_cons_TRUE by auto + then show ?thesis using wfC_trueI wfC_falseI True by auto + next + case False + then show ?thesis using wfG_elims(2)[OF *] by auto + qed + ultimately show ?thesis using wfC_replace_cons by auto +qed + +lemma subst_g_inside_simple: + fixes \\<^sub>1::\ and \\<^sub>2::\ + assumes "wfG P \ (\\<^sub>1@((x,b,c) #\<^sub>\\\<^sub>2))" + shows "(\\<^sub>1@((x,b,c) #\<^sub>\\\<^sub>2))[x::=v]\<^sub>\\<^sub>v = \\<^sub>1[x::=v]\<^sub>\\<^sub>v@\\<^sub>2" +using assms proof(induct \\<^sub>1 rule: \_induct) + case GNil + then show ?case using subst_gv.simps by simp +next + case (GCons x' b' c' G) + hence *:"P; \ \\<^sub>w\<^sub>f (x', b', c') #\<^sub>\ (G @ (x, b, c) #\<^sub>\ \\<^sub>2)" by auto + hence "x\x'" + using GCons append_Cons wfG_cons_fresh2[OF *] by auto + hence "((GCons (x', b', c') G) @ (GCons (x, b, c) \\<^sub>2))[x::=v]\<^sub>\\<^sub>v = + (GCons (x', b', c') (G @ (GCons (x, b, c) \\<^sub>2)))[x::=v]\<^sub>\\<^sub>v" by auto + also have "... = GCons (x', b', c'[x::=v]\<^sub>c\<^sub>v) ((G @ (GCons (x, b, c) \\<^sub>2))[x::=v]\<^sub>\\<^sub>v)" + using subst_gv.simps \x\x'\ by simp + also have "... = (x', b', c'[x::=v]\<^sub>c\<^sub>v) #\<^sub>\ (G[x::=v]\<^sub>\\<^sub>v @ \\<^sub>2)" using GCons * wfG_elims by metis + also have "... = ((x', b', c') #\<^sub>\ G)[x::=v]\<^sub>\\<^sub>v @ \\<^sub>2" using subst_gv.simps \x\x'\ by simp + finally show ?case by blast +qed + +lemma subst_c_TRUE_FALSE: + fixes c::c + assumes "c \ {TRUE,FALSE}" + shows "c[x::=v']\<^sub>c\<^sub>v \ {TRUE, FALSE}" +using assms by(nominal_induct c rule:c.strong_induct,auto simp add: subst_cv.simps) + +lemma lookup_subst: + assumes "Some (b, c) = lookup \ x" and "x \ x'" + shows "\c'. Some (b,c') = lookup \[x'::=v']\<^sub>\\<^sub>v x" +using assms proof(induct \ rule: \_induct) +case GNil + then show ?case by auto +next + case (GCons x1 b1 c1 \1) + then show ?case proof(cases "x1=x'") + case True + then show ?thesis using subst_gv.simps GCons by auto + next + case False + hence *:"((x1, b1, c1) #\<^sub>\ \1)[x'::=v']\<^sub>\\<^sub>v = ((x1, b1, c1[x'::=v']\<^sub>c\<^sub>v) #\<^sub>\ \1[x'::=v']\<^sub>\\<^sub>v)" using subst_gv.simps by auto + then show ?thesis proof(cases "x1=x") + case True + then show ?thesis using lookup.simps * + using GCons.prems(1) by auto + next + case False + then show ?thesis using lookup.simps * + using GCons.prems(1) by (simp add: GCons.hyps assms(2)) + qed + qed +qed + +lemma lookup_subst2: + assumes "Some (b, c) = lookup (\'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\)) x" and "x \ x'" and + "\; \ \\<^sub>w\<^sub>f (\'@((x',b\<^sub>1,c0[z0::=[x']\<^sup>v]\<^sub>c\<^sub>v)#\<^sub>\\))" + shows "\c'. Some (b,c') = lookup (\'[x'::=v']\<^sub>\\<^sub>v@\) x" + using assms lookup_subst subst_g_inside by metis + +lemma wf_subst1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + shows wfV_subst: "\; \; \ \\<^sub>w\<^sub>f v : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \;\\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b" and + wfC_subst: "\; \; \ \\<^sub>w\<^sub>f c \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v" and + wfG_subst: "\; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and + wfT_subst: "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \True" and + "\; \ \\<^sub>w\<^sub>f b \ True " and + wfCE_subst: "\; \; \ \\<^sub>w\<^sub>f ce : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f ce[x::=v']\<^sub>c\<^sub>e\<^sub>v : b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b and c and \ and \ and ts and \ and b and b and td + avoiding: x v' + arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_varI \ \ \ b1 c1 x1) + + show ?case proof(cases "x1=x") + case True + hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = v' " using subst_vv.simps by auto + moreover have "b' = b1" using wfV_varI True lookup_inside_wf + by (metis option.inject prod.inject) + moreover have " \; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v' : b'" using wfV_varI subst_g_inside_simple wf_weakening + append_g_toSetU sup_ge2 wfV_wf by metis + ultimately show ?thesis by auto + next + case False + hence "(V_var x1)[x::=v']\<^sub>v\<^sub>v = (V_var x1) " using subst_vv.simps by auto + moreover have "\; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" using wfV_varI by simp + moreover obtain c1' where "Some (b1, c1') = lookup \[x::=v']\<^sub>\\<^sub>v x1" using wfV_varI False lookup_subst by metis + ultimately show ?thesis using Wellformed.wfV_varI[of \ \ "\[x::=v']\<^sub>\\<^sub>v" b1 c1' x1] by metis + qed +next + case (wfV_litI \ \ l) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfV_pairI \ \ v1 b1 v2 b2) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfV_consI s dclist \ dc x b c \ v) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfV_conspI s bv dclist \ dc x' b' c \ b \ va) + show ?case unfolding subst_vv.simps proof + show \AF_typedef_poly s bv dclist \ set \\ and \(dc, \ x' : b' | c \) \ set dclist\ using wfV_conspI by auto + show \ \ ;\ \\<^sub>w\<^sub>f b \ using wfV_conspI by auto + have "atom bv \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfV_conspI by metis + moreover have "atom bv \ va[x::=v']\<^sub>v\<^sub>v" using wfV_conspI fresh_subst_if by simp + ultimately show \atom bv \ (\, \, \[x::=v']\<^sub>\\<^sub>v, b, va[x::=v']\<^sub>v\<^sub>v)\ unfolding fresh_prodN using wfV_conspI by auto + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f va[x::=v']\<^sub>v\<^sub>v : b'[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by auto + qed +next + case (wfTI z \ \ \ b c) + have " \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \ z : b | c[x::=v']\<^sub>c\<^sub>v \" proof + have \\; \; ((z, b, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ + proof(rule wfTI(9)) + show \(z, b, TRUE) #\<^sub>\ \ = ((z, b, TRUE) #\<^sub>\ \\<^sub>1) @ (x, b', c') #\<^sub>\ \\<^sub>2\ using wfTI append_g.simps by simp + show \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ using wfTI by auto + qed + thus *:\\; \; (z, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ + using subst_gv.simps subst_cv.simps wfTI fresh_x_neq by auto + + have "atom z \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfTI by metis + moreover have "\; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" using wfTI wfX_wfY wfG_elims subst_gv.simps * by metis + ultimately show \atom z \ (\, \, \[x::=v']\<^sub>\\<^sub>v)\ using wfG_fresh_x by metis + show \ \; \ \\<^sub>w\<^sub>f b \ using wfTI by auto + qed + thus ?case using subst_tv.simps wfTI by auto +next + case (wfC_trueI \ \) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfC_falseI \ \) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfC_eqI \ \ \ e1 b e2) + show ?case proof(subst subst_cv.simps,rule) + show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e1[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI subst_dv.simps by auto + show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e2[x::=v']\<^sub>c\<^sub>e\<^sub>v : b " using wfC_eqI by auto + qed +next + case (wfC_conjI \ \ c1 c2) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfC_disjI \ \ c1 c2) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfC_notI \ \ c1) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfC_impI \ \ c1 c2) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfG_nilI \) + then show ?case using subst_cv.simps wf_intros by auto +next + case (wfG_cons1I c \ \ \ y b) + + show ?case proof(cases "x=y") + case True + hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + moreover have "\; \ \\<^sub>w\<^sub>f \" using wfG_cons1I by auto + ultimately show ?thesis by auto + next + case False + have "\\<^sub>1 \ GNil" using wfG_cons1I False by auto + then obtain G where "\\<^sub>1 = (y, b, c) #\<^sub>\ G" using GCons_eq_append_conv wfG_cons1I by auto + hence *:"\ = G @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons1I by auto + hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps False by auto + moreover have "\; \ \\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" proof(rule Wellformed.wfG_cons1I) + show \c[x::=v']\<^sub>c\<^sub>v \ {TRUE, FALSE}\ using wfG_cons1I subst_c_TRUE_FALSE by auto + show \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfG_cons1I * by auto + have "\ = (G @ ((x, b', c') #\<^sub>\GNil)) @ \\<^sub>2" using * append_g_assoc by auto + hence "atom y \ \\<^sub>2" using fresh_suffix \atom y \ \\ by auto + hence "atom y \ v'" using wfG_cons1I wfV_x_fresh by metis + thus \atom y \ \[x::=v']\<^sub>\\<^sub>v\ using fresh_subst_gv wfG_cons1I by auto + have "((y, b, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = (y, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps subst_cv.simps False by auto + thus \ \; \; (y, b, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ using wfG_cons1I(6)[of "(y,b,TRUE) #\<^sub>\G"] * subst_gv.simps + wfG_cons1I by fastforce + show "\; \ \\<^sub>w\<^sub>f b " using wfG_cons1I by auto + qed + ultimately show ?thesis by auto + qed +next + case (wfG_cons2I c \ \ \ y b) + show ?case proof(cases "x=y") + case True + hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v = \" using subst_gv.simps by auto + moreover have "\; \ \\<^sub>w\<^sub>f \" using wfG_cons2I by auto + ultimately show ?thesis by auto + next + case False + have "\\<^sub>1 \ GNil" using wfG_cons2I False by auto + then obtain G where "\\<^sub>1 = (y, b, c) #\<^sub>\ G" using GCons_eq_append_conv wfG_cons2I by auto + hence *:"\ = G @ (x, b', c') #\<^sub>\ \\<^sub>2" using wfG_cons2I by auto + hence "((y, b, c) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v =(y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" using subst_gv.simps False by auto + moreover have "\; \ \\<^sub>w\<^sub>f (y, b, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\\[x::=v']\<^sub>\\<^sub>v" proof(rule Wellformed.wfG_cons2I) + show \c[x::=v']\<^sub>c\<^sub>v \ {TRUE, FALSE}\ using subst_cv.simps wfG_cons2I by auto + show \ \; \ \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfG_cons2I * by auto + have "\ = (G @ ((x, b', c') #\<^sub>\GNil)) @ \\<^sub>2" using * append_g_assoc by auto + hence "atom y \ \\<^sub>2" using fresh_suffix wfG_cons2I by metis + hence "atom y \ v'" using wfG_cons2I wfV_x_fresh by metis + thus \atom y \ \[x::=v']\<^sub>\\<^sub>v\ using fresh_subst_gv wfG_cons2I by auto + show "\; \ \\<^sub>w\<^sub>f b " using wfG_cons2I by auto + qed + ultimately show ?thesis by auto + qed +next + case (wfCE_valI \ \ \ v b) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfCE_eqI \ \ \ v1 b v2) + then show ?case unfolding subst_cev.simps + using Wellformed.wfCE_eqI by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case using Wellformed.wfCE_fstI subst_cev.simps by metis +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case using subst_cev.simps wf_intros by metis +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using subst_vv.simps wf_intros by auto +next + case (wfCE_lenI \ \ \ v1) + then show ?case using subst_vv.simps wf_intros by auto +qed(metis subst_sv.simps wf_intros)+ + +lemma wf_subst2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def + shows "\; \; \; \ ; \ \\<^sub>w\<^sub>f e : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b" and + "\; \; \; \ ; \ \\<^sub>w\<^sub>f s : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \ ;\ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b" and + "\; \; \; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dc ; t \\<^sub>w\<^sub>f subst_branchv cs x v' : b" and + "\; \; \; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \; \; \; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dclist \\<^sub>w\<^sub>f subst_branchlv css x v' : b" and + "\ \\<^sub>w\<^sub>f (\::\) \ True " and + "\; \; \ \\<^sub>w\<^sub>f \ \ \=\\<^sub>1@((x,b',c') #\<^sub>\\\<^sub>2) \ \; \ ; \\<^sub>2 \\<^sub>w\<^sub>f v' : b' \ \ ; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b and b and b and b and \ and \ and ftq and ft + avoiding: x v' + arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 + rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_valI \ \ v b) + then show ?case using subst_vv.simps wf_intros wf_subst1 + by (metis subst_ev.simps(1)) +next + case (wfE_plusI \ \ v1 v2) + then show ?case using subst_vv.simps wf_intros wf_subst1 by auto +next + case (wfE_leqI \ \ \ \ v1 v2) + then show ?case + using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_leqI + by auto +next + case (wfE_eqI \ \ \ \ v1 b v2) + then show ?case + using subst_vv.simps subst_ev.simps subst_ev.simps wf_subst1 Wellformed.wfE_eqI + proof - + show ?thesis + by (metis (no_types) subst_ev.simps(4) wfE_eqI.hyps(1) wfE_eqI.hyps(4) wfE_eqI.hyps(5) wfE_eqI.hyps(6) wfE_eqI.hyps(7) wfE_eqI.prems(1) wfE_eqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_eqI wfV_subst) (* 31 ms *) + qed +next + case (wfE_fstI \ \ v1 b1 b2) + then show ?case using subst_vv.simps subst_ev.simps wf_subst1 Wellformed.wfE_fstI + proof - + show ?thesis + by (metis (full_types) subst_ev.simps(5) wfE_fstI.hyps(1) wfE_fstI.hyps(4) wfE_fstI.hyps(5) wfE_fstI.prems(1) wfE_fstI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_fstI wf_subst1(1)) (* 78 ms *) + qed +next + case (wfE_sndI \ \ v1 b1 b2) + then show ?case + by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_sndI wf_subst1(1)) +next + case (wfE_concatI \ \ \ \ v1 v2) + then show ?case + by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_concatI wf_subst1(1)) +next + case (wfE_splitI \ \ \ \ v1 v2) + then show ?case + by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_splitI wf_subst1(1)) +next + case (wfE_lenI \ \ \ \ v1) +then show ?case + by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_lenI wf_subst1(1)) +next + case (wfE_appI \ \ \ \ f x b c \ s' v) +then show ?case + by (metis (full_types) subst_ev.simps wfE_sndI Wellformed.wfE_appI wf_subst1(1)) +next + case (wfE_appPI \ \ \ \ \ b' bv1 v1 \1 f1 x1 b1 c1 s1) + show ?case proof(subst subst_ev.simps, rule) + show "\ \\<^sub>w\<^sub>f \" using wfE_appPI wfX_wfY by metis + show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v " using wfE_appPI by auto + show "Some (AF_fundef f1 (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s1))) = lookup_fun \ f1" using wfE_appPI by auto + show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v1[x::=v']\<^sub>v\<^sub>v : b1[bv1::=b']\<^sub>b " using wfE_appPI wf_subst1 by auto + show "\; \ \\<^sub>w\<^sub>f b' " using wfE_appPI by auto + have "atom bv1 \ \[x::=v']\<^sub>\\<^sub>v" using fresh_subst_gv_if wfE_appPI by metis + moreover have "atom bv1 \ v1[x::=v']\<^sub>v\<^sub>v" using wfE_appPI fresh_subst_if by simp + moreover have "atom bv1 \ \[x::=v']\<^sub>\\<^sub>v" using wfE_appPI fresh_subst_dv_if by simp + ultimately show "atom bv1 \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, b', v1[x::=v']\<^sub>v\<^sub>v, (b_of \1)[bv1::=b']\<^sub>b)" + using wfE_appPI fresh_prodN by metis + qed +next + case (wfE_mvarI \ \ \ \ \ u \) + have " \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f (AE_mvar u) : b_of \[x::=v']\<^sub>\\<^sub>v" proof + show "\ \\<^sub>w\<^sub>f \ " using wfE_mvarI by auto + show "\; \ ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v " using wfE_mvarI by auto + show "(u, \[x::=v']\<^sub>\\<^sub>v) \ setD \[x::=v']\<^sub>\\<^sub>v" using wfE_mvarI subst_dv_member by auto + qed + thus ?case using subst_ev.simps b_of_subst by auto +next + case (wfD_emptyI \ \) + then show ?case using subst_dv.simps wf_intros wf_subst1 by auto +next + case (wfD_cons \ \ \ \ \ u) + moreover hence "u \ fst ` setD \[x::=v']\<^sub>\\<^sub>v" using subst_dv.simps subst_dv_iff using subst_dv_fst_eq by presburger + ultimately show ?case using subst_dv.simps Wellformed.wfD_cons wf_subst1 by auto +next + case (wfPhi_emptyI \) + then show ?case by auto +next + case (wfPhi_consI f \ \ ft) + then show ?case by auto +next + case (wfS_assertI \ \ \ x2 c \ \ s b) + show ?case unfolding subst_sv.simps proof + show \ \; \; \; (x2, B_bool, c[x::=v']\<^sub>c\<^sub>v) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ + using wfS_assertI(4)[of "(x2, B_bool, c) #\<^sub>\ \\<^sub>1" x ] wfS_assertI by auto + + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f c[x::=v']\<^sub>c\<^sub>v \ using wfS_assertI wf_subst1 by auto + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_assertI wf_subst1 by auto + show \atom x2 \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, c[x::=v']\<^sub>c\<^sub>v, b, s[x::=v']\<^sub>s\<^sub>v)\ + apply(unfold fresh_prodN,intro conjI) + apply(simp add: wfS_assertI )+ + apply(metis fresh_subst_gv_if wfS_assertI) + apply(simp add: fresh_prodN fresh_subst_dv_if wfS_assertI) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_assertI) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\_def wfS_assertI) + by(simp add: fresh_prodN fresh_subst_v_if subst_v_s_def wfS_assertI) + qed +next + case (wfS_letI \ \ \ \ \ e b1 y s b2) + have "\ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f LET y = (e[x::=v']\<^sub>e\<^sub>v) IN (s[x::=v']\<^sub>s\<^sub>v) : b2" + proof + show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f e[x::=v']\<^sub>e\<^sub>v : b1 \ using wfS_letI by auto + have \ \ ; \ ; \ ; ((y, b1, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \ + using wfS_letI(6) wfS_letI append_g.simps by metis + thus \ \ ; \ ; \ ; (y, b1, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b2 \ + using wfS_letI subst_gv.simps by auto + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_letI by auto + show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, e[x::=v']\<^sub>e\<^sub>v, b2)\ + apply(unfold fresh_prodN,intro conjI) + apply(simp add: wfS_letI )+ + apply(metis fresh_subst_gv_if wfS_letI) + apply(simp add: fresh_prodN fresh_subst_dv_if wfS_letI) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_letI) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\_def wfS_letI) + done + qed + thus ?case using subst_sv.simps wfS_letI by auto +next + case (wfS_let2I \ \ \ \ \ s1 \ y s2 b) + have "\ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f LET y : \[x::=v']\<^sub>\\<^sub>v = (s1[x::=v']\<^sub>s\<^sub>v) IN (s2[x::=v']\<^sub>s\<^sub>v) : b" + proof + show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s1[x::=v']\<^sub>s\<^sub>v : b_of (\[x::=v']\<^sub>\\<^sub>v) \ using wfS_let2I b_of_subst by simp + have \ \ ; \ ; \ ; ((y, b_of \, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \ + using wfS_let2I append_g.simps by metis + thus \ \ ; \ ; \ ; (y, b_of \[x::=v']\<^sub>\\<^sub>v, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s2[x::=v']\<^sub>s\<^sub>v : b \ + using wfS_let2I subst_gv.simps append_g.simps using b_of_subst by simp + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_let2I wf_subst1 by metis + show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, s1[x::=v']\<^sub>s\<^sub>v, b, \[x::=v']\<^sub>\\<^sub>v)\ + apply(unfold fresh_prodN,intro conjI) + apply(simp add: wfS_let2I )+ + apply(metis fresh_subst_gv_if wfS_let2I) + apply(simp add: fresh_prodN fresh_subst_dv_if wfS_let2I) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_e_def wfS_let2I) + apply(simp add: fresh_prodN fresh_subst_v_if subst_v_\_def wfS_let2I)+ + done + qed + thus ?case using subst_sv.simps(3) subst_tv.simps wfS_let2I by auto +next + case (wfS_varI \ \ \ \ v u \ \ b s) + show ?case proof(subst subst_sv.simps, auto simp add: u_fresh_xv,rule) + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_varI wf_subst1 by auto + have "b_of (\[x::=v']\<^sub>\\<^sub>v) = b_of \" using b_of_subst by auto + thus \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \[x::=v']\<^sub>\\<^sub>v \ using wfS_varI wf_subst1 by auto + have *:"atom u \ v'" using wfV_supp wfS_varI fresh_def by metis + show \atom u \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, v[x::=v']\<^sub>v\<^sub>v, b)\ + unfolding fresh_prodN apply(auto simp add: wfS_varI) + using wfS_varI fresh_subst_gv * fresh_subst_dv by metis+ + show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; (u, \[x::=v']\<^sub>\\<^sub>v) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ using wfS_varI by auto + qed +next + case (wfS_assignI u \ \ \ \ \ \ v) + show ?case proof(subst subst_sv.simps, rule wf_intros) + show \(u, \[x::=v']\<^sub>\\<^sub>v) \ setD \[x::=v']\<^sub>\\<^sub>v\ using subst_dv_iff wfS_assignI using subst_dv_fst_eq + using subst_dv_member by auto + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_assignI by auto + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : b_of \[x::=v']\<^sub>\\<^sub>v \ using wfS_assignI b_of_subst wf_subst1 by auto + show "\ \\<^sub>w\<^sub>f \ " using wfS_assignI by auto + qed +next + case (wfS_matchI \ \ \ v tid dclist \ \ cs b) + show ?case proof(subst subst_sv.simps, rule wf_intros) + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f v[x::=v']\<^sub>v\<^sub>v : B_id tid \ using wfS_matchI wf_subst1 by auto + show \AF_typedef tid dclist \ set \\ using wfS_matchI by auto + show \ \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dclist \\<^sub>w\<^sub>f subst_branchlv cs x v' : b \ using wfS_matchI by simp + show "\; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v " using wfS_matchI by auto + show "\ \\<^sub>w\<^sub>f \ " using wfS_matchI by auto + qed +next + case (wfS_branchI \ \ \ y \ \ \ s b tid dc) + have " \ ; \ ; \ ; \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v ; tid ; dc ; \ \\<^sub>w\<^sub>f dc y \ (s[x::=v']\<^sub>s\<^sub>v) : b" + proof + have \ \ ; \ ; \ ; ((y, b_of \, TRUE) #\<^sub>\ \)[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ + using wfS_branchI append_g.simps by metis + thus \ \ ; \ ; \ ; (y, b_of \, TRUE) #\<^sub>\ \[x::=v']\<^sub>\\<^sub>v ; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f s[x::=v']\<^sub>s\<^sub>v : b \ + using subst_gv.simps b_of_subst wfS_branchI by simp + show \atom y \ (\, \, \, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, \[x::=v']\<^sub>\\<^sub>v, \)\ + apply(unfold fresh_prodN,intro conjI) + apply(simp add: wfS_branchI )+ + apply(metis fresh_subst_gv_if wfS_branchI) + apply(simp add: fresh_prodN fresh_subst_dv_if wfS_branchI) + apply(metis fresh_subst_gv_if wfS_branchI)+ + done + show \ \; \; \[x::=v']\<^sub>\\<^sub>v \\<^sub>w\<^sub>f \[x::=v']\<^sub>\\<^sub>v \ using wfS_branchI by auto + qed + thus ?case using subst_branchv.simps wfS_branchI by auto +next + case (wfS_finalI \ \ \ \ \ tid dclist' cs b dclist) + then show ?case using subst_branchlv.simps wf_intros by metis +next + case (wfS_cons \ \ \ \ \ tid dclist' cs b css dclist) + then show ?case using subst_branchlv.simps wf_intros by metis + +qed(metis subst_sv.simps wf_subst1 wf_intros)+ + +lemmas wf_subst = wf_subst1 wf_subst2 + +lemma wfG_subst_wfV: + assumes "\; \ \\<^sub>w\<^sub>f \' @ (x, b, c0[z0::=V_var x]\<^sub>c\<^sub>v) #\<^sub>\ \" and "wfV \ \ \ v b" + shows "\; \ \\<^sub>w\<^sub>f \'[x::=v]\<^sub>\\<^sub>v @ \ " + using assms wf_subst subst_g_inside_simple by auto + +lemma wfG_member_subst: + assumes "(x1,b1,c1) \ toSet (\'@\)" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" and "x \ x1" + shows "\c1'. (x1,b1,c1') \ toSet ((\'[x::=v]\<^sub>\\<^sub>v)@\)" +proof - + consider (lhs) "(x1,b1,c1) \ toSet \'" | (rhs) "(x1,b1,c1) \ toSet \" using append_g_toSetU assms by auto + thus ?thesis proof(cases) + case lhs + hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis + hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto + then show ?thesis by auto + next + case rhs + hence "(x1,b1,c1) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto + then show ?thesis by auto + qed +qed + +lemma wfG_member_subst2: + assumes "(x1,b1,c1) \ toSet (\'@((x,b,c) #\<^sub>\\))" and "wfG \ \ (\'@((x,b,c) #\<^sub>\\))" and "x \ x1" + shows "\c1'. (x1,b1,c1') \ toSet ((\'[x::=v]\<^sub>\\<^sub>v)@\)" +proof - + consider (lhs) "(x1,b1,c1) \ toSet \'" | (rhs) "(x1,b1,c1) \ toSet \" using append_g_toSetU assms by auto + thus ?thesis proof(cases) + case lhs + hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v)" using wfG_inside_fresh[THEN subst_gv_member_iff[OF lhs]] assms by metis + hence "(x1,b1,c1[x::=v]\<^sub>c\<^sub>v) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto + then show ?thesis by auto + next + case rhs + hence "(x1,b1,c1) \ toSet (\'[x::=v]\<^sub>\\<^sub>v@\)" using append_g_toSetU by auto + then show ?thesis by auto + qed +qed + +lemma wbc_subst: + fixes \::\ and \'::\ and v::v + assumes "wfC \ \ (\'@((x,b,c') #\<^sub>\\)) c" and "\; \; \ \\<^sub>w\<^sub>f v : b" + shows "\; \; ((\'[x::=v]\<^sub>\\<^sub>v)@\) \\<^sub>w\<^sub>f c[x::=v]\<^sub>c\<^sub>v" +proof - + have "(\'@((x,b,c') #\<^sub>\\))[x::=v]\<^sub>\\<^sub>v = ((\'[x::=v]\<^sub>\\<^sub>v)@\)" using assms subst_g_inside_simple wfC_wf by metis + thus ?thesis using wf_subst1(2)[OF assms(1) _ assms(2)] by metis +qed + +lemma wfG_inside_fresh_suffix: + assumes "wfG P B (\'@(x,b,c) #\<^sub>\\)" + shows "atom x \ \" +proof - + have "wfG P B ((x,b,c) #\<^sub>\\)" using wfG_suffix assms by auto + thus ?thesis using wfG_elims by metis +qed + +lemmas wf_b_subst_lemmas = subst_eb.simps wf_intros + forget_subst subst_b_b_def subst_b_v_def subst_b_ce_def fresh_e_opp_all subst_bb.simps wfV_b_fresh ms_fresh_all(6) + +lemma wf_b_subst1: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows "\ ; B' ; \ \\<^sub>w\<^sub>f v : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and + "\ ; B' ; \ \\<^sub>w\<^sub>f c \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b" and + "\ ; B' \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and + "\ ; B' ; \ \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \True" and + "\ ; B' \\<^sub>w\<^sub>f b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " and + "\ ; B' ; \ \\<^sub>w\<^sub>f ce : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f ce[bv::=b]\<^sub>c\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b' and c and \ and \ and ts and \ and b' and b' and td + avoiding: bv b B + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfB_intI \ \) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfB_boolI \ \) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfB_unitI \ \) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfB_bitvecI \ \) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfB_pairI \ \ b1 b2) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfB_consI \ s dclist \) + then show ?case using subst_bb.simps Wellformed.wfB_consI by simp +next + case (wfB_appI \ ba s bva dclist \) + then show ?case using subst_bb.simps Wellformed.wfB_appI forget_subst wfB_supp + by (metis bot.extremum_uniqueI ex_in_conv fresh_def subst_b_b_def supp_empty_fset) +next + case (wfV_varI \ \1 \ b1 c x) + show ?case unfolding subst_vb.simps proof + show "\ ; B \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfV_varI by auto + show "Some (b1[bv::=b]\<^sub>b\<^sub>b, c[bv::=b]\<^sub>c\<^sub>b) = lookup \[bv::=b]\<^sub>\\<^sub>b x" using subst_b_lookup wfV_varI by simp + qed +next + case (wfV_litI \ \ \ l) + then show ?case using Wellformed.wfV_litI subst_b_base_for_lit by simp +next + case (wfV_pairI \ \1 \ v1 b1 v2 b2) + show ?case unfolding subst_vb.simps proof(subst subst_bb.simps,rule) + show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv::=b]\<^sub>b\<^sub>b" using wfV_pairI by simp + show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : b2[bv::=b]\<^sub>b\<^sub>b " using wfV_pairI by simp + qed +next + case (wfV_consI s dclist \ dc x b' c \' \ v) + show ?case unfolding subst_vb.simps proof(subst subst_bb.simps, rule Wellformed.wfV_consI) + show 1:"AF_typedef s dclist \ set \" using wfV_consI by auto + show 2:"(dc, \ x : b' | c \) \ set dclist" using wfV_consI by auto + have "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" using wfV_consI by auto + moreover hence "supp b' = {}" using 1 2 wfTh_lookup_supp_empty \.supp wfX_wfY by blast + moreover hence "b'[bv::=b]\<^sub>b\<^sub>b = b'" using forget_subst subst_bb_def fresh_def by (metis empty_iff subst_b_b_def) + ultimately show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfV_consI by simp + qed +next + case (wfV_conspI s bva dclist \ dc x b' c \' ba \ v) + have *:"atom bv \ b'" using wfTh_poly_supp_b[of s bva dclist \ dc x b' c] fresh_def wfX_wfY \atom bva \ bv\ + by (metis insert_iff not_self_fresh singleton_insert_inj_eq' subsetI subset_antisym wfV_conspI wfV_conspI.hyps(4) wfV_conspI.prems(2)) + show ?case unfolding subst_vb.simps subst_bb.simps proof + show \AF_typedef_poly s bva dclist \ set \\ using wfV_conspI by auto + show \(dc, \ x : b' | c \) \ set dclist\ using wfV_conspI by auto + thus \ \ ; B \\<^sub>w\<^sub>f ba[bv::=b]\<^sub>b\<^sub>b \ using wfV_conspI by metis + have "atom bva \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfV_conspI by metis + moreover have "atom bva \ ba[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfV_conspI by metis + moreover have "atom bva \ v[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfV_conspI by metis + ultimately show \atom bva \ (\, B, \[bv::=b]\<^sub>\\<^sub>b, ba[bv::=b]\<^sub>b\<^sub>b, v[bv::=b]\<^sub>v\<^sub>b)\ + unfolding fresh_prodN using wfV_conspI fresh_def supp_fset by auto + show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'[bva::=ba[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b \ + using wfV_conspI subst_bb_commute[of bv b' bva ba b] * wfV_conspI by metis + qed +next + case (wfTI z \ \' \' b' c) + show ?case proof(subst subst_tb.simps, rule Wellformed.wfTI) + show "atom z \ (\, B, \'[bv::=b]\<^sub>\\<^sub>b)" using wfTI subst_g_b_x_fresh by simp + show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfTI by auto + show "\ ; B ; (z, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\ \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c[bv::=b]\<^sub>c\<^sub>b " using wfTI by simp + qed +next + case (wfC_eqI \ \' \ e1 b' e2) + thus ?case using Wellformed.wfC_eqI subst_db.simps subst_cb.simps wfC_eqI by metis +next + case (wfG_nilI \ \') + then show ?case using Wellformed.wfG_nilI subst_gb.simps by simp +next + case (wfG_cons1I c' \ \' \' x b') + show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons1I) + show "c'[bv::=b]\<^sub>c\<^sub>b \ {TRUE, FALSE}" using wfG_cons1I(1) + by(nominal_induct c' rule: c.strong_induct,auto+) + show "\ ; B \\<^sub>w\<^sub>f \'[bv::=b]\<^sub>\\<^sub>b " using wfG_cons1I by auto + show "atom x \ \'[bv::=b]\<^sub>\\<^sub>b" using wfG_cons1I subst_g_b_x_fresh by auto + show "\ ; B ; (x, b'[bv::=b]\<^sub>b\<^sub>b, TRUE) #\<^sub>\ \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f c'[bv::=b]\<^sub>c\<^sub>b" using wfG_cons1I by auto + show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons1I by auto + qed +next + case (wfG_cons2I c' \ \' \' x b') + show ?case proof(subst subst_gb.simps, rule Wellformed.wfG_cons2I) + show "c'[bv::=b]\<^sub>c\<^sub>b \ {TRUE, FALSE}" using wfG_cons2I by auto + show "\ ; B \\<^sub>w\<^sub>f \'[bv::=b]\<^sub>\\<^sub>b " using wfG_cons2I by auto + show "atom x \ \'[bv::=b]\<^sub>\\<^sub>b" using wfG_cons2I subst_g_b_x_fresh by auto + show "\ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b " using wfG_cons2I by auto + qed +next + case (wfCE_valI \ \ \ v b) + then show ?case using subst_ceb.simps wf_intros wfX_wfY + by (metis wf_b_subst_lemmas wfCE_b_fresh) +next + case (wfCE_plusI \ \ \ v1 v2) + then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY + by metis +next + case (wfCE_leqI \ \ \ v1 v2) + then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY + by metis +next + case (wfCE_eqI \ \ \ v1 b v2) + then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY + by metis +next + case (wfCE_fstI \ \ \ v1 b1 b2) + then show ?case + by (metis (no_types) subst_bb.simps(5) subst_ceb.simps(3) wfCE_fstI.hyps(2) + wfCE_fstI.prems(1) wfCE_fstI.prems(2) Wellformed.wfCE_fstI) +next + case (wfCE_sndI \ \ \ v1 b1 b2) + then show ?case + by (metis (no_types) subst_bb.simps(5) subst_ceb.simps wfCE_sndI.hyps(2) + wfCE_sndI wfCE_sndI.prems(2) Wellformed.wfCE_sndI) +next + case (wfCE_concatI \ \ \ v1 v2) + then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh + proof - + show ?thesis + using wfCE_concatI.hyps(2) wfCE_concatI.hyps(4) wfCE_concatI.prems(1) wfCE_concatI.prems(2) + Wellformed.wfCE_concatI by auto (* 46 ms *) + qed +next + case (wfCE_lenI \ \ \ v1) + then show ?case using subst_bb.simps subst_ceb.simps wf_intros wfX_wfY wf_b_subst_lemmas wfCE_b_fresh by metis +qed(auto simp add: wf_intros) + +lemma wf_b_subst2: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and s::s and b'::b and ce::ce and td::type_def + and cs::branch_s and css::branch_list + shows "\ ; \ ; B' ; \ ; \ \\<^sub>w\<^sub>f e : b' \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f e[bv::=b]\<^sub>e\<^sub>b : b'[bv::=b]\<^sub>b\<^sub>b" and + "\ ; \ ; \ ; \ ; \ \\<^sub>w\<^sub>f s : b \ True" and + "\ ; \ ; \ ; \ ; \ ; tid ; dc ; t \\<^sub>w\<^sub>f cs : b \ True" and + "\ ; \ ; \ ; \ ; \ ; tid ; dclist \\<^sub>w\<^sub>f css : b \ True" and + "\ \\<^sub>w\<^sub>f (\::\) \ True " and + "\ ; B' ; \ \\<^sub>w\<^sub>f \ \ {|bv|} = B' \ \ ; B \\<^sub>w\<^sub>f b \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" and + "\ ; \ \\<^sub>w\<^sub>f ftq \ True" and + "\ ; \ ; \ \\<^sub>w\<^sub>f ft \ True" +proof(nominal_induct + b' and b and b and b and \ and \ and ftq and ft + avoiding: bv b B +rule:wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.strong_induct) + case (wfE_valI \' \' \' \' \' v' b') + then show ?case unfolding subst_vb.simps subst_eb.simps using wf_b_subst1(1) Wellformed.wfE_valI by auto +next + case (wfE_plusI \ \ \ \ \ v1 v2) + then show ?case unfolding subst_eb.simps + using wf_b_subst_lemmas wf_b_subst1(1) Wellformed.wfE_plusI + proof - + have "\b ba v g f ts. (( ts ; f ; g[bv::=ba]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=ba]\<^sub>v\<^sub>b : b[bv::=ba]\<^sub>b\<^sub>b) \ \ ts ; \ ; g \\<^sub>w\<^sub>f v : b ) \ \ ts ; f \\<^sub>w\<^sub>f ba" + using wfE_plusI.prems(1) wf_b_subst1(1) by force (* 0.0 ms *) + then show "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f [ plus v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_int[bv::=b]\<^sub>b\<^sub>b" + + by (metis wfE_plusI.hyps(1) wfE_plusI.hyps(4) wfE_plusI.hyps(5) wfE_plusI.hyps(6) wfE_plusI.prems(1) wfE_plusI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_plusI wf_b_subst_lemmas(86)) + qed +next + case (wfE_leqI \ \ \ \ \ v1 v2) + then show ?case unfolding subst_eb.simps + using wf_b_subst_lemmas wf_b_subst1 Wellformed.wfE_leqI + proof - + have "\ts f b ba g v. \ (ts ; f \\<^sub>w\<^sub>f b) \ \ (ts ; {|ba|} ; g \\<^sub>w\<^sub>f v : B_int) \ (ts ; f ; g[ba::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[ba::=b]\<^sub>v\<^sub>b : B_int)" + by (metis wf_b_subst1(1) wf_b_subst_lemmas(86)) (* 46 ms *) + then show "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f [ leq v1[bv::=b]\<^sub>v\<^sub>b v2[bv::=b]\<^sub>v\<^sub>b ]\<^sup>e : B_bool[bv::=b]\<^sub>b\<^sub>b" + by (metis (no_types) wfE_leqI.hyps(1) wfE_leqI.hyps(4) wfE_leqI.hyps(5) wfE_leqI.hyps(6) wfE_leqI.prems(1) wfE_leqI.prems(2) wfE_wfS_wfCS_wfCSS_wfPhi_wfD_wfFTQ_wfFT.wfE_leqI wf_b_subst_lemmas(87)) (* 46 ms *) + qed +next + case (wfE_eqI \ \ \ \ \ v1 bb v2) + show ?case unfolding subst_eb.simps subst_bb.simps proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfX_wfY wfE_eqI by metis + show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wfX_wfY wfE_eqI by metis + show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : bb \ using subst_bb.simps wfE_eqI + by (metis (no_types, hide_lams) empty_iff insert_iff wf_b_subst1(1)) + show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v2[bv::=b]\<^sub>v\<^sub>b : bb \ using wfX_wfY wfE_eqI + by (metis insert_iff singleton_iff wf_b_subst1(1) wf_b_subst_lemmas(86) wf_b_subst_lemmas(87) wf_b_subst_lemmas(90)) + show \bb \ {B_bool, B_int, B_unit}\ using wfE_eqI by auto + qed +next + case (wfE_fstI \ \ \ \ \ v1 b1 b2) + then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(84) wf_b_subst1(1) Wellformed.wfE_fstI + by (metis wf_b_subst_lemmas(89)) +next + case (wfE_sndI \ \ \ \ \ v1 b1 b2) + then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_sndI + by (metis wf_b_subst_lemmas(89)) +next + case (wfE_concatI \ \ \ \ \ v1 v2) +then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_concatI + by (metis wf_b_subst_lemmas(91)) +next + case (wfE_splitI \ \ \ \ \ v1 v2) + then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_splitI + by (metis wf_b_subst_lemmas(89) wf_b_subst_lemmas(91)) +next + case (wfE_lenI \ \ \ \ \ v1) + then show ?case unfolding subst_eb.simps using wf_b_subst_lemmas(86) wf_b_subst1(1) Wellformed.wfE_lenI + by (metis wf_b_subst_lemmas(91) wf_b_subst_lemmas(89)) +next + case (wfE_appI \ \ \' \ \ f x b' c \ s v) + hence bf: "atom bv \ b'" using wfPhi_f_simple_wfT wfT_supp bv_not_in_dom_g wfPhi_f_simple_supp_b fresh_def by fast + hence bseq: "b'[bv::=b]\<^sub>b\<^sub>b = b'" using subst_bb.simps wf_b_subst_lemmas by metis + have "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f (AE_app f (v[bv::=b]\<^sub>v\<^sub>b)) : (b_of (\[bv::=b]\<^sub>\\<^sub>b))" + proof + show "\ \\<^sub>w\<^sub>f \" using wfE_appI by auto + show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfE_appI by simp + have "atom bv \ \" using wfPhi_f_simple_wfT[OF wfE_appI(5) wfE_appI(1),THEN wfT_supp] bv_not_in_dom_g fresh_def by force + hence " \[bv::=b]\<^sub>\\<^sub>b = \" using forget_subst subst_b_\_def by metis + thus "Some (AF_fundef f (AF_fun_typ_none (AF_fun_typ x b' c \[bv::=b]\<^sub>\\<^sub>b s))) = lookup_fun \ f" using wfE_appI by simp + show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v[bv::=b]\<^sub>v\<^sub>b : b'" using wfE_appI bseq wf_b_subst1 by metis + qed + then show ?case using subst_eb.simps b_of_subst_bb_commute by simp +next + case (wfE_appPI \ \ \ \ \ b' bv1 v1 \1 f x1 b1 c1 s1) + then have *: "atom bv \ b1" using wfPhi_f_supp(1) wfE_appPI(7,11) + by (metis fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps(1)) + have "\ ; \ ; B ; \[bv::=b]\<^sub>\\<^sub>b ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f AE_appP f b'[bv::=b]\<^sub>b\<^sub>b (v1[bv::=b]\<^sub>v\<^sub>b) : (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b" + proof + show \ \ \\<^sub>w\<^sub>f \ \ using wfE_appPI by auto + show \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b \ using wfE_appPI by auto + show \ \ ; B \\<^sub>w\<^sub>f b'[bv::=b]\<^sub>b\<^sub>b \ using wfE_appPI wf_b_subst1 by auto + have "atom bv1 \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfE_appPI by metis + moreover have "atom bv1 \ b'[bv::=b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis + moreover have "atom bv1 \ v1[bv::=b]\<^sub>v\<^sub>b" using fresh_subst_if subst_b_v_def wfE_appPI by metis + moreover have "atom bv1 \ \[bv::=b]\<^sub>\\<^sub>b" using fresh_subst_if subst_b_\_def wfE_appPI by metis + moreover have "atom bv1 \ (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b\<^sub>b" using fresh_subst_if subst_b_b_def wfE_appPI by metis + ultimately show "atom bv1 \ (\, \, B, \[bv::=b]\<^sub>\\<^sub>b, \[bv::=b]\<^sub>\\<^sub>b, b'[bv::=b]\<^sub>b\<^sub>b, v1[bv::=b]\<^sub>v\<^sub>b, (b_of \1)[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b)" + using wfE_appPI using fresh_def fresh_prodN subst_b_b_def by metis + show \Some (AF_fundef f (AF_fun_typ_some bv1 (AF_fun_typ x1 b1 c1 \1 s1))) = lookup_fun \ f\ using wfE_appPI by auto + + have \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b']\<^sub>b[bv::=b]\<^sub>b\<^sub>b \ + using wfE_appPI subst_b_b_def * wf_b_subst1 by metis + thus \ \ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f v1[bv::=b]\<^sub>v\<^sub>b : b1[bv1::=b'[bv::=b]\<^sub>b\<^sub>b]\<^sub>b \ + using subst_bb_commute subst_b_b_def * by auto + qed + moreover have "atom bv \ b_of \1" proof - + have "supp (b_of \1) \ { atom bv1 }" using wfPhi_f_poly_supp_b_of_t + using b_of.simps wfE_appPI wfPhi_f_supp(5) by simp + thus ?thesis using wfE_appPI + fresh_def fresh_finsert singleton_iff subsetD fresh_def supp_at_base wfE_appPI.hyps by metis + qed + ultimately show ?case using subst_eb.simps(3) subst_bb_commute subst_b_b_def * by simp +next + case (wfE_mvarI \ \ \' \ \ u \) + + have "\ ; \ ; B ; subst_gb \ bv b ; subst_db \ bv b \\<^sub>w\<^sub>f (AE_mvar u)[bv::=b]\<^sub>e\<^sub>b : (b_of (\[bv::=b]\<^sub>\\<^sub>b))" + proof(subst subst_eb.simps,rule Wellformed.wfE_mvarI) + show "\ \\<^sub>w\<^sub>f \ " using wfE_mvarI by simp + show "\ ; B ; \[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b" using wfE_mvarI by metis + show "(u, \[bv::=b]\<^sub>\\<^sub>b) \ setD \[bv::=b]\<^sub>\\<^sub>b" + using wfE_mvarI subst_db.simps set_insert subst_d_b_member by simp + qed + thus ?case using b_of_subst_bb_commute by auto + +next + case (wfS_seqI \ \ \ \ \ s1 s2 b) + then show ?case using subst_bb.simps wf_intros wfX_wfY by metis +next + case (wfD_emptyI \ \' \) + then show ?case using subst_db.simps Wellformed.wfD_emptyI wf_b_subst1 by simp +next + case (wfD_cons \ \' \' \ \ u) + show ?case proof(subst subst_db.simps, rule Wellformed.wfD_cons ) + show "\ ; B ; \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfD_cons by auto + show "\ ; B ; \'[bv::=b]\<^sub>\\<^sub>b \\<^sub>w\<^sub>f \[bv::=b]\<^sub>\\<^sub>b " using wfD_cons wf_b_subst1 by auto + show "u \ fst ` setD \[bv::=b]\<^sub>\\<^sub>b" using wfD_cons subst_b_lookup_d by metis + qed +next + case (wfS_assertI \ \ \ x c \ \ s b) + show ?case by auto +qed(auto) + +lemmas wf_b_subst = wf_b_subst1 wf_b_subst2 + +lemma wfT_subst_wfT: + fixes \::\ and b'::b and bv::bv + assumes "\ ; {|bv|} ; (x,b,c) #\<^sub>\GNil \\<^sub>w\<^sub>f \" and "\ ; B \\<^sub>w\<^sub>f b'" + shows "\ ; B ; (x,b[bv::=b']\<^sub>b\<^sub>b,c[bv::=b']\<^sub>c\<^sub>b) #\<^sub>\GNil \\<^sub>w\<^sub>f (\[bv::=b']\<^sub>\\<^sub>b)" +proof - + have "\ ; B ; ((x,b,c) #\<^sub>\GNil)[bv::=b']\<^sub>\\<^sub>b \\<^sub>w\<^sub>f (\[bv::=b']\<^sub>\\<^sub>b)" + using wf_b_subst assms by metis + thus ?thesis using subst_gb.simps wf_b_subst_lemmas wfCE_b_fresh by metis +qed + +lemma wf_trans: + fixes \::\ and \'::\ and v::v and e::e and c::c and \::\ and ts::"(string*\) list" and \::\ and b::b and ftq::fun_typ_q and ft::fun_typ and ce::ce and td::type_def and s::s + and cs::branch_s and css::branch_list and \::\ + shows "\; \; \ \\<^sub>w\<^sub>f v : b' \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f v : b'" and + "\; \; \ \\<^sub>w\<^sub>f c \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c" and + "\; \ \\<^sub>w\<^sub>f \ \ True" and + "\; \; \ \\<^sub>w\<^sub>f \ \ True" and + "\; \; \ \\<^sub>w\<^sub>f ts \ True" and + "\\<^sub>w\<^sub>f \ \True" and + "\; \ \\<^sub>w\<^sub>f b \ True " and + "\; \; \ \\<^sub>w\<^sub>f ce : b' \ \ = (x, b, c2) #\<^sub>\ G \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f c2 \ \; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f ce : b' " and + "\ \\<^sub>w\<^sub>f td \ True" +proof(nominal_induct + b' and c and \ and \ and ts and \ and b and b' and td + avoiding: c1 + arbitrary: \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 and \\<^sub>1 + rule:wfV_wfC_wfG_wfT_wfTs_wfTh_wfB_wfCE_wfTD.strong_induct) + case (wfV_varI \ \ \ b' c' x') + have wbg: "\; \ \\<^sub>w\<^sub>f (x, b, c1) #\<^sub>\ G " using wfC_wf wfV_varI by simp + show ?case proof(cases "x=x'") + case True + have "Some (b', c1) = lookup ((x, b, c1) #\<^sub>\ G) x'" using lookup.simps wfV_varI using True by auto + then show ?thesis using Wellformed.wfV_varI wbg by simp + next + case False + then have "Some (b', c') = lookup ((x, b, c1) #\<^sub>\ G) x'" using lookup.simps wfV_varI + by simp + then show ?thesis using Wellformed.wfV_varI wbg by simp + qed +next + case (wfV_conspI s bv dclist \ dc x1 b' c \ b1 \ v) + show ?case proof + show \AF_typedef_poly s bv dclist \ set \\ using wfV_conspI by auto + show \(dc, \ x1 : b' | c \) \ set dclist\ using wfV_conspI by auto + show \\; \ \\<^sub>w\<^sub>f b1 \ using wfV_conspI by auto + show \atom bv \ (\, \, (x, b, c1) #\<^sub>\ G, b1, v)\ unfolding fresh_prodN fresh_GCons using wfV_conspI fresh_prodN fresh_GCons by simp + show \\; \; (x, b, c1) #\<^sub>\ G \\<^sub>w\<^sub>f v : b'[bv::=b1]\<^sub>b\<^sub>b\ using wfV_conspI by auto + qed +qed( (auto | metis wfC_wf wf_intros) +) + + +end \ No newline at end of file diff --git a/thys/MiniSail/document/root.bib b/thys/MiniSail/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/MiniSail/document/root.bib @@ -0,0 +1,29 @@ +@article{Armstrong2019, +abstract = {Architecture specifications notionally define the fundamental interface between hardware and software: the envelope of allowed behaviour for processor implementations, and the basic assumptions for software development and verification. But in practice, they are typically prose and pseudocode documents, not rigorous or executable artifacts, leaving software and verification on shaky ground. In this paper, we present rigorous semantic models for the sequential behaviour of large parts of the mainstream ARMv8-A, RISC-V, and MIPS architectures, and the research CHERI-MIPS architecture, that are complete enough to boot operating systems, variously Linux, FreeBSD, or seL4. Our ARMv8-A models are automatically translated from authoritative ARM-internal definitions, and (in one variant) tested against the ARM Architecture Validation Suite. We do this using a custom language for ISA semantics, Sail, with a lightweight dependent type system, that supports automatic generation of emulator code in C and OCaml, and automatic generation of proof-assistant definitions for Isabelle, HOL4, and (currently only for MIPS) Coq. We use the former for validation, and to assess specification coverage. To demonstrate the usability of the latter, we prove (in Isabelle) correctness of a purely functional characterisation of ARMv8-A address translation. We moreover integrate the RISC-V model into the RMEM tool for (user-mode) relaxed-memory concurrency exploration. We prove (on paper) the soundness of the core Sail type system. We thereby take a big step towards making the architectural abstraction actually well-defined, establishing foundations for verification and reasoning.}, +author = {Armstrong, Alasdair and Pulte, Christopher and Flur, Shaked and Stark, Ian and Krishnaswami, Neel and Sewell, Peter and Bauereiss, Thomas and Campbell, Brian and Reid, Alastair and Gray, Kathryn E. and Norton, Robert M. and Mundkur, Prashanth and Wassell, Mark and French, Jon}, +doi = {10.1145/3290384}, +file = {:E$\backslash$:/Mark/PhD/ReadingMaterial/SAIL/sail-popl2019.pdf:pdf}, +issn = {2475-1421}, +journal = {Proceedings of the ACM on Programming Languages}, +number = {POPL}, +pages = {1--31}, +title = {{ISA semantics for ARMv8-A, RISC-V, and CHERI-MIPS}}, +volume = {3}, +year = {2019} +} +@article{Vazou2014, +abstract = {SMT-based checking of refinement types for call-by-value languages is a well-studied subject. Unfortunately, the classical translation of refinement types to verification conditions is unsound under lazy evaluation. When checking an expression, such systems implicitly assume that all the free variables in the expression are bound to values. This property is trivially guaranteed by eager, but does not hold under lazy, evaluation. Thus, to be sound and precise, a refinement type system for Haskell and the corresponding verification conditions must take into account which subset of binders actually reduces to values. We present a stratified type system that labels binders as potentially diverging or not, and that (circularly) uses refinement types to verify the labeling. We have implemented our system in LIQUIDHASKELL and present an experimental evaluation of our approach on more than 10,000 lines of widely used Haskell libraries. We show that LIQUIDHASKELL is able to prove 96{\%} of all recursive functions terminating, while requiring a modest 1.7 lines of termination-annotations per 100 lines of code.}, +annote = {ICFP Video - https://www.youtube.com/watch?v=vqvNQixKr6w}, +archivePrefix = {arXiv}, +arxivId = {arXiv:1604.02480v1}, +author = {Vazou, Niki and Seidel, Eric L and Peyton-jones, Simon}, +doi = {10.1145/2628136.2628161}, +eprint = {arXiv:1604.02480v1}, +file = {:C$\backslash$:/Users/User/AppData/Local/Mendeley Ltd./Mendeley Desktop/Downloaded/Merz, Vanzetto - 2014 - Refinement types for Haskell.pdf:pdf}, +isbn = {9781450328739}, +issn = {15232867}, +journal = {ICFP '14 Proceedings of the 19th ACM SIGPLAN international conference on Functional programming}, +mendeley-groups = {Semantic Tools/RefinementTypes}, +title = {{Refinement Types For Haskell}}, +year = {2014} +} diff --git a/thys/MiniSail/document/root.tex b/thys/MiniSail/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/MiniSail/document/root.tex @@ -0,0 +1,71 @@ +\documentclass[11pt,a4paper]{report} +\usepackage{graphicx,isabelle,isabellesym} + +\usepackage[a4paper,includeheadfoot,margin=2.54cm]{geometry} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + +\begin{document} + +\title{MiniSail} +\author{Mark P. Wassell} +\maketitle + +\begin{abstract} +MiniSail is a kernel language for Sail~\cite{Armstrong2019}, +an instruction set architecture (ISA) specification language. +Sail is an imperative language with a light-weight dependent type system similar to refinement type systems such as~\cite{Vazou2014}. +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. +\end{abstract} + +\tableofcontents + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\begin{center} + \includegraphics[width=\textwidth,height=\textheight,keepaspectratio]{session_graph} +\end{center} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + diff --git a/thys/Public_Announcement_Logic/PAL.thy b/thys/Public_Announcement_Logic/PAL.thy --- a/thys/Public_Announcement_Logic/PAL.thy +++ b/thys/Public_Announcement_Logic/PAL.thy @@ -1,469 +1,479 @@ +(* + File: PAL.thy + Author: Asta Halkjær From + + 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. +*) + theory PAL imports "Epistemic_Logic.Epistemic_Logic" begin section \Syntax\ datatype 'i pfm = FF ("\<^bold>\\<^sub>!") | Pro' id ("Pro\<^sub>!") | Dis \'i pfm\ \'i pfm\ (infixr "\<^bold>\\<^sub>!" 30) | Con \'i pfm\ \'i pfm\ (infixr "\<^bold>\\<^sub>!" 35) | Imp \'i pfm\ \'i pfm\ (infixr "\<^bold>\\<^sub>!" 25) | K' 'i \'i pfm\ ("K\<^sub>!") | Ann \'i pfm\ \'i pfm\ ("[_]\<^sub>! _" [50, 50] 50) abbreviation PIff :: \'i pfm \ 'i pfm \ 'i pfm\ (infixr "\<^bold>\\<^sub>!" 25) where \p \<^bold>\\<^sub>! q \ (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (q \<^bold>\\<^sub>! p)\ section \Semantics\ fun psemantics :: \('i, 'w) kripke \ 'w \ 'i pfm \ bool\ ("_, _ \\<^sub>! _" [50, 50] 50) and restrict :: \('i, 'w) kripke \ 'i pfm \ ('i, 'w) kripke\ where \(M, w \\<^sub>! \<^bold>\\<^sub>!) = False\ | \(M, w \\<^sub>! Pro\<^sub>! x) = \ M w x\ | \(M, w \\<^sub>! (p \<^bold>\\<^sub>! q)) = ((M, w \\<^sub>! p) \ (M, w \\<^sub>! q))\ | \(M, w \\<^sub>! (p \<^bold>\\<^sub>! q)) = ((M, w \\<^sub>! p) \ (M, w \\<^sub>! q))\ | \(M, w \\<^sub>! (p \<^bold>\\<^sub>! q)) = ((M, w \\<^sub>! p) \ (M, w \\<^sub>! q))\ | \(M, w \\<^sub>! K\<^sub>! i p) = (\v \ \ M i w. M, v \\<^sub>! p)\ | \(M, w \\<^sub>! [r]\<^sub>! p) = ((M, w \\<^sub>! r) \ (restrict M r, w \\<^sub>! p))\ | \restrict M p = Kripke (\ M) (\i w. {v. v \ \ M i w \ (M, v \\<^sub>! p)})\ primrec static :: \'i pfm \ bool\ where \static \<^bold>\\<^sub>! = True\ | \static (Pro\<^sub>! _) = True\ | \static (p \<^bold>\\<^sub>! q) = (static p \ static q)\ | \static (p \<^bold>\\<^sub>! q) = (static p \ static q)\ | \static (p \<^bold>\\<^sub>! q) = (static p \ static q)\ | \static (K\<^sub>! i p) = static p\ | \static ([r]\<^sub>! p) = False\ primrec lower :: \'i pfm \ 'i fm\ where \lower \<^bold>\\<^sub>! = \<^bold>\\ | \lower (Pro\<^sub>! x) = Pro x\ | \lower (p \<^bold>\\<^sub>! q) = (lower p \<^bold>\ lower q)\ | \lower (p \<^bold>\\<^sub>! q) = (lower p \<^bold>\ lower q)\ | \lower (p \<^bold>\\<^sub>! q) = (lower p \<^bold>\ lower q)\ | \lower (K\<^sub>! i p) = K i (lower p)\ | \lower ([r]\<^sub>! p) = undefined\ primrec lift :: \'i fm \ 'i pfm\ where \lift \<^bold>\ = \<^bold>\\<^sub>!\ | \lift (Pro x) = Pro\<^sub>! x\ | \lift (p \<^bold>\ q) = (lift p \<^bold>\\<^sub>! lift q)\ | \lift (p \<^bold>\ q) = (lift p \<^bold>\\<^sub>! lift q)\ | \lift (p \<^bold>\ q) = (lift p \<^bold>\\<^sub>! lift q)\ | \lift (K i p) = K\<^sub>! i (lift p)\ lemma lower_semantics: assumes \static p\ shows \(M, w \ lower p) \ (M, w \\<^sub>! p)\ using assms by (induct p arbitrary: w) simp_all lemma lift_semantics: \(M, w \ p) \ (M, w \\<^sub>! lift p)\ by (induct p arbitrary: w) simp_all lemma lower_lift: \lower (lift p) = p\ by (induct p) simp_all lemma lift_lower: \static p \ lift (lower p) = p\ by (induct p) simp_all section \Soundness of Reduction\ primrec reduce' :: \'i pfm \ 'i pfm \ 'i pfm\ where \reduce' r \<^bold>\\<^sub>! = (r \<^bold>\\<^sub>! \<^bold>\\<^sub>!)\ | \reduce' r (Pro\<^sub>! x) = (r \<^bold>\\<^sub>! Pro\<^sub>! x)\ | \reduce' r (p \<^bold>\\<^sub>! q) = (reduce' r p \<^bold>\\<^sub>! reduce' r q)\ | \reduce' r (p \<^bold>\\<^sub>! q) = (reduce' r p \<^bold>\\<^sub>! reduce' r q)\ | \reduce' r (p \<^bold>\\<^sub>! q) = (reduce' r p \<^bold>\\<^sub>! reduce' r q)\ | \reduce' r (K\<^sub>! i p) = (r \<^bold>\\<^sub>! K\<^sub>! i (reduce' r p))\ | \reduce' r ([p]\<^sub>! q) = undefined\ primrec reduce :: \'i pfm \ 'i pfm\ where \reduce \<^bold>\\<^sub>! = \<^bold>\\<^sub>!\ | \reduce (Pro\<^sub>! x) = Pro\<^sub>! x\ | \reduce (p \<^bold>\\<^sub>! q) = (reduce p \<^bold>\\<^sub>! reduce q)\ | \reduce (p \<^bold>\\<^sub>! q) = (reduce p \<^bold>\\<^sub>! reduce q)\ | \reduce (p \<^bold>\\<^sub>! q) = (reduce p \<^bold>\\<^sub>! reduce q)\ | \reduce (K\<^sub>! i p) = K\<^sub>! i (reduce p)\ | \reduce ([r]\<^sub>! p) = reduce' (reduce r) (reduce p)\ lemma static_reduce': \static p \ static r \ static (reduce' r p)\ by (induct p) simp_all lemma static_reduce: \static (reduce p)\ by (induct p) (simp_all add: static_reduce') lemma reduce'_semantics: assumes \static q\ shows \((M, w \\<^sub>! [p]\<^sub>! (q))) = (M, w \\<^sub>! reduce' p q)\ using assms by (induct q arbitrary: w) auto lemma reduce_semantics: \(M, w \\<^sub>! p) = (M, w \\<^sub>! reduce p)\ proof (induct p arbitrary: M w) case (Ann p q) then show ?case using reduce'_semantics static_reduce by fastforce qed simp_all section \Proof System\ primrec peval :: \(id \ bool) \ ('i pfm \ bool) \ 'i pfm \ bool\ where \peval _ _ \<^bold>\\<^sub>! = False\ | \peval g _ (Pro\<^sub>! x) = g x\ | \peval g h (p \<^bold>\\<^sub>! q) = (peval g h p \ peval g h q)\ | \peval g h (p \<^bold>\\<^sub>! q) = (peval g h p \ peval g h q)\ | \peval g h (p \<^bold>\\<^sub>! q) = (peval g h p \ peval g h q)\ | \peval _ h (K\<^sub>! i p) = h (K\<^sub>! i p)\ | \peval _ h ([r]\<^sub>! p) = h ([r]\<^sub>! p)\ abbreviation \ptautology p \ \g h. peval g h p\ inductive PAK :: \('i pfm \ bool) \ 'i pfm \ bool\ ("_ \\<^sub>! _" [50, 50] 50) for A :: \'i pfm \ bool\ where PA1: \ptautology p \ A \\<^sub>! p\ | PA2: \A \\<^sub>! (K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! K\<^sub>! i q)\ | PAx: \A p \ A \\<^sub>! p\ | PR1: \A \\<^sub>! p \ A \\<^sub>! (p \<^bold>\\<^sub>! q) \ A \\<^sub>! q\ | PR2: \A \\<^sub>! p \ A \\<^sub>! K\<^sub>! i p\ | PFF: \A \\<^sub>! ([r]\<^sub>! \<^bold>\\<^sub>! \<^bold>\\<^sub>! (r \<^bold>\\<^sub>! \<^bold>\\<^sub>!))\ | PPro: \A \\<^sub>! ([r]\<^sub>! Pro\<^sub>! x \<^bold>\\<^sub>! (r \<^bold>\\<^sub>! Pro\<^sub>! x))\ | PDis: \A \\<^sub>! ([r]\<^sub>! (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! [r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q)\ | PCon: \A \\<^sub>! ([r]\<^sub>! (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! [r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q)\ | PImp: \A \\<^sub>! (([r]\<^sub>! (p \<^bold>\\<^sub>! q)) \<^bold>\\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q))\ | PK: \A \\<^sub>! (([r]\<^sub>! K\<^sub>! i p) \<^bold>\\<^sub>! (r \<^bold>\\<^sub>! K\<^sub>! i ([r]\<^sub>! p)))\ | PAnn: \A \\<^sub>! p \ A \\<^sub>! [r]\<^sub>! p\ lemma eval_peval: \eval h (g o lift) p = peval h g (lift p)\ by (induct p) simp_all lemma tautology_ptautology: \tautology p \ ptautology (lift p)\ using eval_peval by blast lemma peval_eval: assumes \static p\ shows \eval h g (lower p) = peval h (g o lower) p\ using assms by (induct p) simp_all lemma ptautology_tautology: assumes \static p\ shows \ptautology p \ tautology (lower p)\ using assms peval_eval by blast theorem AK_PAK: \A o lift \ p \ A \\<^sub>! lift p\ by (induct p rule: AK.induct) (auto intro: PAK.intros(1-5) simp: tautology_ptautology) theorem static_completeness: assumes \static p\ \\(M :: ('i :: countable, 'i fm set) kripke) w. M, w \\<^sub>! p\ shows \A \\<^sub>! p\ proof - have \\(M :: ('i :: countable, 'i fm set) kripke) w. M, w \ lower p\ using assms by (simp add: lower_semantics) then have \A o lift \ lower p\ by (simp add: completeness) then have \A \\<^sub>! lift (lower p)\ using AK_PAK by fast then show ?thesis using assms(1) lift_lower by metis qed section \Soundness\ lemma peval_semantics: \peval (val w) (\q. Kripke val r, w \\<^sub>! q) p = (Kripke val r, w \\<^sub>! p)\ by (induct p) simp_all lemma ptautology: assumes \ptautology p\ shows \M, w \\<^sub>! p\ proof - from assms have \peval (g w) (\q. Kripke g r, w \\<^sub>! q) p\ for g r by simp then have \Kripke g r, w \\<^sub>! p\ for g r using peval_semantics by fast then show \M, w \\<^sub>! p\ by (metis kripke.collapse) qed theorem soundness: fixes M :: \('i, 'w) kripke\ assumes \\(M :: ('i, 'w) kripke) w p. A p \ P M \ M, w \\<^sub>! p\ \\(M :: ('i, 'w) kripke) p. P M \ P (restrict M p)\ shows \A \\<^sub>! p \ P M \ M, w \\<^sub>! p\ proof (induct p arbitrary: M w rule: PAK.induct) case (PAnn p r) moreover have \P (restrict M r)\ using PAnn(3) assms(2) by simp ultimately show ?case by simp qed (simp_all add: assms ptautology) corollary \(\_. False) \\<^sub>! p \ M, w \\<^sub>! p\ using soundness[where P=\\_. True\] by metis section \Completeness\ lemma ConE: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ shows \A \\<^sub>! p\ \A \\<^sub>! q\ using assms by (metis PA1 PR1 peval.simps(4-5))+ lemma Iff_Dis: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! p')\ \A \\<^sub>! (q \<^bold>\\<^sub>! q')\ shows \A \\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q'))\ proof - have \A \\<^sub>! ((p \<^bold>\\<^sub>! p') \<^bold>\\<^sub>! (q \<^bold>\\<^sub>! q') \<^bold>\\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q')))\ by (simp add: PA1) then show ?thesis using assms PR1 by blast qed lemma Iff_Con: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! p')\ \A \\<^sub>! (q \<^bold>\\<^sub>! q')\ shows \A \\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q'))\ proof - have \A \\<^sub>! ((p \<^bold>\\<^sub>! p') \<^bold>\\<^sub>! (q \<^bold>\\<^sub>! q') \<^bold>\\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q')))\ by (simp add: PA1) then show ?thesis using assms PR1 by blast qed lemma Iff_Imp: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! p')\ \A \\<^sub>! (q \<^bold>\\<^sub>! q')\ shows \A \\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q'))\ proof - have \A \\<^sub>! ((p \<^bold>\\<^sub>! p') \<^bold>\\<^sub>! (q \<^bold>\\<^sub>! q') \<^bold>\\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q')))\ by (simp add: PA1) then show ?thesis using assms PR1 by blast qed lemma Iff_sym: \(A \\<^sub>! (p \<^bold>\\<^sub>! q)) = (A \\<^sub>! (q \<^bold>\\<^sub>! p))\ proof - have \A \\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (q \<^bold>\\<^sub>! p))\ by (simp add: PA1) then show ?thesis using PR1 ConE by blast qed lemma Iff_Iff: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! p')\ \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ shows \A \\<^sub>! (p' \<^bold>\\<^sub>! q)\ proof - have \ptautology ((p \<^bold>\\<^sub>! p') \<^bold>\\<^sub>! (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q))\ by (metis peval.simps(4-5)) with PA1 have \A \\<^sub>! ((p \<^bold>\\<^sub>! p') \<^bold>\\<^sub>! (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! (p' \<^bold>\\<^sub>! q))\ . then show ?thesis using assms PR1 by blast qed lemma K'_A2': \A \\<^sub>! (K\<^sub>! i (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i q)\ proof - have \A \\<^sub>! (K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! K\<^sub>! i q)\ using PA2 by fast moreover have \A \\<^sub>! ((P \<^bold>\\<^sub>! Q \<^bold>\\<^sub>! R) \<^bold>\\<^sub>! (Q \<^bold>\\<^sub>! P \<^bold>\\<^sub>! R))\ for P Q R by (simp add: PA1) ultimately show ?thesis using PR1 by fast qed lemma K'_map: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ shows \A \\<^sub>! (K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i q)\ proof - note \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ then have \A \\<^sub>! K\<^sub>! i (p \<^bold>\\<^sub>! q)\ using PR2 by fast moreover have \A \\<^sub>! (K\<^sub>! i (p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i q)\ using K'_A2' by fast ultimately show ?thesis using PR1 by fast qed lemma ConI: assumes \A \\<^sub>! p\ \A \\<^sub>! q\ shows \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ proof - have \A \\<^sub>! (p \<^bold>\\<^sub>! q \<^bold>\\<^sub>! p \<^bold>\\<^sub>! q)\ by (simp add: PA1) then show ?thesis using assms PR1 by blast qed lemma Iff_wk: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! q)\ shows \A \\<^sub>! ((r \<^bold>\\<^sub>! p) \<^bold>\\<^sub>! (r \<^bold>\\<^sub>! q))\ proof - have \A \\<^sub>! ((p \<^bold>\\<^sub>! q) \<^bold>\\<^sub>! ((r \<^bold>\\<^sub>! p) \<^bold>\\<^sub>! (r \<^bold>\\<^sub>! q)))\ by (simp add: PA1) then show ?thesis using assms PR1 by blast qed lemma Iff_reduce': assumes \static p\ shows \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! reduce' r p)\ using assms proof (induct p rule: pfm.induct) case FF then show ?case by (simp add: PFF) next case (Pro' x) then show ?case by (simp add: PPro) next case (Dis p q) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q \<^bold>\\<^sub>! reduce' r (p \<^bold>\\<^sub>! q))\ using Iff_Dis by force moreover have \A \\<^sub>! (([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q) \<^bold>\\<^sub>! ([r]\<^sub>! (p \<^bold>\\<^sub>! q)))\ using PDis Iff_sym by fastforce ultimately show ?case using PA1 PR1 Iff_Iff by blast next case (Con p q) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q \<^bold>\\<^sub>! reduce' r (p \<^bold>\\<^sub>! q))\ using Iff_Con by force moreover have \A \\<^sub>! (([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q) \<^bold>\\<^sub>! ([r]\<^sub>! (p \<^bold>\\<^sub>! q)))\ using PCon Iff_sym by fastforce ultimately show ?case using PA1 PR1 Iff_Iff by blast next case (Imp p q) then have \A \\<^sub>! (([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q) \<^bold>\\<^sub>! reduce' r (p \<^bold>\\<^sub>! q))\ using Iff_Imp by force moreover have \A \\<^sub>! (([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q) \<^bold>\\<^sub>! ([r]\<^sub>! (p \<^bold>\\<^sub>! q)))\ using PImp Iff_sym by fastforce ultimately show ?case using PA1 PR1 Iff_Iff by blast next case (K' i p) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! reduce' r p)\ by simp then have \A \\<^sub>! (K\<^sub>! i ([r]\<^sub>! p) \<^bold>\\<^sub>! K\<^sub>! i (reduce' r p))\ using K'_map ConE ConI by metis moreover have \A \\<^sub>! ([r]\<^sub>! K\<^sub>! i p \<^bold>\\<^sub>! r \<^bold>\\<^sub>! K\<^sub>! i ([r]\<^sub>! p))\ using PK . ultimately have \A \\<^sub>! ([r]\<^sub>! K\<^sub>! i p \<^bold>\\<^sub>! r \<^bold>\\<^sub>! K\<^sub>! i (reduce' r p))\ by (meson Iff_Iff Iff_sym Iff_wk) then show ?case by simp next case (Ann r p) then show ?case by simp qed lemma Iff_Ann1: assumes r: \A \\<^sub>! (r \<^bold>\\<^sub>! r')\ and \static p\ shows \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r']\<^sub>! p)\ using assms(2-) proof (induct p) case FF have \A \\<^sub>! ((r \<^bold>\\<^sub>! r') \<^bold>\\<^sub>! ((r \<^bold>\\<^sub>! \<^bold>\\<^sub>!) \<^bold>\\<^sub>! (r' \<^bold>\\<^sub>! \<^bold>\\<^sub>!)))\ by (auto intro: PA1) then have \A \\<^sub>! ((r \<^bold>\\<^sub>! \<^bold>\\<^sub>!) \<^bold>\\<^sub>! (r' \<^bold>\\<^sub>! \<^bold>\\<^sub>!))\ using r PR1 by blast then show ?case by (meson PFF Iff_Iff Iff_sym) next case (Pro' x) have \A \\<^sub>! ((r \<^bold>\\<^sub>! r') \<^bold>\\<^sub>! ((r \<^bold>\\<^sub>! Pro\<^sub>! x) \<^bold>\\<^sub>! (r' \<^bold>\\<^sub>! Pro\<^sub>! x)))\ by (auto intro: PA1) then have \A \\<^sub>! ((r \<^bold>\\<^sub>! Pro\<^sub>! x) \<^bold>\\<^sub>! (r' \<^bold>\\<^sub>! Pro\<^sub>! x))\ using r PR1 by blast then show ?case by (meson PPro Iff_Iff Iff_sym) next case (Dis p q) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q \<^bold>\\<^sub>! [r']\<^sub>! p \<^bold>\\<^sub>! [r']\<^sub>! q)\ by (simp add: Iff_Dis) then show ?case by (meson PDis Iff_Iff Iff_sym) next case (Con p q) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q \<^bold>\\<^sub>! [r']\<^sub>! p \<^bold>\\<^sub>! [r']\<^sub>! q)\ by (simp add: Iff_Con) then show ?case by (meson PCon Iff_Iff Iff_sym) next case (Imp p q) then have \A \\<^sub>! (([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! q) \<^bold>\\<^sub>! ([r']\<^sub>! p \<^bold>\\<^sub>! [r']\<^sub>! q))\ by (simp add: Iff_Imp) then show ?case by (meson PImp Iff_Iff Iff_sym) next case (K' i p) then have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r']\<^sub>! p)\ by simp then have \A \\<^sub>! (K\<^sub>! i ([r]\<^sub>! p) \<^bold>\\<^sub>! K\<^sub>! i ([r']\<^sub>! p))\ using K'_map ConE ConI by metis then show ?case by (meson Iff_Iff Iff_Imp Iff_sym PK r) next case (Ann s p) then show ?case by simp qed lemma Iff_Ann2: assumes \A \\<^sub>! (p \<^bold>\\<^sub>! p')\ shows \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! [r]\<^sub>! p')\ using assms PAnn ConE ConI PImp PR1 by metis lemma Iff_reduce: \A \\<^sub>! (p \<^bold>\\<^sub>! reduce p)\ proof (induct p) case (Dis p q) then show ?case by (simp add: Iff_Dis) next case (Con p q) then show ?case by (simp add: Iff_Con) next case (Imp p q) then show ?case by (simp add: Iff_Imp) next case (K' i p) have \A \\<^sub>! (K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i (reduce p))\ \A \\<^sub>! (K\<^sub>! i (reduce p) \<^bold>\\<^sub>! K\<^sub>! i p)\ using K' K'_map ConE by fast+ then have \A \\<^sub>! (K\<^sub>! i p \<^bold>\\<^sub>! K\<^sub>! i (reduce p))\ using ConI by blast then show ?case by simp next case (Ann r p) have \A \\<^sub>! ([reduce r]\<^sub>! reduce p \<^bold>\\<^sub>! reduce' (reduce r) (reduce p))\ using Iff_reduce' static_reduce by blast moreover have \A \\<^sub>! ([r]\<^sub>! reduce p \<^bold>\\<^sub>! [reduce r]\<^sub>! reduce p)\ using Ann(1) Iff_Ann1 static_reduce by blast ultimately have \A \\<^sub>! ([r]\<^sub>! reduce p \<^bold>\\<^sub>! reduce' (reduce r) (reduce p))\ using Iff_Iff Iff_sym by blast moreover have \A \\<^sub>! ([r]\<^sub>! reduce p \<^bold>\\<^sub>! [r]\<^sub>! p)\ by (meson Ann(2) static_reduce Iff_Ann2 Iff_sym) ultimately have \A \\<^sub>! ([r]\<^sub>! p \<^bold>\\<^sub>! reduce' (reduce r) (reduce p))\ using Iff_Iff by blast then show ?case by simp qed (simp_all add: PA1) theorem completeness: assumes \\(M :: ('i :: countable, 'i fm set) kripke) w. M, w \\<^sub>! p\ shows \A \\<^sub>! p\ proof - have \\(M :: ('i :: countable, 'i fm set) kripke) w. M, w \\<^sub>! reduce p\ using assms reduce_semantics by fast moreover have \static (reduce p)\ using static_reduce by fast ultimately have \A \\<^sub>! reduce p\ using static_completeness by blast moreover have \A \\<^sub>! (p \<^bold>\\<^sub>! reduce p)\ using Iff_reduce by blast ultimately show ?thesis using ConE(2) PR1 by blast qed -end \ No newline at end of file +end diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,605 +1,608 @@ 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 BNF_CC BNF_Operations BTree Banach_Steinhaus 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 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 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_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties -CSP_RefTK 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 Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos -DOM_Components 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_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF 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 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 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl +IsaGeoCoq Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model -IsaGeoCoq 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 -Knuth_Bendix_Order 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 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 -Modular_Assembly_Kit_Security 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 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 Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regression_Test_Selection Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_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 -Shadow_DOM 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 TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models 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 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 VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL diff --git a/thys/SpecCheck/README.md b/thys/SpecCheck/README.md new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/README.md @@ -0,0 +1,61 @@ +# SpecCheck + +SpecCheck is a [QuickCheck](https://en.wikipedia.org/wiki/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. + +## Quick Usage +1. Import `SpecCheck.SpecCheck` into your environment. +2. Write specifications using the ML invocation: `check show gen name prop ctxt seed` where + * `show` converts values into `Pretty.T` types to show the failing inputs. See `show/`. + * `gen` is the value generator used for the test. See `generators/`. + * `name` is the name of the test case + * `prop` is the property to be tested. See `property.ML` + * `seed` is the initial seed for the generator. + +You can also choose to omit the show method for rapid testing or add a shrinking method à la +QuickCheck to get better counterexamples. See `speccheck.ML`. + +An experimental alternative allows you to specify tests using strings: +1. Import `SpecCheck_Dynamic.Dynamic` into your environment. +2. `check_property "ALL x. P x"` where `P x` is some ML code evaluating to a boolean + +Examples can be found in `examples/`. + +## Notes + +SpecCheck is based on [QCheck](https://github.com/league/qcheck), a testing framework for Standard ML by +[Christopher League](https://contrapunctus.net/league/). +As Isabelle/ML provides a rich and uniform ML platform, some features where removed or adapted, in particular: + +1. Isabelle/ML provides common data structures, which we can use in the +tool's implementation for storing data and printing output. + +2. Implementations in Isabelle/ML checked with this tool commonly use Isabelle/ML's `int` type +(which corresponds ML's `IntInf.int`), but do not use other integer types in ML such as ML's `Int.int`, +`Word.word`, and others. + +3. As Isabelle makes heavy use of parallelism, we avoid reference types. + +## Next Steps + +* Implement more shrinking methods for commonly used types +* Implement sizing methods (cf. QuickCheck's `sized`) + +## License + +The source code originated from Christopher League's QCheck, which is +licensed under the 2-clause BSD license. The current source code is +licensed under the compatible 3-clause BSD license of Isabelle. + +## Authors + +* Lukas Bulwahn +* Nicolai Schaffroth +* Sebastian Willenbrink +* Kevin Kappelmann diff --git a/thys/SpecCheck/ROOT b/thys/SpecCheck/ROOT new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/ROOT @@ -0,0 +1,30 @@ +chapter AFP + +session SpecCheck (AFP) = "Pure" + + description + \SpecCheck is a specification-based testing environment for ML programs. It is based on QCheck + (\<^url>\https://github.com/league/qcheck/\) by Christopher League (\<^url>\https://contrapunctus.net/\). + It got adapted and extended to fit into the Isabelle/ML framework and resemble the very + successful QuickCheck (\<^url>\https://en.wikipedia.org/wiki/QuickCheck\) more closely.\ + options [timeout = 300] + + directories + dynamic + examples + generators + output_styles + "show" + shrink + + theories + SpecCheck_Generators + SpecCheck_Output_Style + SpecCheck_Show + SpecCheck_Shrink + SpecCheck + SpecCheck_Dynamic + theories [document = false] + SpecCheck_Examples + + document_files + "root.tex" diff --git a/thys/SpecCheck/SpecCheck.thy b/thys/SpecCheck/SpecCheck.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/SpecCheck.thy @@ -0,0 +1,17 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \SpecCheck\ +theory SpecCheck +imports + SpecCheck_Generators + SpecCheck_Show + SpecCheck_Shrink + SpecCheck_Output_Style +begin + +paragraph \Summary\ +text \The SpecCheck (specification based) testing environment and Lecker testing framework.\ + +ML_file \speccheck.ML\ +ML_file \lecker.ML\ + +end diff --git a/thys/SpecCheck/SpecCheck_Base.thy b/thys/SpecCheck/SpecCheck_Base.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/SpecCheck_Base.thy @@ -0,0 +1,18 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \SpecCheck Base\ +theory SpecCheck_Base +imports Pure +begin + +paragraph \Summary\ +text \Basic setup for SpecCheck.\ + +ML_file \util.ML\ + +ML_file \speccheck_base.ML\ +ML_file \property.ML\ +ML_file \configuration.ML\ + +ML_file \random.ML\ + +end diff --git a/thys/SpecCheck/configuration.ML b/thys/SpecCheck/configuration.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/configuration.ML @@ -0,0 +1,41 @@ +(* Title: SpecCheck/configuration.ML + Author: Kevin Kappelmann + +Configuration options for SpecCheck. +*) + +signature SPECCHECK_CONFIGURATION = +sig + + (*maximum number of successful tests before succeeding*) + val max_success : int Config.T + (*maximum number of discarded tests per successful test before giving up*) + val max_discard_ratio : int Config.T + (*maximum number of shrinks per counterexample*) + val max_shrinks : int Config.T + (*number of counterexamples shown*) + val num_counterexamples : int Config.T + (*sort counterexamples by size*) + val sort_counterexamples : bool Config.T + (*print timing etc. depending on style*) + val show_stats : bool Config.T + +end + +structure SpecCheck_Configuration : SPECCHECK_CONFIGURATION = +struct + +val max_success = Attrib.setup_config_int \<^binding>\speccheck_max_success\ (K 100) + +val max_discard_ratio = Attrib.setup_config_int \<^binding>\speccheck_max_discard_ratio\ (K 10) + +val max_shrinks = Attrib.setup_config_int \<^binding>\speccheck_max_shrinks\ (K 10000) + +val num_counterexamples = Attrib.setup_config_int \<^binding>\speccheck_num_counterexamples\ (K 1) + +val sort_counterexamples = + Attrib.setup_config_bool \<^binding>\speccheck_sort_counterexamples\ (K true) + +val show_stats = Attrib.setup_config_bool \<^binding>\speccheck_show_stats\ (K true) + +end diff --git a/thys/SpecCheck/document/root.tex b/thys/SpecCheck/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/document/root.tex @@ -0,0 +1,39 @@ +\documentclass[11pt,a4paper]{article} +\usepackage[T1]{fontenc} +\usepackage{isabelle,isabellesym} + +% 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{SpecCheck -- Specification-Based Testing for Isabelle/ML} +\author{Kevin Kappelmann, Lukas Bulwahn, and Sebastian Willenbrink} +\maketitle + +\begin{abstract} +SpecCheck is a \href{https://en.wikipedia.org/wiki/QuickCheck}{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. +\end{abstract} + +\tableofcontents + +% include generated text of all theories +\input{session} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} diff --git a/thys/SpecCheck/dynamic/SpecCheck_Dynamic.thy b/thys/SpecCheck/dynamic/SpecCheck_Dynamic.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/dynamic/SpecCheck_Dynamic.thy @@ -0,0 +1,16 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Dynamic Generators\ +theory SpecCheck_Dynamic +imports SpecCheck +begin + +paragraph \Summary\ +text \Generators and show functions for SpecCheck that are dynamically derived from a given ML input +string. This approach can be handy to quickly test a function during development, but it lacks +customisability and is very brittle. See @{file "../examples/SpecCheck_Examples.thy"} for some +examples contrasting this approach to the standard one (specifying generators as ML code).\ + +ML_file \dynamic_construct.ML\ +ML_file \speccheck_dynamic.ML\ + +end diff --git a/thys/SpecCheck/dynamic/dynamic_construct.ML b/thys/SpecCheck/dynamic/dynamic_construct.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/dynamic/dynamic_construct.ML @@ -0,0 +1,188 @@ +(* Title: SpecCheck/dynamic/dynamic_construct.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + +Dynamic construction of generators and show functions (returned as strings that need to be compiled) +from a given string representing ML code to be tested as a SpecCheck test. +*) + +signature SPECCHECK_DYNAMIC_CONSTRUCT = +sig + val register : string * (string * string) -> theory -> theory + type mltype + val parse_pred : string -> string * mltype + val build_check : Proof.context -> string -> mltype * string -> string + (*val safe_check : string -> mltype * string -> string*) + val string_of_bool : bool -> string + val string_of_ref : ('a -> string) -> 'a Unsynchronized.ref -> string +end; + +structure SpecCheck_Dynamic_Construct : SPECCHECK_DYNAMIC_CONSTRUCT = +struct + +(* Parsing ML types *) + +datatype mltype = Var | Con of string * mltype list | Tuple of mltype list; + +(*Split string into tokens for parsing*) +fun split s = + let + fun split_symbol #"(" = "( " + | split_symbol #")" = " )" + | split_symbol #"," = " ," + | split_symbol #":" = " :" + | split_symbol c = Char.toString c + fun is_space c = c = #" " + in String.tokens is_space (String.translate split_symbol s) end; + +(*Accept anything that is not a recognized symbol*) +val scan_name = Scan.one (fn s => not (String.isSubstring s "(),*->;")); + +(*Turn a type list into a nested Con*) +fun make_con [] = raise Empty + | make_con [c] = c + | make_con (Con (s, _) :: cl) = Con (s, [make_con cl]); + +(*Parse a type*) +fun parse_type s = (parse_fun || parse_tuple || parse_type_single) s + +and parse_type_arg s = (parse_tuple || parse_type_single) s + +and parse_type_single s = (parse_con || parse_type_basic) s + +and parse_type_basic s = (parse_var || $$ "(" |-- parse_type --| $$ ")") s + +and parse_list s = + ($$ "(" |-- parse_type -- Scan.repeat1 ($$ "," |-- parse_type) --| $$ ")" >> op::) s + +and parse_var s = (Scan.one (String.isPrefix "'") >> (fn _ => Var)) s + +and parse_con s = ((parse_con_nest + || parse_type_basic -- parse_con_nest >> (fn (b, Con (t, _) :: tl) => Con (t, [b]) :: tl) + || parse_list -- parse_con_nest >> (fn (l, Con (t, _) :: tl) => Con (t, l) :: tl)) + >> (make_con o rev)) s + +and parse_con_nest s = Scan.unless parse_var (Scan.repeat1 (scan_name >> (fn t => Con (t, [])))) s + +and parse_fun s = (parse_type_arg -- $$ "->" -- parse_type >> (fn ((a, f), r) => Con (f, [a, r]))) s + +and parse_tuple s = (parse_type_single -- Scan.repeat1 ($$ "*" |-- parse_type_single) + >> (fn (t, tl) => Tuple (t :: tl))) s; + +(*Parse entire type + name*) +fun parse_function s = + let + val p = $$ "val" |-- scan_name --| ($$ "=" -- $$ "fn" -- $$ ":") + val (name, ty) = p (split s) + val stop = Scan.stopper (fn _ => ";") (fn s => s = ";"); + val (typ, _) = Scan.finite stop parse_type ty + in (name, typ) end; + +(*Create desired output*) +fun parse_pred s = + let + val (name, Con ("->", t :: _)) = parse_function s + in (name, t) end; + +(* Construct Generators and Pretty Printers *) + +(*copied from smt_config.ML *) +fun string_of_bool b = if b then "true" else "false" + +fun string_of_ref f r = f (!r) ^ " ref"; + +val initial_content = Symtab.make [ + ("bool", ("SpecCheck_Generator.bernoulli 0.5", "Gen_Construction.string_of_bool")), + ("option", ("SpecCheck_Generator.option (SpecCheck_Generator.bernoulli (2.0 / 3.0))", + "ML_Syntax.print_option")), + ("list", ("SpecCheck_Generator.unfold_while (K (SpecCheck_Generator.bernoulli (2.0 / 3.0)))", + " ML_Syntax.print_list")), + ("unit", ("gen_unit", "fn () => \"()\"")), + ("int", ("SpecCheck_Generator.range_int (~2147483647,2147483647)", "string_of_int")), + ("real", ("SpecCheck_Generator.real", "string_of_real")), + ("char", ("SpecCheck_Generator.char", "fn c => \"#'\" ^ (Char.toString c) ^ \"'\"")), + ("string", ("SpecCheck_Generator.string (SpecCheck_Generator.nonneg 100) SpecCheck_Generator.char", + "ML_Syntax.print_string")), + ("->", ("SpecCheck_Generator.function' o snd", "fn (_, _) => fn _ => \"fn\"")), + ("typ", ("SpecCheck_Generator.typ'' (SpecCheck_Generator.return 8) (SpecCheck_Generator.nonneg 4) (SpecCheck_Generator.nonneg 4) (1,1,1)", + "Pretty.string_of o Syntax.pretty_typ (Context.the_local_context ())")), + ("term", ("SpecCheck_Generator.term_tree (fn h => fn _ => " + ^ "let val ngen = SpecCheck_Generator.nonneg (Int.max (0, 4-h))\n" + ^ " val aterm_gen = SpecCheck_Generator.aterm' (SpecCheck_Generator.return 8) ngen (1,1,1,0)\n" + ^ "in SpecCheck_Generator.zip aterm_gen ngen end)", + "Pretty.string_of o Syntax.pretty_term (Context.the_local_context ())"))] + +structure Data = Theory_Data +( + type T = (string * string) Symtab.table + val empty = initial_content + val extend = I + fun merge data : T = Symtab.merge (K true) data +) + +fun data_of ctxt tycon = + (case Symtab.lookup (Data.get (Proof_Context.theory_of ctxt)) tycon of + SOME data => data + | NONE => error ("No generator and printer defined for ML type constructor " ^ quote tycon)) + +val generator_of = fst oo data_of +val printer_of = snd oo data_of + +fun register (ty, data) = Data.map (Symtab.update (ty, data)) + +(* +fun remove_gen ty = gen_table := AList.delete (op =) ty (!gen_table); +*) + +fun combine dict [] = dict + | combine dict dicts = enclose "(" ")" dict ^ " " ^ enclose "(" ")" (commas dicts) + +fun compose_generator _ Var = "SpecCheck_Generator.range_int (~2147483647, 2147483647)" + | compose_generator ctxt (Con (s, types)) = + combine (generator_of ctxt s) (map (compose_generator ctxt) types) + | compose_generator ctxt (Tuple t) = + let + fun tuple_body t = space_implode "" + (map + (fn (ty, n) => implode ["val (x", string_of_int n, ", r", string_of_int n, ") = ", + compose_generator ctxt ty, " r", string_of_int (n - 1), " "]) + (t ~~ (1 upto (length t)))) + fun tuple_ret a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a)) + in + "fn r0 => let " ^ tuple_body t ^ + "in ((" ^ tuple_ret (length t) ^ "), r" ^ string_of_int (length t) ^ ") end" + end + +fun compose_printer _ Var = "Int.toString" + | compose_printer ctxt (Con (s, types)) = + combine (printer_of ctxt s) (map (compose_printer ctxt) types) + | compose_printer ctxt (Tuple t) = + let + fun tuple_head a = commas (map (fn n => "x" ^ string_of_int n) (1 upto a)) + fun tuple_body t = space_implode " ^ \", \" ^ " + (map (fn (ty, n) => "(" ^ compose_printer ctxt ty ^ ") x" ^ string_of_int n) + (t ~~ (1 upto (length t)))) + in implode ["fn (", tuple_head (length t), ") => \"(\" ^ ", tuple_body t, " ^ \")\""] end + +(*produce compilable string*) +fun build_check ctxt name (ty, spec) = implode ["SpecCheck.check (Pretty.str o (", + compose_printer ctxt ty, ")) (", compose_generator ctxt ty, ") \"", name, + "\" (SpecCheck_Property.prop (", spec, + ")) (Context.the_local_context ()) (SpecCheck_Random.new ());"] + +(*produce compilable string - non-eqtype functions*) +(* +fun safe_check name (ty, spec) = + let + val default = + (case AList.lookup (op =) (!gen_table) "->" of + NONE => ("gen_function_rand", "fn (_, _) => fn _ => \"fn\"") + | SOME entry => entry) + in + (gen_table := + AList.update (op =) ("->", ("gen_function_safe", "fn (_, _) => fn _ => \"fn\"")) (!gen_table); + build_check name (ty, spec) before + gen_table := AList.update (op =) ("->", default) (!gen_table)) + end; +*) + +end; diff --git a/thys/SpecCheck/dynamic/speccheck_dynamic.ML b/thys/SpecCheck/dynamic/speccheck_dynamic.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/dynamic/speccheck_dynamic.ML @@ -0,0 +1,81 @@ +(* Title: SpecCheck/dynamic/speccheck_dynamic.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + +This file allows to run SpecCheck tests specified as a string representing ML code. + +TODO: this module is not very well tested. +*) + +signature SPECCHECK_DYNAMIC = +sig + val check_dynamic : Proof.context -> string -> unit +end + +structure SpecCheck_Dynamic : SPECCHECK_DYNAMIC = +struct + +(*call the compiler and pass resulting type string to the parser*) +fun determine_type ctxt s = + let + val return = Unsynchronized.ref "return" + val context : ML_Compiler0.context = + {name_space = #name_space ML_Env.context, + print_depth = SOME 1000000, + here = #here ML_Env.context, + print = fn r => return := r, + error = #error ML_Env.context} + val _ = + Context.setmp_generic_context (SOME (Context.Proof ctxt)) + (fn () => + ML_Compiler0.ML context + {line = 0, file = "generated code", verbose = true, debug = false} s) () + in SpecCheck_Dynamic_Construct.parse_pred (! return) end; + +(*call the compiler and run the test*) +fun run_test ctxt s = + Context.setmp_generic_context (SOME (Context.Proof ctxt)) + (fn () => + ML_Compiler0.ML ML_Env.context + {line = 0, file = "generated code", verbose = false, debug = false} s) (); + +(*split input into tokens*) +fun input_split s = + let + fun dot c = c = #"." + fun space c = c = #" " + val (head, code) = Substring.splitl (not o dot) (Substring.full s) + in + (String.tokens space (Substring.string head), + Substring.string (Substring.dropl dot code)) + end; + +(*create the function from the input*) +fun make_fun s = + let + val scan_param = Scan.one (fn s => s <> ";") + fun parameters s = Scan.repeat1 scan_param s + val p = $$ "ALL" |-- parameters + val (split, code) = input_split s + val stop = Scan.stopper (fn _ => ";") (fn s => s = ";"); + val (params, _) = Scan.finite stop p split + in "fn (" ^ commas params ^ ") => " ^ code end; + +(*read input and perform the test*) +fun gen_check_property check ctxt s = + let + val func = make_fun s + val (_, ty) = determine_type ctxt func + in run_test ctxt (check ctxt "Dynamic Test" (ty, func)) end; + +val check_dynamic = gen_check_property SpecCheck_Dynamic_Construct.build_check +(*val check_property_safe = gen_check_property Construct_Gen.safe_check*) + +(*perform test for specification function*) +(*fun gen_check_property_f check ctxt s = + let + val (name, ty) = determine_type ctxt s + in run_test ctxt (check ctxt name (ty, s)) end; + +val check_property_f = gen_check_property_f Gen_Dynamic.build_check*) +(*val check_property_safe_f_ = gen_check_property_f Construct_Gen.safe_check*) +end; diff --git a/thys/SpecCheck/examples/SpecCheck_Examples.thy b/thys/SpecCheck/examples/SpecCheck_Examples.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/examples/SpecCheck_Examples.thy @@ -0,0 +1,129 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Examples\ +theory SpecCheck_Examples +imports SpecCheck_Dynamic +begin + +subsection \List examples\ + +ML \ +open SpecCheck +open SpecCheck_Dynamic +structure Gen = SpecCheck_Generator +structure Prop = SpecCheck_Property +structure Show = SpecCheck_Show +structure Shrink = SpecCheck_Shrink +structure Random = SpecCheck_Random +\ + +ML_command \ +let + val int_gen = Gen.range_int (~10000000, 10000000) + val size_gen = Gen.nonneg 10 + val check_list = check_shrink (Show.list Show.int) (Shrink.list Shrink.int) + (Gen.list size_gen int_gen) + fun list_test (k, f, xs) = + AList.lookup (op=) (AList.map_entry (op=) k f xs) k = Option.map f (AList.lookup (op=) xs k) + + val list_test_show = Show.zip3 Show.int Show.none (Show.list (Show.zip Show.int Show.int)) + val list_test_gen = Gen.zip3 int_gen (Gen.function' int_gen) + (Gen.list size_gen (Gen.zip int_gen int_gen)) +in + Lecker.test_group @{context} (Random.new ()) [ + Prop.prop (fn xs => rev xs = xs) |> check_list "rev = I", + Prop.prop (fn xs => rev (rev xs) = xs) |> check_list "rev o rev = I", + Prop.prop list_test |> check list_test_show list_test_gen "lookup map equiv map lookup" + ] +end +\ + +text \The next three examples roughly correspond to the above test group (except that there's no +shrinking). Compared to the string-based method, the method above is more flexible - you can change +your generators, shrinking methods, and show instances - and robust - you are not reflecting strings +(which might contain typos) but entering type-checked code. In exchange, it is a bit more work to +set up the generators. However, in practice, one shouldn't rely on default generators in most cases +anyway.\ + +ML_command \ +check_dynamic @{context} "ALL xs. rev xs = xs"; +\ + +ML_command \ +check_dynamic @{context} "ALL xs. rev (rev xs) = xs"; +\ + +subsection \AList Specification\ + +ML_command \ +(*map_entry applies the function to the element*) +check_dynamic @{context} (implode + ["ALL k f xs. AList.lookup (op =) (AList.map_entry (op =) k f xs) k = ", + "Option.map f (AList.lookup (op =) xs k)"]) +\ + +ML_command \ +(*update always results in an entry*) +check_dynamic @{context} "ALL k v xs. AList.defined (op =) (AList.update (op =) (k, v) xs) k"; +\ + +ML_command \ +(*update always writes the value*) +check_dynamic @{context} + "ALL k v xs. AList.lookup (op =) (AList.update (op =) (k, v) xs) k = SOME v"; +\ + +ML_command \ +(*default always results in an entry*) +check_dynamic @{context} "ALL k v xs. AList.defined (op =) (AList.default (op =) (k, v) xs) k"; +\ + +ML_command \ +(*delete always removes the entry*) +check_dynamic @{context} "ALL k xs. not (AList.defined (op =) (AList.delete (op =) k xs) k)"; +\ + +ML_command \ +(*default writes the entry iff it didn't exist*) +check_dynamic @{context} (implode + ["ALL k v xs. (AList.lookup (op =) (AList.default (op =) (k, v) xs) k = ", + "(if AList.defined (op =) xs k then AList.lookup (op =) xs k else SOME v))"]) +\ + +subsection \Examples on Types and Terms\ + +ML_command \ +check_dynamic @{context} "ALL f g t. map_types (g o f) t = (map_types f o map_types g) t"; +\ + +ML_command \ +check_dynamic @{context} "ALL f g t. map_types (f o g) t = (map_types f o map_types g) t"; +\ + + +text \One would think this holds:\ + +ML_command \ +check_dynamic @{context} "ALL t ts. strip_comb (list_comb (t, ts)) = (t, ts)" +\ + +text \But it only holds with this precondition:\ + +ML_command \ +check_dynamic @{context} + "ALL t ts. case t of _ $ _ => true | _ => strip_comb (list_comb (t, ts)) = (t, ts)" +\ + +subsection \Some surprises\ + +ML_command \ +check_dynamic @{context} "ALL Ts t. type_of1 (Ts, t) = fastype_of1 (Ts, t)" +\ + + +ML_command \ +val thy = \<^theory>; +check_dynamic (Context.the_local_context ()) + "ALL t u. if Pattern.matches thy (t, u) then Term.could_unify (t, u) else true" +\ + +end diff --git a/thys/SpecCheck/generators/SpecCheck_Generators.thy b/thys/SpecCheck/generators/SpecCheck_Generators.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/SpecCheck_Generators.thy @@ -0,0 +1,20 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Generators\ +theory SpecCheck_Generators +imports SpecCheck_Base +begin + +paragraph \Summary\ +text \Generators for SpecCheck take a state and return a pair consisting of a generated value and +a new state. Refer to @{file "gen_base.ML"} for the most basic combinators.\ + +ML_file \gen_types.ML\ +ML_file \gen_base.ML\ +ML_file \gen_text.ML\ +ML_file \gen_int.ML\ +ML_file \gen_real.ML\ +ML_file \gen_function.ML\ +ML_file \gen_term.ML\ +ML_file \gen.ML\ + +end diff --git a/thys/SpecCheck/generators/gen.ML b/thys/SpecCheck/generators/gen.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen.ML @@ -0,0 +1,15 @@ +(* Title: SpecCheck/generators/gen.ML + Author: Kevin Kappelmann, TU Muenchen + +Structure containing all generators. +*) +structure SpecCheck_Generator = +struct +open SpecCheck_Gen_Types +open SpecCheck_Gen_Base +open SpecCheck_Gen_Text +open SpecCheck_Gen_Real +open SpecCheck_Gen_Int +open SpecCheck_Gen_Function +open SpecCheck_Gen_Term +end diff --git a/thys/SpecCheck/generators/gen_base.ML b/thys/SpecCheck/generators/gen_base.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_base.ML @@ -0,0 +1,244 @@ +(* Title: SpecCheck/generators/gen_base.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Basic utility functions to create and combine generators. +*) + +signature SPECCHECK_GEN_BASE = +sig + include SPECCHECK_GEN_TYPES + + (*generator that always returns the passed value*) + val return : 'a -> ('a, 's) gen_state + val join : (('a, 's) gen_state, 's) gen_state -> ('a, 's) gen_state + + val zip : ('a, 's) gen_state -> ('b, 's) gen_state -> ('a * 'b, 's) gen_state + val zip3 : ('a, 's) gen_state -> ('b, 's) gen_state -> ('c, 's) gen_state -> + ('a * 'b * 'c, 's) gen_state + val zip4 : ('a, 's) gen_state -> ('b, 's) gen_state -> ('c, 's) gen_state -> ('d, 's) gen_state -> + ('a * 'b * 'c * 'd, 's) gen_state + val map : ('a -> 'b) -> ('a, 's) gen_state -> ('b, 's) gen_state + val map2 : ('a -> 'b -> 'c) -> ('a, 's) gen_state -> ('b, 's) gen_state -> ('c, 's) gen_state + val map3 : ('a -> 'b -> 'c -> 'd) -> ('a, 's) gen_state -> ('b, 's) gen_state -> + ('c, 's) gen_state -> ('d, 's) gen_state + val map4 : ('a -> 'b -> 'c -> 'd -> 'e) -> ('a, 's) gen_state -> ('b, 's) gen_state -> + ('c, 's) gen_state -> ('d, 's) gen_state -> ('e, 's) gen_state + + (*ensures that all generated values fulfill the predicate; loops if the predicate is never + satisfied by the generated values.*) + val filter : ('a -> bool) -> ('a, 's) gen_state -> ('a, 's) gen_state + val filter_bounded : int -> ('a -> bool) -> ('a, 's) gen_state -> ('a, 's) gen_state + + (*bernoulli random variable with parameter p \ [0,1]*) + val bernoulli : real -> bool gen + + (*random variable following a binomial distribution with parameters p \ [0,1] and n \ 0*) + val binom_dist : real -> int -> int gen + + (*range_int (x,y) r returns a value in [x;y]*) + val range_int : int * int -> int gen + + (*randomly generates one of the given values*) + val elements : 'a vector -> 'a gen + (*randomly uses one of the given generators*) + val one_of : 'a gen vector -> 'a gen + + (*randomly generates one of the given values*) + val elementsL : 'a list -> 'a gen + (*randomly uses one of the given generators*) + val one_ofL : 'a gen list -> 'a gen + + (*chooses one of the given generators with a weighted random distribution.*) + val one_ofW : (int * 'a gen) vector -> 'a gen + (*chooses one of the given values with a weighted random distribution.*) + val elementsW : (int * 'a) vector -> 'a gen + + (*chooses one of the given generators with a weighted random distribution.*) + val one_ofWL : (int * 'a gen) list -> 'a gen + (*chooses one of the given values with a weighted random distribution.*) + val elementsWL : (int * 'a) list -> 'a gen + + (*creates a vector of length as returned by the passed int generator*) + val vector : (int, 's) gen_state -> ('a, 's) gen_state -> ('a vector, 's) gen_state + + (*creates a list of length as returned by the passed int generator*) + val list : (int, 's) gen_state -> ('a, 's) gen_state -> ('a list, 's) gen_state + + (*generates elements until the passed (generator) predicate fails; + returns a list of all values that satisfied the predicate*) + val unfold_while : ('a -> (bool, 's) gen_state) -> ('a, 's) gen_state -> ('a list, 's) gen_state + (*generates a random permutation of the given list*) + val shuffle : 'a list -> 'a list gen + + (*generates SOME value if passed bool generator returns true and NONE otherwise*) + val option : (bool, 's) gen_state -> ('a, 's) gen_state -> ('a option, 's) gen_state + + (*seq gen init_state creates a sequence where gen takes a state and returns the next element and + an updated state. The sequence stops when the first NONE value is returned by the generator.*) + val seq : ('a option, 's * SpecCheck_Random.rand) gen_state -> 's -> ('a Seq.seq) gen + + (*creates a generator that returns all elements of the sequence in order*) + val of_seq : ('a option, 'a Seq.seq) gen_state + + val unit : (unit, 's) gen_state + val ref_gen : ('a, 's) gen_state -> ('a Unsynchronized.ref, 's) gen_state + + (*variant i g creates the ith variant of a given generator. It raises an error if i is negative.*) + val variant : (int, 'b) cogen + val cobool : (bool, 'b) cogen + val colist : ('a, 'b) cogen -> ('a list, 'b) cogen + val cooption : ('a, 'b) cogen -> ('a option, 'b) cogen + +end + +structure SpecCheck_Gen_Base : SPECCHECK_GEN_BASE = +struct + +open SpecCheck_Gen_Types + +val return = pair +fun join g = g #> (fn (g', r') => g' r') + +fun zip g1 g2 = + g1 #-> (fn x1 => g2 #-> (fn x2 => pair (x1, x2))) +fun zip3 g1 g2 g3 = + zip g1 (zip g2 g3) #-> (fn (x1, (x2, x3)) => pair (x1, x2, x3)) +fun zip4 g1 g2 g3 g4 = + zip (zip g1 g2) (zip g3 g4) #-> (fn ((x1, x2), (x3, x4)) => pair (x1, x2, x3, x4)) + +fun map f g = g #>> f +fun map2 f = map (uncurry f) oo zip +fun map3 f = map (fn (a,b,c) => f a b c) ooo zip3 +fun map4 f = map (fn (a,b,c,d) => f a b c d) oooo zip4 + +fun filter p gen r = + let fun loop (x, r) = if p x then (x, r) else loop (gen r) + in loop (gen r) end + +fun filter_bounded bound p gen r = + let fun loop 0 _ = raise Fail "Generator failed to satisfy predicate" + | loop n (x, r) = if p x then (x, r) else loop (n-1) (gen r) + in loop bound (gen r) end + +fun bernoulli p r = let val (x,r) = SpecCheck_Random.real_unit r in (x <= p, r) end + +fun binom_dist p n r = + let + fun binom_dist_unchecked _ 0 r = (0, r) + | binom_dist_unchecked p n r = + let fun int_of_bool b = if b then 1 else 0 + in + bernoulli p r + |>> int_of_bool + ||> binom_dist_unchecked p (n-1) + |> (fn (x,(acc,r)) => (acc + x, r)) + end + in if n < 0 then raise Domain else binom_dist_unchecked p n r end + +fun range_int (min, max) r = + if min > max + then raise Fail (SpecCheck_Util.spaces ["Range_Int:", string_of_int min, ">", string_of_int max]) + else + SpecCheck_Random.real_unit r + |>> (fn s => Int.min (min + Real.floor (s * Real.fromInt (max - min + 1)), max)) + +fun one_of v r = + let val (i, r) = range_int (0, Vector.length v - 1) r + in Vector.sub (v, i) r end + +local +datatype ('a,'b) either = Left of 'a | Right of 'b +in +fun one_ofW (v : (int * 'a gen) vector) r = + let + val weight_total = Vector.foldl (fn ((freq,_),acc) => freq + acc) 0 v + fun selectGen _ (_, Right gen) = Right gen + | selectGen rand ((weight, gen), Left acc) = + let val acc = acc + weight + in if acc < rand + then Left acc + else Right gen + end + val (threshhold, r) = range_int (1, weight_total) r + val gen = case Vector.foldl (selectGen threshhold) (Left 0) v of + Right gen => gen + | _ => raise Fail "Error in one_ofW - did you pass an empty vector or invalid frequencies?" + in gen r end +end + +fun elements v = one_of (Vector.map return v) +fun elementsW v = one_ofW (Vector.map (fn p => p ||> return) v) + +fun one_ofL l = one_of (Vector.fromList l) +fun one_ofWL l = one_ofW (Vector.fromList l) +fun elementsL l = elements (Vector.fromList l) +fun elementsWL l = elementsW (Vector.fromList l) + +fun list length_g g r = + let val (l, r) = length_g r + in fold_map (K g) (map_range I l) r end + +fun unfold_while bool_g_of_elem g r = + let + fun unfold_while_accum (xs, r) = + let + val (x, r) = g r + val (continue, r) = bool_g_of_elem x r + in + if continue + then unfold_while_accum (x::xs, r) + else (xs, r) + end + in unfold_while_accum ([], r) |>> rev end + +fun shuffle xs r = + let + val (ns, r) = list (return (length xs)) SpecCheck_Random.real_unit r + val real_ord = make_ord (op <=) + val xs = ListPair.zip (ns, xs) |> sort (real_ord o apply2 fst) |> List.map snd + in (xs, r) end + +fun vector length_g g = list length_g g #>> Vector.fromList + +fun option bool_g g r = + case bool_g r of + (false, r) => (NONE, r) + | (true, r) => map SOME g r + +fun seq gen xq r = + let + val (r1, r2) = SpecCheck_Random.split r + fun gen_next p () = case gen p of + (NONE, _) => NONE + | (SOME v, p) => SOME (v, Seq.make (gen_next p)) + in (Seq.make (gen_next (xq, r1)), r2) end + +fun of_seq xq = case Seq.pull xq of + SOME (x, xq) => (SOME x, xq) + | NONE => (NONE, Seq.empty) + +fun unit s = return () s + +fun ref_gen gen r = let val (value, r) = gen r + in (Unsynchronized.ref value, r) end + +fun variant i g r = + if i < 0 then raise Subscript + else + let + fun nth i r = + let val (r1, r2) = SpecCheck_Random.split r + in if i = 0 then r1 else nth (i-1) r2 end + in g (nth i r) end + +fun cobool false = variant 0 + | cobool true = variant 1 + +fun colist _ [] = variant 0 + | colist co (x::xs) = colist co xs o co x o variant 1 + +fun cooption _ NONE = variant 0 + | cooption co (SOME x) = co x o variant 1 + +end diff --git a/thys/SpecCheck/generators/gen_function.ML b/thys/SpecCheck/generators/gen_function.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_function.ML @@ -0,0 +1,41 @@ +(* Title: SpecCheck/generators/gen_function.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Random generators for functions. +*) + +signature SPECCHECK_GEN_FUNCTION = sig + val function : ('a, 'b) SpecCheck_Gen_Types.cogen -> 'b SpecCheck_Gen_Types.gen -> + ('a -> 'b) SpecCheck_Gen_Types.gen + val function' : 'b SpecCheck_Gen_Types.gen -> (''a -> 'b) SpecCheck_Gen_Types.gen +end + +structure SpecCheck_Gen_Function : SPECCHECK_GEN_FUNCTION = +struct + +fun function cogen gen r = + let + val (r1, r2) = SpecCheck_Random.split r + fun g x = fst (cogen x gen r1) + in (g, r2) end + +fun function' gen r = + let + val (internal, external) = SpecCheck_Random.split r + val seed = Unsynchronized.ref internal + val table = Unsynchronized.ref [] + fun new_entry k = + let + val (new_val, new_seed) = gen (!seed) + val _ = seed := new_seed + val _ = table := AList.update (op =) (k, new_val) (!table) + in new_val end + in + (fn v1 => + case AList.lookup (op =) (!table) v1 of + NONE => new_entry v1 + | SOME v2 => v2, external) + end + +end diff --git a/thys/SpecCheck/generators/gen_int.ML b/thys/SpecCheck/generators/gen_int.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_int.ML @@ -0,0 +1,38 @@ +(* Title: SpecCheck/generators/gen_int.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Random generators for ints. +*) + +signature SPECCHECK_GEN_INT = sig + + (*pos m generates an integer in [1, m]*) + val pos : int -> int SpecCheck_Gen_Types.gen + (*neg m generates an integer in [m, 1]*) + val neg : int -> int SpecCheck_Gen_Types.gen + (*nonneg m generates an integer in [0, m]*) + val nonneg : int -> int SpecCheck_Gen_Types.gen + (*nonpos m generates an integer in [m, 0]*) + val nonpos : int -> int SpecCheck_Gen_Types.gen + + val coint : (int, 'b) SpecCheck_Gen_Types.cogen + +end + +structure SpecCheck_Gen_Int : SPECCHECK_GEN_INT = +struct + +open SpecCheck_Gen_Base + +fun pos m = range_int (1, m) +fun neg m = range_int (m, ~1) +fun nonneg m = range_int (0, m) +fun nonpos m = range_int (m, 0) + +fun coint n = + if n = 0 then variant 0 + else if n < 0 then coint (~n) o variant 1 + else coint (n div 2) o variant 2 + +end diff --git a/thys/SpecCheck/generators/gen_real.ML b/thys/SpecCheck/generators/gen_real.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_real.ML @@ -0,0 +1,65 @@ +(* Title: SpecCheck/generators/gen_real.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Random generators for reals. +*) + +signature SPECCHECK_GEN_REAL = sig + + (*range_real (x,y) r returns a value in [x, y]*) + val range_real : real * real -> real SpecCheck_Gen_Types.gen + + val real : real SpecCheck_Gen_Types.gen + + val real_pos : real SpecCheck_Gen_Types.gen + val real_neg : real SpecCheck_Gen_Types.gen + + val real_nonneg : real SpecCheck_Gen_Types.gen + val real_nonpos : real SpecCheck_Gen_Types.gen + + val real_finite : real SpecCheck_Gen_Types.gen +end + +structure SpecCheck_Gen_Real : SPECCHECK_GEN_REAL = +struct + +open SpecCheck_Gen_Base +open SpecCheck_Gen_Text + +fun range_real (min, max) r = + if min > max + then + raise Fail (SpecCheck_Util.spaces ["Range_Real:", string_of_real min, ">", string_of_real max]) + else SpecCheck_Random.real_unit r |>> (fn s => min + (s * max - s * min)) + +val digits = string (range_int (1, Real.precision)) (range_char (#"0", #"9")) + +val {exp=minExp,...} = Real.toManExp Real.minPos +val {exp=maxExp,...} = Real.toManExp Real.posInf + +val ratio = 99 + +fun mk r = + let + val (a, r) = digits r + val (b, r) = digits r + val (e, r) = range_int (minExp div 4, maxExp div 4) r + val x = String.concat [a, ".", b, "E", Int.toString e] + in + (the (Real.fromString x), r) + end + +val real_pos = one_ofWL ((ratio, mk) :: + List.map ((pair 1) o return) [Real.posInf, Real.maxFinite, Real.minPos, Real.minNormalPos]) + +val real_neg = map Real.~ real_pos + +val real_nonneg = one_ofWL [(1, return 0.0), (ratio, real_pos)] +val real_nonpos = one_ofWL [(1, return 0.0), (ratio, real_neg)] + +val real = one_ofL [real_nonneg, real_nonpos] + +val real_finite = filter Real.isFinite real + +end diff --git a/thys/SpecCheck/generators/gen_term.ML b/thys/SpecCheck/generators/gen_term.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_term.ML @@ -0,0 +1,269 @@ +(* Title: SpecCheck/generators/gen_term.ML + Author: Sebastian Willenbrink and Kevin Kappelmann TU Muenchen + +Generators for terms and types. +*) +signature SPECCHECK_GEN_TERM = sig + (* sort generators *) + + (*first parameter determines the number of classes to pick*) + val sort : (int, 's) SpecCheck_Gen_Types.gen_state -> (class, 's) SpecCheck_Gen_Types.gen_state + -> (sort, 's) SpecCheck_Gen_Types.gen_state + val dummyS : (sort, 's) SpecCheck_Gen_Types.gen_state + + (* name generators *) + (*parameters: a base name and a generator for the number of variants to choose from based on then + passed base name*) + val basic_name : string -> int SpecCheck_Gen_Types.gen -> string SpecCheck_Gen_Types.gen + + val indexname : (string, 's) SpecCheck_Gen_Types.gen_state -> + (int, 's) SpecCheck_Gen_Types.gen_state -> (indexname, 's) SpecCheck_Gen_Types.gen_state + + (*a variant with base name "k"*) + val type_name : int SpecCheck_Gen_Types.gen -> string SpecCheck_Gen_Types.gen + + (*creates a free type variable name from a passed basic name generator*) + val tfree_name : (string, 's) SpecCheck_Gen_Types.gen_state -> + (string, 's) SpecCheck_Gen_Types.gen_state + (*chooses a variant with base name "'a"*) + val tfree_name' : int SpecCheck_Gen_Types.gen -> string SpecCheck_Gen_Types.gen + + (*creates a type variable name from a passed basic name (e.g. "a") generator*) + val tvar_name : (indexname, 's) SpecCheck_Gen_Types.gen_state -> + (indexname, 's) SpecCheck_Gen_Types.gen_state + (*chooses a variant with base name "'a"*) + val tvar_name' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + indexname SpecCheck_Gen_Types.gen + + (*chooses a variant with base name "c"*) + val const_name : int SpecCheck_Gen_Types.gen -> string SpecCheck_Gen_Types.gen + (*chooses a variant with base name "f"*) + val free_name : int SpecCheck_Gen_Types.gen -> string SpecCheck_Gen_Types.gen + (*chooses a variant with base name "v*) + val var_name : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + indexname SpecCheck_Gen_Types.gen + + (* typ generators *) + + val tfree : (string, 's) SpecCheck_Gen_Types.gen_state -> + (sort, 's) SpecCheck_Gen_Types.gen_state -> (typ, 's) SpecCheck_Gen_Types.gen_state + (*uses tfree_name' and dummyS to create a free type variable*) + val tfree' : int SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen + + val tvar : (indexname, 's) SpecCheck_Gen_Types.gen_state -> + (sort, 's) SpecCheck_Gen_Types.gen_state -> (typ, 's) SpecCheck_Gen_Types.gen_state + (*uses tvar' and dummyS to create a type variable*) + val tvar' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + typ SpecCheck_Gen_Types.gen + + (*atyp tfree_gen tvar_gen (weight_tfree, weight_tvar)*) + val atyp : typ SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen -> (int * int) -> + typ SpecCheck_Gen_Types.gen + (*uses tfree' and tvar' to create an atomic type*) + val atyp' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> (int * int) -> + typ SpecCheck_Gen_Types.gen + + (*type' type_name_gen arity_gen tfree_gen tvar_gen (weight_type, weight_tfree, weight_tvar)*) + val type' : string SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + typ SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen -> + (int * int * int) -> typ SpecCheck_Gen_Types.gen + (*uses type_name to generate a type*) + val type'' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + typ SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen -> (int * int * int) -> + typ SpecCheck_Gen_Types.gen + + (*typ type_gen tfree_gen tvar_gen (wtype, wtfree, wtvar)*) + val typ : typ SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen -> + typ SpecCheck_Gen_Types.gen -> (int * int * int) -> typ SpecCheck_Gen_Types.gen + (*uses type'' for its type generator*) + val typ' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + typ SpecCheck_Gen_Types.gen -> typ SpecCheck_Gen_Types.gen -> (int * int * int) -> + typ SpecCheck_Gen_Types.gen + (*uses typ' with tfree' and tvar' parameters*) + val typ'' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + int SpecCheck_Gen_Types.gen -> (int * int * int) -> typ SpecCheck_Gen_Types.gen + + val dummyT : (typ, 's) SpecCheck_Gen_Types.gen_state + + (* term generators *) + + val const : (string, 's) SpecCheck_Gen_Types.gen_state -> + (typ, 's) SpecCheck_Gen_Types.gen_state -> (term, 's) SpecCheck_Gen_Types.gen_state + (*uses const_name and dummyT to create a constant*) + val const' : int SpecCheck_Gen_Types.gen -> term SpecCheck_Gen_Types.gen + + val free : (string, 's) SpecCheck_Gen_Types.gen_state -> + (typ, 's) SpecCheck_Gen_Types.gen_state -> (term, 's) SpecCheck_Gen_Types.gen_state + (*uses free_name and dummyT to create a free variable*) + val free' : int SpecCheck_Gen_Types.gen -> term SpecCheck_Gen_Types.gen + + val var : (indexname, 's) SpecCheck_Gen_Types.gen_state -> + (typ, 's) SpecCheck_Gen_Types.gen_state -> (term, 's) SpecCheck_Gen_Types.gen_state + (*uses var_name and dummyT to create a variable*) + val var' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + term SpecCheck_Gen_Types.gen + + val bound : (int, 's) SpecCheck_Gen_Types.gen_state -> (term, 's) SpecCheck_Gen_Types.gen_state + + (*aterm const_gen free_gen var_gen bound_gen + (weight_const, weight_free, weight_var, weight_bound*) + val aterm : term SpecCheck_Gen_Types.gen -> term SpecCheck_Gen_Types.gen -> + term SpecCheck_Gen_Types.gen -> term SpecCheck_Gen_Types.gen -> (int * int * int * int) -> + term SpecCheck_Gen_Types.gen + (*uses const', free', and var' to create an atomic term*) + val aterm' : int SpecCheck_Gen_Types.gen -> int SpecCheck_Gen_Types.gen -> + (int * int * int * int) -> term SpecCheck_Gen_Types.gen + + (*term_tree f init_state - where "f height index state" returns "((term, num_args), new_state)" - + generates a term by applying f to every node and expanding that node depending + on num_args returned by f. + Traversal order: function \ first argument \ ... \ last argument + The tree is returned in its applicative term form: (...((root $ child1) $ child2) .. $ childn). + + Arguments of f: + - height describes the distance from the root (starts at 0) + - index describes the global index in that tree layer, left to right (0 \ index < width) + - state is passed along according to above traversal order + + Return value of f: + - term is the term whose arguments will be generated next. + - num_args specifies how many arguments should be passed to the term. + - new_state is passed along according to the traversal above.*) + val term_tree : (int -> int -> (term * int, 's) SpecCheck_Gen_Types.gen_state) -> + (term, 's) SpecCheck_Gen_Types.gen_state + + (*In contrast to term_tree, f now takes a (term, index_of_argument) list which specifies the path + from the root to the current node.*) + val term_tree_path : ((term * int) list -> (term * int, 's) SpecCheck_Gen_Types.gen_state) -> + (term, 's) SpecCheck_Gen_Types.gen_state + +end + +structure SpecCheck_Gen_Term : SPECCHECK_GEN_TERM = +struct + +structure Gen = SpecCheck_Gen_Base + +fun sort size_gen = Gen.list size_gen +fun dummyS s = Gen.return Term.dummyS s + +fun basic_name name num_variants_gen = + num_variants_gen + #>> (fn i => name ^ "_" ^ string_of_int i) + +fun indexname basic_name_gen = Gen.zip basic_name_gen + +fun type_name num_variants_gen = basic_name "k" num_variants_gen + +fun tfree_name basic_name_gen = Gen.map (curry op^"'") basic_name_gen +fun tfree_name' num_variants_gen = tfree_name (basic_name "a" num_variants_gen) + +fun tvar_name indexname_gen = Gen.map (curry op^"'" |> apfst) indexname_gen +fun tvar_name' num_variants_gen = + tvar_name o indexname (basic_name "a" num_variants_gen) + +fun const_name num_variants_gen = basic_name "c" num_variants_gen +fun free_name num_variants_gen = basic_name "v" num_variants_gen +fun var_name num_variants_gen = indexname (free_name num_variants_gen) + +(* types *) + +fun tfree name_gen = Gen.map TFree o Gen.zip name_gen +fun tfree' num_variants_gen = + tfree_name' num_variants_gen + |> (fn name_gen => tfree name_gen dummyS) + +fun tvar idx_gen = Gen.map TVar o Gen.zip idx_gen +fun tvar' num_variants_gen = + tvar_name' num_variants_gen + #> (fn name_gen => tvar name_gen dummyS) + +fun atyp tfree_gen tvar_gen (wtfree, wtvar) = + Gen.one_ofWL [(wtfree, tfree_gen), (wtvar, tvar_gen)] +fun atyp' num_variants_gen = atyp (tfree' num_variants_gen) o tvar' num_variants_gen + +fun type' type_name_gen arity_gen tfree_gen tvar_gen (weights as (wtype, wtfree, wtvar)) = + (*eta-abstract to avoid strict evaluation, causing an infinite loop*) + [(wtype, fn r => type' type_name_gen arity_gen tfree_gen tvar_gen weights r), + (wtfree, fn r => tfree_gen r), (wtvar, fn r => tvar_gen r)] + |> Gen.one_ofWL + |> Gen.list arity_gen + |> Gen.zip type_name_gen + |> Gen.map Type + +fun type'' num_variants_gen = type_name num_variants_gen |> type' + +fun typ type_gen tfree_gen tvar_gen (wtype, wtfree, wtvar) = + Gen.one_ofWL [(wtype, type_gen), (wtfree, tfree_gen), (wtvar, tvar_gen)] +fun typ' num_variants_gen arity_gen tfree_gen tvar_gen weights = + typ (type'' num_variants_gen arity_gen tfree_gen tvar_gen weights) tfree_gen tvar_gen weights +fun typ'' num_variants_gen arity_gen = + typ' num_variants_gen arity_gen (tfree' num_variants_gen) o tvar' num_variants_gen + +fun dummyT s = Gen.return Term.dummyT s + +(* terms *) + +fun const name_gen = Gen.map Const o Gen.zip name_gen +fun const' num_variants_gen = + const_name num_variants_gen + |> (fn name_gen => const name_gen dummyT) + +fun free name_gen = Gen.map Free o Gen.zip name_gen +fun free' num_variants_gen = + free_name num_variants_gen + |> (fn name_gen => free name_gen dummyT) + +fun var idx_gen = Gen.map Var o Gen.zip idx_gen +fun var' num_variants_gen = + var_name num_variants_gen + #> (fn name_gen => var name_gen dummyT) + +fun bound int_gen = Gen.map Bound int_gen + +fun aterm const_gen free_gen var_gen bound_gen (wconst, wfree, wvar, wbound) = + Gen.one_ofWL [(wconst, const_gen), (wfree, free_gen), (wvar, var_gen), (wbound, bound_gen)] +fun aterm' num_variants_gen index_gen = + aterm (const' num_variants_gen) (free' num_variants_gen) (var' num_variants_gen index_gen) + (bound num_variants_gen) + +(*nth_map has no default*) +fun nth_map_default 0 f _ (x::xs) = f x :: xs + | nth_map_default 0 f d [] = [f d] + | nth_map_default n f d [] = replicate (n-1) d @ [f d] + | nth_map_default n f d (x::xs) = x :: nth_map_default (n-1) f d xs + +fun term_tree term_gen state = + let + fun nth_incr n = nth_map_default n (curry op+ 1) (~1) + fun build_tree indices height state = + let + (*indices stores the number of nodes visited so far at each height*) + val indices = nth_incr height indices + val index = nth indices height + (*generate the term for the current node*) + val ((term, num_args), state) = term_gen height index state + fun build_child (children, indices, state) = + build_tree indices (height + 1) state + |> (fn (child, indices, state) => (child :: children, indices, state)) + (*generate the subtrees for each argument*) + val (children, indices, state) = fold (K build_child) (1 upto num_args) ([], indices, state) + in (Term.list_comb (term, (rev children)), indices, state) end + in + build_tree [] 0 state + |> (fn (term, _, state) => (term, state)) + end + +fun term_tree_path f init_state = + let + fun build_tree path state = + let + val ((term, num_args), state) = f path state + fun build_children i (args, state) = + build_tree ((term, i) :: path) state + |>> (fn x => x :: args) + val (children, state) = fold build_children (0 upto num_args-1) ([], state) + in (Term.list_comb (term, (rev children)), state) end + in build_tree [] init_state end + +end diff --git a/thys/SpecCheck/generators/gen_text.ML b/thys/SpecCheck/generators/gen_text.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_text.ML @@ -0,0 +1,70 @@ +(* Title: SpecCheck/generators/gen_text.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Random generators for chars and strings. +*) + +signature SPECCHECK_GEN_TEXT = sig + + val range_char : char * char -> char SpecCheck_Gen_Types.gen + val char : char SpecCheck_Gen_Types.gen + + val char_of : string -> char SpecCheck_Gen_Types.gen + + val string : int SpecCheck_Gen_Types.gen -> char SpecCheck_Gen_Types.gen -> + string SpecCheck_Gen_Types.gen + + val substring : string SpecCheck_Gen_Types.gen -> substring SpecCheck_Gen_Types.gen + + val cochar : (char, 'b) SpecCheck_Gen_Types.cogen + val costring : (string, 'b) SpecCheck_Gen_Types.cogen + val cosubstring : (substring, 'b) SpecCheck_Gen_Types.cogen + + val digit : char SpecCheck_Gen_Types.gen + val lowercase_letter : char SpecCheck_Gen_Types.gen + val uppercase_letter : char SpecCheck_Gen_Types.gen + val letter : char SpecCheck_Gen_Types.gen +end + +structure SpecCheck_Gen_Text : SPECCHECK_GEN_TEXT = +struct + +open SpecCheck_Gen_Base + +type char = Char.char +type string = String.string +type substring = Substring.substring + +fun range_char (lo, hi) = map Char.chr (range_int (Char.ord lo, Char.ord hi)) +val char = range_char (Char.minChar, Char.maxChar) + +fun char_of s = + one_of (Vector.tabulate (String.size s, fn i => return (String.sub (s, i)))) + +fun string length_g g = list length_g g #>> CharVector.fromList + +fun substring gen r = + let + val (s, r) = gen r + val (i, r) = range_int (0, String.size s) r + val (j, r) = range_int (0, String.size s - i) r + in + (Substring.substring (s, i, j), r) + end + +fun cochar c = + if Char.ord c = 0 then variant 0 + else cochar (Char.chr (Char.ord c div 2)) o variant 1 + +fun cosubstring s = colist cochar (Substring.explode s) + +fun costring s = cosubstring (Substring.full s) + +val digit = range_char (#"0", #"9") + +val lowercase_letter = range_char (#"a", #"z") +val uppercase_letter = range_char (#"A", #"Z") +val letter = one_ofL [lowercase_letter, uppercase_letter] + +end diff --git a/thys/SpecCheck/generators/gen_types.ML b/thys/SpecCheck/generators/gen_types.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/generators/gen_types.ML @@ -0,0 +1,34 @@ +(* Title: SpecCheck/generators/gen_types.ML + Author: Kevin Kappelmann + +Shared type definitions for SpecCheck generators. +*) + +signature SPECCHECK_GEN_TYPES = +sig + + (*consumes a state and returns a new state along with a generated value*) + type ('a, 's) gen_state = 's -> 'a * 's + (*consumes a random seed and returns an unused one along with a generated value*) + type 'a gen = ('a, SpecCheck_Random.rand) gen_state + + (*a cogenerator produces new generators depending on an input element and an existing generator.*) + type ('a, 'b, 's) cogen_state = 'a -> ('b, 's) gen_state -> ('b, 's) gen_state + + (*a cogenerator produces new generators depending on an input element and an existing generator.*) + type ('a, 'b) cogen = ('a, 'b, SpecCheck_Random.rand) cogen_state + +end + +structure SpecCheck_Gen_Types : SPECCHECK_GEN_TYPES = +struct + +type ('a, 's) gen_state = 's -> 'a * 's + +type 'a gen = ('a, SpecCheck_Random.rand) gen_state + +type ('a, 'b, 's) cogen_state = 'a -> ('b, 's) gen_state -> ('b, 's) gen_state + +type ('a, 'b) cogen = ('a, 'b, SpecCheck_Random.rand) cogen_state + +end diff --git a/thys/SpecCheck/lecker.ML b/thys/SpecCheck/lecker.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/lecker.ML @@ -0,0 +1,23 @@ +(* Title: SpecCheck/lecker.ML + Author: Kevin Kappelmann + +Testing framework that lets you combine SpecCheck tests into test suites. + +TODO: this file can be largely extended to become a pendant of Haskell's tasty: +https://hackage.haskell.org/package/tasty +*) + +signature LECKER = +sig + (*the first parameter to test_group will usually be a context*) + val test_group : 'a -> 's -> ('a -> 's -> 's) list -> 's +end + +structure Lecker : LECKER = +struct + +fun test_group _ s [] = s + | test_group fixed_param s (t::ts) = + fold (fn t => (Pretty.writeln (Pretty.para ""); t fixed_param)) ts (t fixed_param s) + +end diff --git a/thys/SpecCheck/output_styles/SpecCheck_Output_Style.thy b/thys/SpecCheck/output_styles/SpecCheck_Output_Style.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/output_styles/SpecCheck_Output_Style.thy @@ -0,0 +1,20 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Output Styles\ +theory SpecCheck_Output_Style +imports + SpecCheck_Base + SpecCheck_Show +begin + + +paragraph \Summary\ +text \Output styles for SpecCheck take the result of a test run and process it accordingly, +e.g. by printing it or storing it to a file.\ + + +ML_file \output_style_types.ML\ +ML_file \output_style_perl.ML\ +ML_file \output_style_custom.ML\ +ML_file \output_style.ML\ + +end diff --git a/thys/SpecCheck/output_styles/output_style.ML b/thys/SpecCheck/output_styles/output_style.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/output_styles/output_style.ML @@ -0,0 +1,20 @@ +(* Title: SpecCheck/output_styles/output_style.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Output styles for presenting SpecCheck results. +*) + +signature SPECCHECK_DEFAULT_OUTPUT_STYLE = +sig + include SPECCHECK_OUTPUT_STYLE_TYPES + val default : 'a output_style +end + +structure SpecCheck_Default_Output_Style : SPECCHECK_DEFAULT_OUTPUT_STYLE = +struct + +open SpecCheck_Output_Style_Types +val default = SpecCheck_Output_Style_Custom.style + +end diff --git a/thys/SpecCheck/output_styles/output_style_custom.ML b/thys/SpecCheck/output_styles/output_style_custom.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/output_styles/output_style_custom.ML @@ -0,0 +1,83 @@ +(* Title: SpecCheck/output_style/output_style_custom.ML + Author: Lukas Bulwahn, Nicolai Schaffroth, Kevin Kappelmann TU Muenchen + Author: Christopher League + +Custom-made, QuickCheck-inspired output style for SpecCheck. +*) + +structure SpecCheck_Output_Style_Custom : SPECCHECK_OUTPUT_STYLE = +struct + +open SpecCheck_Base + +fun print_success stats = + let + val num_success_tests = string_of_int (#num_success_tests stats) + val num_discarded_tests = #num_discarded_tests stats + val discarded_str = + if num_discarded_tests = 0 + then "." + else SpecCheck_Util.spaces [";", string_of_int num_discarded_tests, "discarded."] + in + implode ["OK, passed ", num_success_tests, " tests", discarded_str] + |> writeln + end + +fun print_gave_up stats = + let + val num_success_tests = string_of_int (#num_success_tests stats) + val num_discarded_tests = string_of_int (#num_discarded_tests stats) + in + SpecCheck_Util.spaces ["Gave up! Passed only", num_success_tests, "test(s);", num_discarded_tests, + "discarded test(s)."] + |> warning + end + +fun print_failure_data ctxt show_opt failure_data = + case #the_exception failure_data of + SOME exn => + cat_lines ["Exception during test run:", exnMessage exn] + |> warning + | NONE => case show_opt of + NONE => () + | SOME show => + let + val sort_counterexamples = Config.get ctxt SpecCheck_Configuration.sort_counterexamples + val maybe_sort = if sort_counterexamples then sort (int_ord o apply2 size) else I + val counterexamples = + #counterexamples failure_data + |> map (Pretty.string_of o show) + |> maybe_sort + in fold (fn x => fn _ => warning x) counterexamples () end + +fun print_failure ctxt show_opt (stats, failure_data) = + ((SpecCheck_Util.spaces ["Failed! Falsified (after", string_of_int (num_tests stats), "test(s) and ", + string_of_int (num_shrinks stats), "shrink(s)):"] |> warning); + print_failure_data ctxt show_opt failure_data) + +fun print_stats ctxt stats total_time = + let + val show_stats = Config.get ctxt SpecCheck_Configuration.show_stats + (*the time spent in the test function in relation to the total time spent; + the latter includes generating test cases and overhead from the framework*) + fun show_time {elapsed, ...} = + implode ["Time: ", Time.toString elapsed, "s (run) / ", Time.toString (#elapsed total_time), + "s (total)"] + in + if not show_stats + then () + else writeln (show_time (#timing stats)) + end + +fun style show_opt ctxt name total_time result = + let val stats = stats_of_result result + in + writeln (SpecCheck_Util.spaces ["Testing:", name]); + (case result of + Success _ => print_success stats + | Gave_Up _ => print_gave_up stats + | Failure data => print_failure ctxt show_opt data); + print_stats ctxt stats total_time + end + +end diff --git a/thys/SpecCheck/output_styles/output_style_perl.ML b/thys/SpecCheck/output_styles/output_style_perl.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/output_styles/output_style_perl.ML @@ -0,0 +1,54 @@ +(* Title: SpecCheck/output_styles/output_style_perl.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Perl output styles for SpecCheck. +*) + +structure SpecCheck_Output_Style_Perl : SPECCHECK_OUTPUT_STYLE = +struct + +open SpecCheck_Configuration +open SpecCheck_Base + +fun style show_opt ctxt name timing result = + let + val sort_counterexamples = Config.get ctxt sort_counterexamples + val maybe_sort = if sort_counterexamples then sort (int_ord o apply2 size) else I + + val stats = stats_of_result result + val num_failed_tests = #num_failed_tests stats + + fun code (Success _) = "ok" + | code (Gave_Up _) = "Gave up!" + | code (Failure _) = "FAILED" + + fun ratio stats = + let + val num_success_tests = #num_success_tests stats + in + if num_failed_tests = 0 + then implode ["(", string_of_int num_success_tests, " passed)"] + else implode ["(", string_of_int num_success_tests, "/", + string_of_int (num_success_tests + num_failed_tests), " passed)"] + end + + val result_string = name ^ ".\n" ^ code result ^ " " ^ ratio stats + + fun show_counterexamples counterexamples = + case show_opt of + SOME show => + (case maybe_sort (map (Pretty.string_of o show) counterexamples) of + [] => () + | es => (warning "Counterexamples:"; fold (fn x => fn _ => warning x) es ())) + | NONE => () + + in + case result of + Success _ => writeln result_string + | Gave_Up _ => warning result_string + | Failure (_, failure_data) => + (warning result_string; show_counterexamples (#counterexamples failure_data)) + end + +end diff --git a/thys/SpecCheck/output_styles/output_style_types.ML b/thys/SpecCheck/output_styles/output_style_types.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/output_styles/output_style_types.ML @@ -0,0 +1,24 @@ +(* Title: SpecCheck/output_styles/output_style_types.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +Shared types for SpecCheck output styles. +*) +signature SPECCHECK_OUTPUT_STYLE_TYPES = +sig + type 'a output_style = 'a SpecCheck_Show.show option -> Proof.context -> string -> + Timing.timing -> 'a SpecCheck_Base.result -> unit +end + +structure SpecCheck_Output_Style_Types : SPECCHECK_OUTPUT_STYLE_TYPES = +struct + +type 'a output_style = 'a SpecCheck_Show.show option -> Proof.context -> string -> Timing.timing -> + 'a SpecCheck_Base.result -> unit + +end + +signature SPECCHECK_OUTPUT_STYLE = +sig + val style : 'a SpecCheck_Output_Style_Types.output_style +end diff --git a/thys/SpecCheck/property.ML b/thys/SpecCheck/property.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/property.ML @@ -0,0 +1,46 @@ +(* Title: SpecCheck/property.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +The base module of testable properties. +A property is the type of values that SpecCheck knows how to test. +Properties not only test whether a given predicate holds, but, for example, can also have +preconditions. +*) + +signature SPECCHECK_PROPERTY = +sig + + type 'a pred = 'a -> bool + (*the type of values testable by SpecCheck*) + type 'a prop + (*transforms a predicate into a testable property*) + val prop : 'a pred -> 'a prop + (*implication for properties: if the first argument evaluates to false, the test case is + discarded*) + val implies : 'a pred -> 'a prop -> 'a prop + (*convenient notation for `implies` working on predicates*) + val ==> : 'a pred * 'a pred -> 'a prop + +end + +structure SpecCheck_Property : SPECCHECK_PROPERTY = +struct + +type 'a pred = 'a -> bool +type 'a prop = 'a -> SpecCheck_Base.result_single + +fun apply f x = SpecCheck_Base.Result (f x) + (*testcode may throw arbitrary exceptions; interrupts must not be caught!*) + handle exn => if Exn.is_interrupt exn then Exn.reraise exn else SpecCheck_Base.Exception exn + +fun prop f x = apply f x + +fun implies cond prop x = + if cond x + then prop x + else SpecCheck_Base.Discard + +fun ==> (p1, p2) = implies p1 (prop p2) + +end diff --git a/thys/SpecCheck/random.ML b/thys/SpecCheck/random.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/random.ML @@ -0,0 +1,61 @@ +(* Title: SpecCheck/random.ML + Author: Lukas Bulwahn and Nicolai Schaffroth, TU Muenchen + Author: Christopher League + +A Lehmer random number generator: +https://en.wikipedia.org/wiki/Lehmer_random_number_generator +We use int to avoid any float imprecision problems (and the seed is an int anyway). +The parameters "a" and "m" are selected according to the recommendation in above article; +they are an an improved version of the so-called "minimal standard" (MINSTD) generator. + +This file only contains those functions that rely on the internal integer representation of rand. +*) + +signature SPECCHECK_RANDOM = +sig + type rand + (*creates a new random seed*) + val new : unit -> rand + (*creates a new random seed from a given one*) + val next : rand -> rand + (*use this function for reproducible randomness; inputs \ 0 are mapped to 1*) + val deterministic_seed : int -> rand + + (*returns a real in the unit interval [0;1]; theoretically, with 2^31-2 equidistant discrete + values*) + val real_unit : rand -> real * rand + + (*splits a seed into two new independent seeds*) + val split : rand -> rand * rand +end + +structure SpecCheck_Random : SPECCHECK_RANDOM = +struct + +type rand = int + +val a = 48271 +val m = 2147483647 (* 2^31 - 1 *) + +fun next seed = (seed * a) mod m + +(*TODO: Time is not sufficiently random when polled rapidly!*) +fun new () = + Time.now () + |> Time.toMicroseconds + |> (fn x => Int.max (1, x mod m)) (*The seed must be within [1;m)*) + |> next + +fun deterministic_seed r = Int.max (1, r mod m) + +fun real_unit r = ((Real.fromInt (r - 1)) / (Real.fromInt (m - 2)), next r) + +(*TODO: In theory, the current implementation could return two seeds directly adjacent in the +sequence of the pseudorandom number generator. Practically, however, it should be good enough.*) +fun split r = + let + val r0 = next r + val r1 = r - r0 + in (next r0, next r1) end + +end diff --git a/thys/SpecCheck/show/SpecCheck_Show.thy b/thys/SpecCheck/show/SpecCheck_Show.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/show/SpecCheck_Show.thy @@ -0,0 +1,15 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Show\ +theory SpecCheck_Show +imports Pure +begin + +paragraph \Summary\ +text \Show functions (pretty-printers) for SpecCheck take a value and return a @{ML_type Pretty.T} +representation of the value. Refer to @{file "show_base.ML"} for some basic examples.\ + +ML_file \show_types.ML\ +ML_file \show_base.ML\ +ML_file \show.ML\ + +end diff --git a/thys/SpecCheck/show/show.ML b/thys/SpecCheck/show/show.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/show/show.ML @@ -0,0 +1,9 @@ +(* Title: SpecCheck/show/show.ML + Author: Kevin Kappelmann, TU Muenchen + +Structure containing all show functions. +*) +structure SpecCheck_Show = +struct +open SpecCheck_Show_Base +end diff --git a/thys/SpecCheck/show/show_base.ML b/thys/SpecCheck/show/show_base.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/show/show_base.ML @@ -0,0 +1,47 @@ +(* Title: SpecCheck/show/show_base.ML + Author: Kevin Kappelmann + +Basic utility functions to create and combine show functions. +*) + +signature SPECCHECK_SHOW_BASE = +sig + include SPECCHECK_SHOW_TYPES + + val none : 'a show + val char : char show + val string : string show + val int : int show + val real : real show + val bool : bool show + val list : 'a show -> ('a list) show + val option : 'a show -> ('a option) show + + val zip : 'a show -> 'b show -> ('a * 'b) show + val zip3 : 'a show -> 'b show -> 'c show -> ('a * 'b * 'c) show + val zip4 : 'a show -> 'b show -> 'c show -> 'd show -> ('a * 'b * 'c * 'd) show + +end + +structure SpecCheck_Show_Base : SPECCHECK_SHOW_BASE = +struct + +open SpecCheck_Show_Types + +fun none _ = Pretty.str "" +val char = Pretty.enclose "'" "'" o single o Pretty.str o Char.toString +val string = Pretty.quote o Pretty.str +val int = Pretty.str o string_of_int +val real = Pretty.str o string_of_real +fun bool b = Pretty.str (if b then "true" else "false") +fun list show = Pretty.list "[" "]" o map show +fun option _ NONE = Pretty.str "NONE" + | option show (SOME v) = Pretty.block [Pretty.str "SOME ", show v] + +fun pretty_tuple ps = ps |> Pretty.commas |> Pretty.enclose "(" ")" + +fun zip showA showB (a, b) = pretty_tuple [showA a, showB b] +fun zip3 showA showB showC (a, b, c) = pretty_tuple [showA a, showB b, showC c] +fun zip4 showA showB showC showD (a, b, c, d) = pretty_tuple [showA a, showB b, showC c, showD d] + +end diff --git a/thys/SpecCheck/show/show_types.ML b/thys/SpecCheck/show/show_types.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/show/show_types.ML @@ -0,0 +1,17 @@ +(* Title: SpecCheck/show/show_types.ML + Author: Kevin Kappelmann + +Shared type definitions for SpecCheck showable types. +*) + +signature SPECCHECK_SHOW_TYPES = +sig + type 'a show = 'a -> Pretty.T +end + +structure SpecCheck_Show_Types : SPECCHECK_SHOW_TYPES = +struct + +type 'a show = 'a -> Pretty.T + +end diff --git a/thys/SpecCheck/shrink/SpecCheck_Shrink.thy b/thys/SpecCheck/shrink/SpecCheck_Shrink.thy new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/shrink/SpecCheck_Shrink.thy @@ -0,0 +1,15 @@ +\<^marker>\creator "Kevin Kappelmann"\ +section \Shrinkers\ +theory SpecCheck_Shrink +imports SpecCheck_Generators +begin + +paragraph \Summary\ +text \Shrinkers for SpecCheck take a value and return a sequence of smaller values derived from it. +Refer to @{file "shrink_base.ML"} for some basic examples.\ + +ML_file \shrink_types.ML\ +ML_file \shrink_base.ML\ +ML_file \shrink.ML\ + +end diff --git a/thys/SpecCheck/shrink/shrink.ML b/thys/SpecCheck/shrink/shrink.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/shrink/shrink.ML @@ -0,0 +1,11 @@ +(* Title: SpecCheck/shrink/shrink.ML + Author: Kevin Kappelmann, TU Muenchen + +Structure containing all shrink functions. +*) +structure SpecCheck_Shrink = +struct + +open SpecCheck_Shrink_Base + +end diff --git a/thys/SpecCheck/shrink/shrink_base.ML b/thys/SpecCheck/shrink/shrink_base.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/shrink/shrink_base.ML @@ -0,0 +1,89 @@ +(* Title: SpecCheck/shrink/shrink_base.ML + Author: Kevin Kappelmann + +Basic utility functions to create and combine shrink functions. +*) + +signature SPECCHECK_SHRINK_BASE = +sig + include SPECCHECK_SHRINK_TYPES + + val none : 'a shrink + + val product : 'a shrink -> 'b shrink -> ('a * 'b) shrink + val product3 : 'a shrink -> 'b shrink -> 'c shrink -> ('a * 'b * 'c) shrink + val product4 : 'a shrink -> 'b shrink -> 'c shrink -> 'd shrink -> ('a * 'b * 'c * 'd) shrink + + val int : int shrink + + val list : 'a shrink -> ('a list shrink) + val list' : ('a list) shrink + + val term : term shrink + +end + +structure SpecCheck_Shrink_Base : SPECCHECK_SHRINK_BASE = +struct +open SpecCheck_Shrink_Types + +fun none _ = Seq.empty + +fun product_seq xq yq (x, y) = + let + val yqy = Seq.append yq (Seq.single y) + val zq1 = Seq.maps (fn x => Seq.map (pair x) yqy) xq + val zq2 = Seq.map (pair x) yq + in Seq.append zq1 zq2 end + +fun product shrinkA shrinkB (a, b) = product_seq (shrinkA a) (shrinkB b) (a, b) + +fun product3 shrinkA shrinkB shrinkC (a, b, c) = + product shrinkA shrinkB (a, b) + |> (fn abq => product_seq abq (shrinkC c) ((a,b),c)) + |> Seq.map (fn ((a,b),c) => (a,b,c)) + +fun product4 shrinkA shrinkB shrinkC shrinkD (a, b, c, d) = + product3 shrinkA shrinkB shrinkC (a, b, c) + |> (fn abcq => product_seq abcq (shrinkD d) ((a,b,c),d)) + |> Seq.map (fn ((a,b,c),d) => (a,b,c,d)) + +(*bit-shift right until it hits zero (some special care needs to be taken for negative numbers*) +fun int 0 = Seq.empty + | int i = + let + val absi = abs i + val signi = Int.sign i + fun seq 0w0 () = NONE + | seq w () = + let + val next_value = signi * IntInf.~>> (absi, w) + val next_word = Word.- (w, 0w1) + in SOME (next_value, Seq.make (seq next_word)) end + val w = Word.fromInt (IntInf.log2 (abs i)) + in Seq.cons 0 (Seq.make (seq w)) end + +fun list _ [] = Seq.single [] + | list elem_shrink [x] = Seq.cons [] (Seq.map single (elem_shrink x)) + | list elem_shrink (x::xs) = + let + val elems = Seq.append (elem_shrink x) (Seq.single x) + val seq = Seq.maps (fn xs => Seq.map (fn x => x :: xs) elems) (list elem_shrink xs) + in Seq.cons [] seq end + +fun list' xs = list none xs + +fun term (t1 $ t2) = + let + val s1 = Seq.append (term t1) (Seq.single t1) + val s2 = Seq.append (term t2) (Seq.single t2) + val s3 = Seq.map (op$) (product term term (t1, t2)) + in Seq.append s1 (Seq.append s2 s3) end + | term (Abs (x, T, t)) = + let + val s1 = Seq.append (term t) (Seq.single t) + val s2 = Seq.map (fn t => Abs (x, T, t)) (term t) + in Seq.append s1 s2 end + | term _ = Seq.empty + +end diff --git a/thys/SpecCheck/shrink/shrink_types.ML b/thys/SpecCheck/shrink/shrink_types.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/shrink/shrink_types.ML @@ -0,0 +1,17 @@ +(* Title: SpecCheck/shrink/shrink_types.ML + Author: Kevin Kappelmann + +Shared type definitions for SpecCheck shrinkable types. +*) + +signature SPECCHECK_SHRINK_TYPES = +sig + type 'a shrink = 'a -> 'a Seq.seq +end + +structure SpecCheck_Shrink_Types : SPECCHECK_SHRINK_TYPES = +struct + +type 'a shrink = 'a -> 'a Seq.seq + +end diff --git a/thys/SpecCheck/speccheck.ML b/thys/SpecCheck/speccheck.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/speccheck.ML @@ -0,0 +1,205 @@ +(* Title: SpecCheck/speccheck.ML + Author: Lukas Bulwahn, Nicolai Schaffroth, and Kevin Kappelmann TU Muenchen + Author: Christopher League + +Specification-based testing of ML programs. +*) + +signature SPECCHECK = +sig + + (*tries to shrink a given (failing) input to a smaller, failing input*) + val try_shrink : 'a SpecCheck_Property.prop -> 'a SpecCheck_Shrink.shrink -> 'a -> int -> + SpecCheck_Base.stats -> ('a * SpecCheck_Base.stats) + + (*runs a property for a given test case and returns the result and the updated the statistics*) + val run_a_test : 'a SpecCheck_Property.prop -> 'a -> SpecCheck_Base.stats -> + SpecCheck_Base.result_single * SpecCheck_Base.stats + + (*returns the new state after executing the test*) + val check_style : 'a SpecCheck_Output_Style_Types.output_style -> + ('a SpecCheck_Show.show) option -> 'a SpecCheck_Shrink.shrink -> + ('a, 's) SpecCheck_Generator.gen_state -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 's -> 's + val check_shrink : 'a SpecCheck_Show.show -> 'a SpecCheck_Shrink.shrink -> + ('a, 's) SpecCheck_Generator.gen_state -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 's -> 's + val check : 'a SpecCheck_Show.show -> ('a, 's) SpecCheck_Generator.gen_state -> string -> + 'a SpecCheck_Property.prop -> Proof.context -> 's -> 's + val check_base : ('a, 's) SpecCheck_Generator.gen_state -> string -> + 'a SpecCheck_Property.prop -> Proof.context -> 's -> 's + + (*returns all unprocessed elements of the sequence*) + val check_seq_style : 'a SpecCheck_Output_Style_Types.output_style -> + ('a SpecCheck_Show.show) option -> 'a Seq.seq -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 'a Seq.seq + val check_seq : 'a SpecCheck_Show.show -> 'a Seq.seq -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 'a Seq.seq + val check_seq_base : 'a Seq.seq -> string -> 'a SpecCheck_Property.prop -> Proof.context -> + 'a Seq.seq + + (*returns all unused elements of the list*) + val check_list_style : 'a SpecCheck_Output_Style_Types.output_style -> + ('a SpecCheck_Show.show) option -> 'a list -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 'a Seq.seq + val check_list : 'a SpecCheck_Show.show -> 'a list -> string -> 'a SpecCheck_Property.prop -> + Proof.context -> 'a Seq.seq + val check_list_base : 'a list -> string -> 'a SpecCheck_Property.prop -> Proof.context -> + 'a Seq.seq + +end + +structure SpecCheck : SPECCHECK = +struct + +open SpecCheck_Base + +fun run_a_test prop input { + num_success_tests, + num_failed_tests, + num_discarded_tests, + num_recently_discarded_tests, + num_success_shrinks, + num_failed_shrinks, + timing + } = + let + val (time, result) = Timing.timing (fn () => prop input) () + val is_success = case result of Result is_success => is_success | _ => false + val is_discard = case result of Discard => true | _ => false + val is_failure = not is_discard andalso not is_success + + val num_success_tests = num_success_tests + (if is_success then 1 else 0) + val num_failed_tests = num_failed_tests + (if is_failure then 1 else 0) + val num_discarded_tests = num_discarded_tests + (if is_discard then 1 else 0) + val num_recently_discarded_tests = if is_discard then num_recently_discarded_tests + 1 else 0 + val timing = add_timing timing time + + val stats = { + num_success_tests = num_success_tests, + num_failed_tests = num_failed_tests, + num_discarded_tests = num_discarded_tests, + num_recently_discarded_tests = num_recently_discarded_tests, + num_success_shrinks = num_success_shrinks, + num_failed_shrinks = num_failed_shrinks, + timing = timing + } + in (result, stats) end + +fun add_num_success_shrinks stats n = { + num_success_tests = #num_success_tests stats, + num_failed_tests = #num_failed_tests stats, + num_discarded_tests = #num_discarded_tests stats, + num_recently_discarded_tests = #num_recently_discarded_tests stats, + num_success_shrinks = #num_success_shrinks stats + n, + num_failed_shrinks = #num_failed_shrinks stats, + timing = #timing stats +} + +fun add_num_failed_shrinks stats n = { + num_success_tests = #num_success_tests stats, + num_failed_tests = #num_failed_tests stats, + num_discarded_tests = #num_discarded_tests stats, + num_recently_discarded_tests = #num_recently_discarded_tests stats, + num_success_shrinks = #num_success_shrinks stats, + num_failed_shrinks = #num_failed_shrinks stats + n, + timing = #timing stats +} + +fun try_shrink prop shrink input max_shrinks stats = + let + fun is_failure input = case run_a_test prop input empty_stats |> fst of + Result is_success => not is_success + | Discard => false + | Exception _ => false (*do not count exceptions as a failure because the shrinker might + just have broken some invariant of the function*) + fun find_first_failure xq pulls_left stats = + if pulls_left <= 0 + then (NONE, pulls_left, stats) + else + case Seq.pull xq of + NONE => (NONE, pulls_left, stats) + | SOME (x, xq) => + if is_failure x + then (SOME x, pulls_left - 1, add_num_success_shrinks stats 1) + else find_first_failure xq (pulls_left - 1) (add_num_failed_shrinks stats 1) + in + (*always try the first successful branch and abort without backtracking once no further + shrink is possible.*) + case find_first_failure (shrink input) max_shrinks stats of + (*recursively shrink*) + (SOME input, shrinks_left, stats) => try_shrink prop shrink input shrinks_left stats + | (NONE, _, stats) => (input, stats) + end + +fun test output_style init_stats show_opt shrink opt_gen name prop ctxt state = + let + val max_discard_ratio = Config.get ctxt SpecCheck_Configuration.max_discard_ratio + val max_success = Config.get ctxt SpecCheck_Configuration.max_success + (*number of counterexamples to generate before stopping the test*) + val num_counterexamples = + let val conf_num_counterexamples = + Config.get ctxt SpecCheck_Configuration.num_counterexamples + in if conf_num_counterexamples > 0 then conf_num_counterexamples else ~1 end + val max_shrinks = Config.get ctxt SpecCheck_Configuration.max_shrinks + + fun run_tests opt_gen state stats counterexamples = + (*stop the test run if enough (successful) tests were run or counterexamples were found*) + if #num_success_tests stats >= max_success then (Success stats, state) + else if num_tests stats >= max_success orelse + #num_failed_tests stats = num_counterexamples + then (Failure (stats, failure_data counterexamples), state) + else (*test input and attempt to shrink on failure*) + let val (input_opt, state) = opt_gen state + in + case input_opt of + NONE => + if #num_failed_tests stats > 0 + then (Failure (stats, failure_data counterexamples), state) + else (Success stats, state) + | SOME input => + let val (result, stats) = run_a_test prop input stats + in + case result of + Result true => run_tests opt_gen state stats counterexamples + | Result false => + let val (counterexample, stats) = try_shrink prop shrink input max_shrinks stats + in run_tests opt_gen state stats (counterexample :: counterexamples) end + | Discard => + if #num_recently_discarded_tests stats > max_discard_ratio + then + if #num_failed_tests stats > 0 + then (Failure (stats, failure_data counterexamples), state) + else (Gave_Up stats, state) + else run_tests opt_gen state stats counterexamples + | Exception exn => + (Failure (stats, failure_data_exn (input :: counterexamples) exn), state) + end + end + in + Timing.timing (fn _ => run_tests opt_gen state init_stats []) () + |> uncurry (apfst o output_style show_opt ctxt name) + |> snd + end + +fun check_style style show_opt shrink = + test style empty_stats show_opt shrink o SpecCheck_Generator.map SOME + +fun check_shrink show = check_style SpecCheck_Default_Output_Style.default (SOME show) +fun check show = check_shrink show SpecCheck_Shrink.none + +fun check_base gen = + check_style SpecCheck_Default_Output_Style.default NONE SpecCheck_Shrink.none gen + +fun check_seq_style style show_opt xq name prop ctxt = + test style empty_stats show_opt SpecCheck_Shrink.none SpecCheck_Generator.of_seq name prop ctxt + xq + +fun check_seq show = check_seq_style SpecCheck_Default_Output_Style.default (SOME show) +fun check_seq_base xq = check_seq_style SpecCheck_Default_Output_Style.default NONE xq + +fun check_list_style style show = check_seq_style style show o Seq.of_list +fun check_list show = check_seq show o Seq.of_list +fun check_list_base xs = check_seq_base (Seq.of_list xs) + +end diff --git a/thys/SpecCheck/speccheck_base.ML b/thys/SpecCheck/speccheck_base.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/speccheck_base.ML @@ -0,0 +1,111 @@ +(* Title: SpecCheck/speccheck_base.ML + Author: Kevin Kappelmann + +Types returned by single tests and complete test runs as well as simple utility methods on them. +*) + +signature SPECCHECK_BASE = +sig + + datatype result_single = Result of bool | Discard | Exception of exn + + type stats = { + num_success_tests : int, + num_failed_tests : int, + num_discarded_tests : int, + num_recently_discarded_tests : int, + num_success_shrinks : int, + num_failed_shrinks : int, + timing : Timing.timing + } + + val add_timing : Timing.timing -> Timing.timing -> Timing.timing + + val empty_stats : stats + val num_tests : stats -> int + val num_shrinks : stats -> int + + type 'a failure_data = { + counterexamples : 'a list, + the_exception : exn option + } + + val failure_data : 'a list -> 'a failure_data + val failure_data_exn : 'a list -> exn -> 'a failure_data + + datatype 'a result = + Success of stats | + Gave_Up of stats | + Failure of stats * 'a failure_data + + val stats_of_result : 'a result -> stats + +end + +structure SpecCheck_Base : SPECCHECK_BASE = +struct + +datatype result_single = Result of bool | Discard | Exception of exn + +type stats = { + num_success_tests : int, + num_failed_tests : int, + num_discarded_tests : int, + num_recently_discarded_tests : int, + num_success_shrinks : int, + num_failed_shrinks : int, + timing : Timing.timing +} + +val empty_stats = { + num_success_tests = 0, + num_failed_tests = 0, + num_discarded_tests = 0, + num_recently_discarded_tests = 0, + num_success_shrinks = 0, + num_failed_shrinks = 0, + timing = { + cpu = Time.zeroTime, + elapsed = Time.zeroTime, + gc = Time.zeroTime + } +} + +fun add_timing {elapsed = elapsed1, cpu = cpu1, gc = gc1} + {elapsed = elapsed2, cpu = cpu2, gc = gc2} = { + elapsed = elapsed1 + elapsed2, + cpu = cpu1 + cpu2, + gc = gc1 + gc2 +} + +fun num_tests {num_success_tests, num_failed_tests, ...} = + num_success_tests + num_failed_tests + +fun num_shrinks {num_success_shrinks, num_failed_shrinks, ...} = + num_success_shrinks + num_failed_shrinks + +type 'a failure_data = { + counterexamples : 'a list, + the_exception : exn option +} + +fun failure_data counterexamples = { + counterexamples = counterexamples, + the_exception = NONE +} + +fun failure_data_exn counterexamples exn = { + counterexamples = counterexamples, + the_exception = SOME exn +} + +datatype 'a result = + Success of stats | + Gave_Up of stats | + Failure of stats * 'a failure_data + +fun stats_of_result (Success stats) = stats + | stats_of_result (Gave_Up stats) = stats + | stats_of_result (Failure (stats, _)) = stats + +end diff --git a/thys/SpecCheck/util.ML b/thys/SpecCheck/util.ML new file mode 100644 --- /dev/null +++ b/thys/SpecCheck/util.ML @@ -0,0 +1,19 @@ +(* Title: SpecCheck/util.ML + Author: Kevin Kappelmann + +General utility functions for SpecCheck. +*) + +signature SPECCHECK_UTIL = +sig + +val spaces : string list -> string + +end + +structure SpecCheck_Util : SPECCHECK_UTIL = +struct + +val spaces = space_implode " " + +end diff --git a/thys/Van_der_Waerden/Digits.thy b/thys/Van_der_Waerden/Digits.thy new file mode 100644 --- /dev/null +++ b/thys/Van_der_Waerden/Digits.thy @@ -0,0 +1,249 @@ +theory Digits + imports Complex_Main +begin + +section \Representation of integers in different bases\ + +text \\ + +text \ First, we look at some useful lemmas for splitting sums. \ +lemma split_sum_first_elt_less: assumes "ni\{n..i\{Suc n ..jjj=i..jj\{.. {i..i < n\ by (intro sum.cong) auto + also have "\ = (\jj=i..jjj=i..In order to use representation of numbers in a basis \base\ and to calculate the conversion +to and from integers, we introduce the following locale.\ +locale digits = + fixes base :: nat + assumes base_pos: "base > 0" +begin + +text \Conversion from basis base to integers: \from_digits n d\ + +\begin{tabular}{lcp{8cm}} + n:& \nat\& length of representation in basis base\\ + d:& \nat \ nat\& function of digits in basis base where \d i\ is the $i$-th digit in basis base\\ + output:& \nat\& natural number corresponding to $d(n-1) \dots d(0)$ as integer\\ +\end{tabular} +\ +fun from_digits :: "nat \ (nat \ nat) \ nat" where + "from_digits 0 d = 0" +| "from_digits (Suc n) d = d 0 + base * from_digits n (d \ Suc)" + +text \Alternative definition using sum:\ +lemma from_digits_altdef: "from_digits n d = (\iDigit in basis base of some integer number: \digit x i\ + +\begin{tabular}{lcp{8cm}} + x:& \nat\& integer\\ + i:& \nat\& index\\ + output:& \nat\& $i$-th digit of representation in basis base of $x$\\ +\end{tabular} +\ +fun digit :: "nat \ nat \ nat" where + "digit x 0 = x mod base" +| "digit x (Suc i) = digit (x div base) i" + +text \Alternative definition using divisor and modulo:\ +lemma digit_altdef: "digit x i = (x div (base ^ i)) mod base" + by (induction x i rule: digit.induct) (auto simp: div_mult2_eq) + +text \Every digit must be smaller that the base.\ +lemma digit_less_base: "digit x i < base" + using base_pos by (auto simp: digit_altdef) + +text \A representation in basis \base\ of length $n$ must be less than $\base\ ^ n$.\ +lemma from_digits_less: + assumes "\i Suc) \ base ^ n -1" using 2 + by (metis One_nat_def Suc_leI Suc_pred base_pos comp_apply + less_Suc_eq_le zero_less_power) + moreover have "d 0 \ base -1" using 2 + by (metis One_nat_def Suc_pred base_pos less_Suc_eq_0_disj + less_Suc_eq_le) + ultimately have "d 0 + base * from_digits n (d \ Suc) \ + base - 1 + base * (base^(n) - 1)" + by (simp add: add_mono_thms_linordered_semiring(1)) + then show "from_digits (Suc n) d < base ^ Suc n" + using base_pos by (auto simp:comp_def) + (metis Suc_pred add_gr_0 le_imp_less_Suc mult_Suc_right + zero_less_power) +qed auto + +text \Lemmas for \mod\ and \div\ in number systems of basis \base\:\ +lemma mod_base: assumes "\i. i d i < base" "n>0" + shows "from_digits n d mod base = d 0 " +proof - + have "(\iii. d i * base ^ i mod base)" "n-1"] + unfolding from_digits_altdef by simp +qed + +lemma mod_base_i: + assumes "\i. i d i < base" "n>0" "ij=i..j=i..j=i..i. i d i < base" "n>0" "ij=i..{i..jjjj=i..jjj=i..ij = i..Conversions are inverse to each other.\ +lemma digit_from_digits: + assumes "\j. j d j < base" "n>0" "iix + by auto (metis less_mult_imp_div_less x_n_def) + then have "x_n mod base = x_n" by simp + have x_less_i_eq_x_i:"x mod base^n div base ^i mod base = + x div base^i mod base" if "i = x mod base^n div base^i mod base" + using div_distrib[where a="x div base^n" and b = "x mod base^n"] + that by auto + finally show ?thesis by simp + qed + have "x = (x_n mod base)*base^n + x_less" + unfolding \x_n mod base=x_n\ + using x_n_def x_less_def div_mod_decomp by blast + also have "\ = (x div base^n mod base) * base^n + + (\iStronger formulation of above lemma.\ +lemma from_digits_digit': + "from_digits n (digit x) = x mod (base ^ n)" + unfolding from_digits_altdef digit_altdef +proof (induction n arbitrary: x) + case 0 + then show ?case by simp +next + case (Suc n) + define x_less where "x_less = x mod base^n" + define x_n where "x_n = x div base^n mod base" + have "x_less < base^n" using x_less_def base_pos + mod_less_divisor by presburger + then have IH_x_less: + "(\i = x mod base^n div base^i mod base" + using div_distrib[where a="x div base^n" and b = "x mod base^n"] + that by auto + finally show ?thesis by simp + qed + have "x mod base^Suc n = x_n*base^n + x_less" + by (metis mod_mult2_eq mult.commute power_Suc2 x_less_def x_n_def) + also have "\ = (x div base^n mod base) * base^n + + (\iVan der Waerden's Theorem\ + +text \In combinatorics, Van der Waerden's Theorem is about arithmetic progressions of a certain +length of the same colour in a colouring of an interval. In order to state the theorem and to +prove it, we need to formally introduce arithmetic progressions. We will express $k$-colourings as +functions mapping an integer interval to the set $\{0,\dots , k-1 \}$ of colours.\ + +subsection \Arithmetic progressions\ + +text \A sequence of integer numbers with the same step size is called an arithmetic progression. + We say an $m$-fold arithmetic progression is an arithmetic progression with multiple step +lengths.\ + +text \ Arithmetic progressions are defined in the following using the variables: + +\begin{tabular}{lcp{8cm}} +$start$:& \int\& starting value\\ +$step$:& \nat\& positive integer for step length\\ +$i$:& \nat\& $i$-th value in the arithmetic progression \\ +\end{tabular}\ + +definition arith_prog :: "int \ nat \ nat \ int" + where "arith_prog start step i = start + int (i * step)" + +text \ An $m$-fold arithmetic progression (which we will also call a multi-arithmetic progression) +is defined in the following using the variables: + +\begin{tabular}{lcp{8cm}} +$dims$:& \nat\& number of dimensions/step directions of $m$-fold arithmetic progression\\ +$start$:& \int\& starting value\\ +$steps$:& \nat \ nat\& function of steps, returns step in $i$-th dimension for $i\in[0..nat \ nat\& function of coefficients, returns coefficient in $i$-th dimension for + $i\in[0.. + +definition multi_arith_prog :: + "nat \ int \ (nat \ nat) \ (nat \ nat) \ int" + where "multi_arith_prog dims start steps c = + start + int (\iAn $m$-fold arithmetic progression of dimension $1$ is also an arithmetic progression and + vice versa. This is shown in the following lemmas.\ +lemma multi_to_arith_prog: + "multi_arith_prog 1 start steps c = + arith_prog start (steps 0) (c 0)" + unfolding multi_arith_prog_def arith_prog_def by auto + +lemma arith_prog_to_multi: + "arith_prog start step c = + multi_arith_prog 1 start (\_. step) (\_. c)" + unfolding multi_arith_prog_def arith_prog_def by auto + +text \To show that an arithmetic progression is well-defined, we introduce the following predicate. +It assures that \arith_prog start step ` [0.. is contained in the integer interval $[a..b]$.\ +definition is_arith_prog_on :: + "nat \ int \ nat \ int \ int \ bool" + where "is_arith_prog_on l start step a b \ + (start \ a \ arith_prog start step (l-1) \ b)" + +text \Furthermore, we have monotonicity for arithmetic progressions.\ +lemma arith_prog_mono: + assumes "c \ c'" + shows "arith_prog start step c \ arith_prog start step c'" + using assms unfolding arith_prog_def by (auto intro: mult_mono) + +text \Now, we state the well-definedness of an arithmetic progression of length $l$ in an integer +interval $[a..b]$. +Indeed, \is_arith_prog_on\ guarantees that every element of \arith_prog start step\ of length $l$ + lies in $[a..b]$.\ +lemma is_arith_prog_onD: + assumes "is_arith_prog_on l start step a b" + assumes "c \ {0.. {a..b}" +proof - + have "arith_prog start step 0 \ arith_prog start step c" + by (rule arith_prog_mono) auto + hence "arith_prog start step c \ a" + using assms by (simp add: arith_prog_def is_arith_prog_on_def + add_increasing2) + moreover have "arith_prog start step (l-1) \ + arith_prog start step c" + by (rule arith_prog_mono) (use assms(2) in auto) + hence "arith_prog start step c \ b" + using assms unfolding arith_prog_def is_arith_prog_on_def + by linarith + ultimately show ?thesis + by auto +qed + +text \We also need a predicate for an $m$-fold arithmetic progression to be well-defined. +It assures that \multi_arith_prog start step ` [0.. is contained in $[a..b]$.\ +definition is_multi_arith_prog_on :: + "nat \ nat \ int \ (nat \ nat) \ int \ int \ bool" + where "is_multi_arith_prog_on l m start steps a b \ + (start \ a \ multi_arith_prog m start steps (\_. l-1) \ b)" + +text \Moreover, we have monotonicity for $m$-fold arithmetic progressions as well.\ +lemma multi_arith_prog_mono: + assumes "\i. i < m \ c i \ c' i" + shows "multi_arith_prog m start steps c \ + multi_arith_prog m start steps c'" + using assms unfolding multi_arith_prog_def + by (auto intro!: sum_mono intro: mult_right_mono) + +text \Finally, we get the well-definedness for $m$-fold arithmetic progressions of length $l$. +Here, \is_multi_arith_prog_on\ guarantees that every element of \multi_arith_prog start step\ + of length $l$ lies in $[a..b]$.\ +lemma is_multi_arith_prog_onD: + assumes "is_multi_arith_prog_on l m start steps a b" + assumes "c \ {0.. {0.. {a..b}" +proof - + have "multi_arith_prog m start steps (\_. 0) \ + multi_arith_prog m start steps c" + by (rule multi_arith_prog_mono) auto + hence "multi_arith_prog m start steps c \ a" + using assms by (simp add: multi_arith_prog_def + is_multi_arith_prog_on_def) + moreover have "multi_arith_prog m start steps (\_. l-1) \ + multi_arith_prog m start steps c" + by (rule multi_arith_prog_mono) (use assms in force) + hence "multi_arith_prog m start steps c \ b" + using assms by (simp add: multi_arith_prog_def + is_multi_arith_prog_on_def) + ultimately show ?thesis + by auto +qed + + +subsection \Van der Waerden's Theorem\ + +text \The property for a number $n$ to fulfill Van der Waerden's theorem is the following:\\ +For a $k$-colouring col of $[a..b]$ there exist +\begin{itemize} +\item $start$: starting value of an arithmetic progression +\item $step$: step length of an arithmetic progression +\item $j$: colour +\end{itemize} +such that \arith_prog start step\ is a valid arithmetic progression of length $l$ lying +in $[a..b]$ of the same colour $j$. + +The following variables will be used:\\ +\begin{tabular}{lcp{8cm}} +$k$:& \nat\& number of colours in segment colouring on $[a..b]$\\ +$l$:& \nat\& length of arithmetic progression\\ +$n$:& \nat\& number fulfilling Van der Waerden's Theorem\\ +\end{tabular} +\ +definition vdw :: + "nat \ nat \ nat \ bool" + where "vdw k l n \ + (\a b col. b + 1 \ a + int n \ col \ {a..b} \ {.. + (\j start step. j < k \ step > 0 \ + is_arith_prog_on l start step a b \ + arith_prog start step ` {.. col -` {j} \ {a..b}))" + +text \To better work with the property of Van der Waerden's theorem, we introduce an + elimination rule.\ +lemma vdwE: + assumes "vdw k l n" + "b + 1 \ a + int n" + "col \ {a..b} \ {.. 0" + "is_arith_prog_on l start step a b" + "arith_prog start step ` {.. col -` {j} \ {a..b}" + using assms that unfolding vdw_def by metis + +text \Van der Waerden's theorem implies that the number fulfilling it is positive. This is show +in the following lemma.\ +lemma vdw_imp_pos: + assumes "vdw k l n" + "l > 0" + shows "n > 0" +proof (rule Nat.gr0I) + assume [simp]: "n = 0" + show False + using assms + by (elim vdwE[where a = 1 and b = 0 and col = "\_. 0"]) + (auto simp: lessThan_empty_iff) +qed + +text \Van der Waerden's Theorem is trivial for a non-existent colouring. +It also makes no sense for arithmetic progressions of length 0.\ +lemma vdw_0_left [simp, intro]: "n>0 \ vdw 0 l n" + by (auto simp: vdw_def) + +text \In the case of $k=1$, Van der Waerden's Theorem holds. Then every number has the same colour, +hence also the arithmetic progression. A possible choice for the number fulfilling Van der +Waerden Theorem is $l$.\ +lemma vdw_1_left: + assumes "l>0" + shows "vdw 1 l l" +unfolding vdw_def +proof (safe, goal_cases) + case (1 a b col) + have "arith_prog a 1 ` {.. {a..b}" + using 1(1) by (auto simp: arith_prog_def) + also have "{a..b} = col -` {0} \ {a..b}" + using 1(2) by auto + finally have "arith_prog a 1 ` {.. col -` {0} \ {a..b}" + by auto + moreover have "is_arith_prog_on l a 1 a b" + unfolding is_arith_prog_on_def arith_prog_def using 1 assms + by auto + ultimately show "\j start step. j < 1 \ 0 < step \ + is_arith_prog_on l start step a b \ + arith_prog start step ` {.. col -` {j} \ {a..b}" + by auto +qed + +text \In the case $l=1$, Van der Waerden's Theorem holds. As the length of the arithmetic +progression is $1$, it consists of just one element. Thus every nonempty integer interval fulfills +the Van der Waerden property. We can prove $N_{k,1}$ to be $1$.\ +lemma vdw_1_right: "vdw k 1 1" +unfolding vdw_def +proof safe + fix a b :: int and col :: "int \ nat" + assume *: "a + int 1 \ b + 1" "col \ {a..b} \ {.. col -` {col a} \ {a..b}" + using * by auto + finally have "arith_prog a 1 ` {..<1} \ col -` {col a} \ {a..b}" + by auto + moreover have "is_arith_prog_on 1 a 1 a b" + unfolding is_arith_prog_on_def arith_prog_def + using * by auto + ultimately show "\j start step. + j < k \ 0 < step \ is_arith_prog_on 1 start step a b \ + arith_prog start step ` {..<1} \ col -` {j} \ {a..b}" + using \col a by blast +qed + +text \In the case $l=2$, Van der Waerden's Theorem holds as well. Here, any two distinct numbers +form an arithmetic progression of length $2$. Thus we only have to find two numbers with the same +colour. +Using the pigeonhole principle on $k+1$ values, we can find two integers with the same colour.\ +lemma vdw_2_right: "vdw k 2 (k+1)" +unfolding vdw_def +proof safe + fix a b :: int and col :: "int \ nat" + assume *: "a + int (k + 1) \ b + 1" "col \ {a..b} \ {.. {.. card {a..b}" using *(1) by auto + ultimately have "card (col ` {a..b}) < card {a..b}" using * + by (metis card_lessThan card_mono finite_lessThan le_less_trans + less_add_one not_le) + then have "\ inj_on col {a..b}" using pigeonhole[of col "{a..b}"] + by auto + then obtain start start_step + where pigeon: "col start = col start_step" + "start < start_step" + "start \ {a..b}" + "start_step \ {a..b}" + using inj_onI[of "{a..b}" col] + by (metis not_less_iff_gr_or_eq) + + define step where "step = nat (start_step - start)" + define j where "j = col start" + + have "j < k" unfolding j_def using *(2) pigeon(3) by auto + moreover have "0 < step" unfolding step_def using pigeon(2) by auto + moreover have "is_arith_prog_on 2 start step a b" + unfolding is_arith_prog_on_def arith_prog_def step_def + using pigeon by auto + moreover { + have "arith_prog start step i \ {start, start_step}" if "i<2" for i + using that arith_prog_def step_def by (auto simp: less_2_cases_iff) + also have "\ \ col -` {j} \ {a..b}" + using pigeon unfolding j_def by auto + finally have "arith_prog start step ` {..<2} \ col -` {j} \ {a..b}" + by auto + } + ultimately show "\j start step. + j < k \ + 0 < step \ + is_arith_prog_on 2 start step a b \ + arith_prog start step ` {..<2} \ col -` {j} \ {a..b}" by blast +qed + +text \In order to prove Van der Waerden's Theorem, we first prove a slightly different lemma. +The statement goes as follows:\\ +For a $k$-colouring $col$ on $[a..b]$ there exist +\begin{itemize} +\item $start$: starting value of an arithmetic progression +\item $steps$: step length of an arithmetic progression +\end{itemize} +such that \f = multi_arith_prog m start step\ is a valid $m$-fold arithmetic progression of +length $l$ lying in $[a..b]$ such that for every $snat\& number of colours in segment colouring of $[a..b]$\\ +$m$:& \nat\& dimension of $m$-fold arithmetic progression\\ +$l$:& \nat\& $l+1$ is length of $m$-fold arithmetic progression\\ +$n$:& \nat\& number fulfilling \vdw_lemma\\\ +\end{tabular} +\ +definition vdw_lemma :: "nat \ nat \ nat \ nat \ bool" where + "vdw_lemma k m l n \ + (\a b col. b + 1 \ a + int n \ col \ {a..b} \ {.. + (\start steps. (\i 0) \ + is_multi_arith_prog_on (l+1) m start steps a b \ ( + let f = multi_arith_prog m start steps + in (\c \ {0.. {0..l}. \s j \ s. c j < l) \ + col (f c) = col (f (\i. if i \ s then 0 else c i))))))" + +text \To better work with this property, we introduce an elimination rule for \vdw_lemma\.\ +lemma vdw_lemmaE: + fixes a b :: int + assumes "vdw_lemma k m l n" + "b + 1 \ a + int n" "col \ {a..b} \ {..i. i < m \ steps i > 0" + "is_multi_arith_prog_on (l+1) m start steps a b" + "let f = multi_arith_prog m start steps + in \c \ {0.. {0..l}. \s j \ s. c j < l) \ + col (f c) = col (f (\i. if i \ s then 0 else c i))" + using assms that unfolding vdw_lemma_def by blast + +text \To simplify the following proof, we show the following formula.\ +lemma sum_mod_poly: + assumes "(k::nat)>0" + shows "(k - 1) * (\ n\{..nn = int k ^ q - 1" + by (induction q) (auto simp: algebra_simps) + also have "\ < int (k ^ q)" + by simp + finally show ?thesis by linarith +qed + +text \The proof of Van der Waerden's Theorem now proceeds in three steps:\\ +\begin{itemize} +\item Firstly, we show that the \vdw\ property for all $k$ proves the \vdw_lemma\ for fixed $l$ but +arbitrary $k$ and $m$. This is done by induction over $m$. +\item Secondly, we show that \vdw_lemma\ implies the induction step of \vdw\ using the pigeonhole +principle. +\item Lastly, we combine the previous steps in an induction over $l$ to show Van der Waerden's +Theorem in the general setting. +\end{itemize}\ + +text \Firstly, we need to show that \vdw\ for arbitrary $k$ implies \vdw_lemma\ for fixed $l$. +As mentioned earlier, we use induction over $m$.\ +lemma vdw_imp_vdw_lemma: + fixes l + assumes vdw_assms: "\k'. k'>0 \ \n_k'. vdw k' l n_k'" + and "l \ 2" + and "m > 0" + and "k > 0" + shows "\N. vdw_lemma k m l N" +using \m>0\ \k>0\ proof (induction m rule: less_induct) + case (less m) + consider "m=1" | "m>1" using less.prems by linarith + then show ?case + proof cases + text \ Case $m=1$: Show \vdw_lemma\ for arithmetic progression, Induction start. \ + assume "m = 1" + + obtain n where vdw: "vdw k l n" using vdw_assms \k>0\ by blast + define N where "N = 2*n" + have "l>0" and "l>1" using \l\2\ by auto + + have "vdw_lemma k m l N" + unfolding vdw_lemma_def + proof (safe, goal_cases) + case (1 a b col) + text \ Divide $[a..b]$ in two intervals $I_1$, $I_2$ of same length and obtain arithmetic + progression of length $l$ in $I_1$. \ + have col_restr: "col \ {a..a + int n - 1} \ {.. 0" + "is_arith_prog_on l start step a (a + int n -1)" + "arith_prog start step ` {.. + col -` {j} \ {a..a + int n - 1}" + using vdw 1 unfolding N_def by (elim vdwE)(auto simp:is_arith_prog_on_def) + have range_prog_lessThan_l: + "arith_prog start step i \ {a..a + int n -1}" if "i < l" for i + using that prog by auto + + have "{a..a + int n-1}\{a..b}" using N_def "1"(1) by auto + then have "a + 2* int n - 1 \ b" using 1(1) unfolding N_def + by auto + + text \ Show that \arith_prog start step\ is an arithmetic progression of length $l+1$ + in $[a..b]$. \ + have prog_in_ivl: "arith_prog start step i \ {a..b}" + if "i \ l" for i + proof (cases "i=l") + case False + have "i{a..a + int n-1}\{a..b}\ by force + next + case True + text \ Show $\step\\leq |I_1|$ then have \arith_prog start step (l+1)\[a..b]\ as + \arith_prog start step (l+1) = arith_prog start step l + step\ \ + have "start \ {a..a + int n -1}" + using range_prog_lessThan_l[of 0] + unfolding arith_prog_def by (simp add: \0 < l\) + moreover have "start + int step \ {a..a + int n -1}" + using range_prog_lessThan_l[of 1] + unfolding arith_prog_def by (metis \1 < l\ mult.left_neutral) + ultimately have "step \ n" by auto + have "arith_prog start step (l-1) \ {a..a + int n -1}" + using range_prog_lessThan_l[of "l-1"] unfolding arith_prog_def + using \0 < l\ diff_less less_numeral_extra(1) by blast + moreover have "arith_prog start step l = + arith_prog start step (l-1) + int step" + unfolding arith_prog_def using \0 < l\ mult_eq_if by force + ultimately have "arith_prog start step l \ {a..b}" + using \step\n\ N_def \a + 2* int n -1 \ b\ by auto + then show ?thesis using range_prog_lessThan_l using True + by force + qed + + have col_prog_eq: "col (arith_prog start step k) = j" + if "k < l" for k + using prog that by blast + + define steps :: "nat \ nat" where steps_def: "steps = (\i. step)" + define f where "f = multi_arith_prog 1 start steps" + + have rel_prop_1: + "col (f c) = col (f (\i. if i < s then 0 else c i))" + if "c \ {0..<1} \ {0..l}" "s<1" "\j\s. c j < l" for c s + using that by auto + + have arith_prog_on: + "is_multi_arith_prog_on (l+1) m start steps a b" + using prog(3) unfolding is_arith_prog_on_def is_multi_arith_prog_on_def + using \m=1\ arith_prog_to_multi steps_def prog_in_ivl by auto + + show ?case + by (rule exI[of _ start], rule exI[of _ steps]) + (use rel_prop_1 \step > 0\ \m = 1\ arith_prog_on col_prog_eq + multi_to_arith_prog in \auto simp: f_def Let_def steps_def\) + qed + then show ?case .. + + next + text \ Case $m>1$: Show \vdw_lemma\ for $m$-fold arithmetic progression, + Induction step $(m-1) \longrightarrow m$. \ + assume "m>1" + + obtain q where vdw_lemma_IH:"vdw_lemma k (m-1) l q" + using \1 < m\ less by force + have "k^q>0" using \k>0\ by auto + obtain n_kq where vdw: "vdw (k^q) l n_kq" + using vdw_assms \k^q>0\ by blast + define N where "N = q + 2 * n_kq" + + text \Idea: $[a..b] = I_1 \cup I_2$ where $|I_1| = 2*n_{k,q}$ and $|I_2| = q$. + Divide $I_1$ into blocks of length $q$ and define a new colouring on the set of + $q$-blocks where the colour of the block is the $k$-basis representation where + the $i$-th digit corresponds to the colour of the $i$-th element in the block. + Get an arithmetic progression of $q$-blocks of length $l+1$ in $I_1$, such that + the first $l$ $q$-blocks have the same colour. + The step of the block-arithmetic progression is going to be the additional + step in the induction over $m$. \ + + have "vdw_lemma k m l N" + unfolding vdw_lemma_def + proof (safe, goal_cases) + case (1 a b col) + have "n_kq>0" using vdw_imp_pos vdw \l\2\ by auto + then have "N>0" by (simp add:N_def) + then have "a\b" using 1 by auto + then have "k>0" using 1 by (intro Nat.gr0I) force + have "l>0" and "l>1" using \l\2\ by auto + interpret digits k by (simp add: \0 < k\ digits_def) + define col1 where "col1 = (\ x. from_digits q (\y. col (x + y)))" + have range_col1: "col1\{a..a + int n_kq - 1} \ {..{a..a + int n_kq - 1}" + then have col_xn:"col (x + int n)\{.. k - 1" + if "nk>0\ by (auto) + have "(\n + (\n = (k-1) * (\n < k^q" using sum_mod_poly \k>0\ by auto + finally show "col1 x 0" + "is_arith_prog_on l start step a (a + int n_kq - 1)" + "arith_prog start step ` {.. + col1 -` {j} \ {a..a + int n_kq -1}" + using vdw range_col1 by (elim vdwE) (auto simp: \k>0\) + + have range_prog_lessThan_l: + "arith_prog start step i \ {a..a + int n_kq -1}" + if "i < l" for i + using that prog by auto + + have prog_in_ivl: + "arith_prog start step i \ {a..a + 2 * int n_kq -1}" + if "i \ l" for i + proof (cases "i=l") + case False + then have "i {a..a + int n_kq -1}" + using range_prog_lessThan_l[of 0] unfolding arith_prog_def + by (simp add: \0 < l\) + moreover have "start + step \ {a..a + int n_kq -1}" + using range_prog_lessThan_l[of 1] unfolding arith_prog_def + by (metis \1 < l\ mult.left_neutral) + ultimately have "step \ n_kq" by auto + have "arith_prog start step (l-1) \ {a..a + int n_kq -1}" + using range_prog_lessThan_l[of "l-1"] unfolding arith_prog_def + using \0 < l\ diff_less less_numeral_extra(1) by blast + moreover have "arith_prog start step l = + arith_prog start step (l-1) + step" + unfolding arith_prog_def using \0 < l\ mult_eq_if by force + ultimately have "arith_prog start step l \ + {a..a + 2 * int n_kq - 1}" + using \step\n_kq\ by auto + then show ?thesis using range_prog_lessThan_l using True + by force + qed + + have col_prog_eq: "col1 (arith_prog start step k) = j" + if "k < l" for k + using prog that by blast + + have digit_col1:"digit (col1 x) y = col (x+int y)" + if "x\{a..{..j'. j' x+j'\{a..b}" + using "1"(1) N_def that(1) by force + then have "\j'. j' (\y. col (x+int y)) j' < k" + using 1 that by auto + then show "digit (from_digits q (\xa. col (x + int xa))) y = + col (x + int y)" + using digit_from_digits that 1 by auto + qed + + text \ Impact on the colour when taking the block-step. \ + have one_step_more: + "col (arith_prog start' step i) = digit j (nat (start'-start))" + if "start'\{start..{.. start'" using that by simp + have shift_arith_prog: + "arith_prog start step i + (start' - start) = + arith_prog start' step i" + unfolding arith_prog_def by simp + define diff where "diff = nat (start'-start)" + have "diff \{..{a..a + 2 * int n_kq-1}" + using prog(4) that by auto + ultimately show ?thesis + using digit_col1[where x = "arith_prog start step i" + and y = "diff"] + prog 1 \diff \{.. by auto + qed + then show ?thesis unfolding diff_def 1 + by (auto simp: \start\start'\ shift_arith_prog) + qed + + have one_step_more': "col (arith_prog start' step i) = + col (arith_prog start' step 0)" + if "start'\{start..{.. start + int q - 1 + 1" by linarith + have "{start..start + int q-1} \ {a..b}" + using prog N_def 1(1) by (force simp: arith_prog_def is_arith_prog_on_def) + then have col': "col \ {start..start + int q-1} \ {.. Obtain an $(m-1)$-fold arithmetic progression in the starting $q$-bolck of the + block arithmetic progression. \ + obtain start_m steps_m where + step_m_pos: "\i. i < m - 1 \ 0 < steps_m i" and + is_multi_arith_prog: "is_multi_arith_prog_on (l+1) (m - 1) + start_m steps_m start (start + int q - 1)" and + g_aux: "let g = multi_arith_prog (m - 1) start_m steps_m + in \c\{0.. {0..l}. \sj\s. c j < l) \ + col (g c) = col (g (\i. if i \ s then 0 else c i))" + by (rule vdw_lemmaE[OF vdw_lemma_IH start_q col']) blast + + define g where "g = multi_arith_prog (m-1) start_m steps_m" + have g: "col (g c) = col (g (\i. if i \ s then 0 else c i))" + if "c \ {0..<(m-1)} \ {0..l}" "s < m - 1" "\j \ s. c j < l" + for c s using g_aux that unfolding g_def Let_def by blast + + have range_g: "g c \ {start..start + int q - 1}" + if "c \ {0.. {0..<(l+1)}" for c + using is_multi_arith_prog_onD[OF is_multi_arith_prog that] + by (auto simp: g_def) + + text \Obtain an $m$-fold arithmetic progression by adding the block-step.\ + define steps :: "nat \ nat" where steps_def: + "steps = (\i. (if i=0 then step else steps_m (i-1)))" + define f where "f = multi_arith_prog m start_m steps" + have f_step_g: "f c = int (c 0*step) + g (c \ Suc)" for c + proof - + have "f c = start_m + int (\i = start_m + int (c 0 * steps 0) + + int (\i = start_m + int (c 0 * step) + + int (\i Show that this $m$-fold arithmetic progression fulfills all needed properties. \ + have steps_gr_0: "\i start_m" using is_multi_arith_prog + unfolding is_multi_arith_prog_on_def + using is_arith_prog_on_def prog(3) by force + moreover { + have "f (\_. l) = arith_prog (g ((\_. l) \ Suc)) step l" + using f_step_g unfolding arith_prog_def by auto + also have "g ((\_. l) \ Suc) \ start + q" + using range_g[of "(\_. l) \ Suc"] by auto + then have "arith_prog (g ((\_. l) \ Suc)) step l \ + arith_prog start step l + q" + unfolding arith_prog_def by auto + also have "\\ b" using prog_in_ivl[of l] + using is_multi_arith_prog unfolding is_multi_arith_prog_on_def + using "1"(1) N_def by auto + finally have "f (\_. l) \ b" by auto + } + ultimately show ?thesis + unfolding is_multi_arith_prog_on_def f_def by auto + qed + + text \ Show the relational property for all $s$. \ + have rel_prop_1: + "col (f c) = col (f (\i. if i \ s then 0 else c i))" + if "c \ {0.. {0..l}" "sj\s. c j < l" for c s + proof (cases "s = 0") + case True + have "c 0 < l" using that(3) True by auto + have range_c_Suc: "c \ Suc \ {0.. {0..l}" + using that(1) by auto + have "f c = arith_prog (g (c \ Suc)) step (c 0)" + using f_step_g unfolding arith_prog_def by auto + then have "col (f c) = col (arith_prog (g (c \ Suc)) step 0)" + using one_step_more'[of "g (c \ Suc)" "c 0"] \c 0 < l\ + range_g[of "c \ Suc"] range_c_Suc + atLeastLessThanSuc_atLeastAtMost by auto + also { + have "(\xx=1..x. x-1)" "Suc"]) + (auto simp: steps_def split:if_splits) + also have "\ = (\x Suc)) step 0 = + f (\i. if i \ s then 0 else c i)" + unfolding f_def g_def multi_arith_prog_def arith_prog_def + using True by auto + } + finally show ?thesis by auto + next + case False + hence s_greater_0: "s > 0" by auto + have range_c_Suc: "c \ Suc \ {0.. {0..l}" + using that(1) by auto + have "c 0 < l" using \s>0\ that by auto + have g_IH: + "col (g c') = col (g (\i. if i \ s' then 0 else c' i))" + if "c' \ {0.. {0..l}" "s'j\s'. c' j < l" + for c' s' + using g_aux that unfolding multi_arith_prog_def g_def + by (auto simp: Let_def) + have g_shift_IH: "col (g (c \ Suc)) = + col (g ((\i. if i\{1..t} then 0 else c i) \ Suc))" + if "c \ {1.. {0..l}" "t\{1..j\{1..t}. c j < l" + for c t + proof - + have "(\i. (if i \ t - 1 then 0 else (c \ Suc) i)) = + (\i. (if i \ {1..t} then 0 else c i)) \ Suc" + using that by (auto split: if_splits simp:fun_eq_iff) + then have right: + "g (\i. if i \ (t-1) then 0 else (c \ Suc) i) = + g ((\i. if i\{1..t} then 0 else c i) \ Suc)" by auto + have "(c \ Suc)\ {0.. {0..l}" using that(1) by auto + moreover have "t-1j\t-1. (c \ Suc) j < l" using that by auto + ultimately have "col (g (c \ Suc)) = + col (g (\i. (if i \ t-1 then 0 else (c \ Suc) i)))" + using g_IH[of "(c \ Suc)" "t-1"] by auto + with right show ?thesis by auto + qed + + have "col (f c) = col (int (c 0 * step) + g (c \ Suc))" + using f_step_g by simp + also have "int (c 0 * step) + g (c \ Suc) = + arith_prog (g (c \ Suc)) step (c 0)" + by (simp add: arith_prog_def) + also have "col \ = col (arith_prog (g (c \ Suc)) step 0)" + using one_step_more'[of "g (c \ Suc)" "c 0"] \c 0 < l\ + range_g[of "c \ Suc"] range_c_Suc + atLeastLessThanSuc_atLeastAtMost by auto + also have "\ = col (g (c \ Suc))" + unfolding arith_prog_def by auto + also have "\ = col (g ((\i. if i\{1..s} then 0 else c i) \ + Suc))" using g_shift_IH[of "c" s] \s>0\ that by force + also have "\ = col ((\c. int (c 0 * step) + + g (c \ Suc))(\i. if i\s then 0 else c i))" + by (auto simp: g_def multi_arith_prog_def) + also have "\ = col (f (\i. if i \ s then 0 else c i))" + unfolding f_step_g by auto + finally show ?thesis by simp + qed + + show ?case + by (rule exI[of _ start_m], rule exI[of _ steps]) + (use steps_gr_0 is_multi_on_f rel_prop_1 in + \auto simp: f_def Let_def steps_def\) + qed + then show ?case .. + qed +qed + + + + +text \ Secondly, we show that \vdw_lemma\ implies the induction step of Van der Waerden's Theorem +using the pigeonhole principle. \ +lemma vdw_lemma_imp_vdw: + assumes "vdw_lemma k k l N" + shows "vdw k (Suc l) N" +unfolding vdw_def proof (safe, goal_cases) +text \Idea: Proof uses pigeonhole principle to guarantee the existence of an arithmetic + progression of length $l+1$ with the same colour. \ + case (1 a b col) + obtain start steps where prog: + "\i. i < k \ steps i > 0" + "is_multi_arith_prog_on (l+1) k start steps a b" + "let f = multi_arith_prog k start steps + in \c \ {0.. {0..l}. \s j \ s. c j < l) \ + col (f c) = col (f (\i. if i \ s then 0 else c i))" + using assms 1 + by (elim vdw_lemmaE[where a=a and b=b and col=col and m=k + and k=k and l=l and n=N]) auto + + text \ Obtain a $k$-fold arithmetic progression $f$ of length $l$ from assumptions. \ + define f where "f = multi_arith_prog k start steps" + have rel_propE: "col (f c) = col (f (\i. if i \ s then 0 else c i))" + if "c \ {0.. {0..l}" "s j \ s. c j < l" + for c s + using prog(3) that unfolding f_def Let_def by auto + + text \There are $k+1$ values $a_r = f(0,\dots,0,l,\dots,l)$ with $0\leq r\leq k$ zeros.\ + define a_r where "a_r = (\r. f (\i. (if i {a..b}" unfolding a_r_def f_def + by (intro is_multi_arith_prog_onD[OF prog(2)]) auto + thus ?thesis using 1 by blast + qed + then have "(col \ a_r) ` {.. {.. a_r) ` {.. card {.. inj_on (col \ a_r) {.. a_r" "{..Using the pigeonhole principle get $r_1$ and $r_2$ where $a_{r_1}$ and $a_{r_2}$ have the + same colour.\ + then obtain r1 r2 where pigeon_cols: + "r1\{..{.. a_r) r1 = (col \ a_r) r2" + by (metis (mono_tags, lifting) linear linorder_inj_onI) + text \ Show that the following function $h$ is an arithmetic progression which fulfills all + properties for Van der Waerden's Theorem. \ + define h where + "h = (\x. f (\i. (if ir1 + by (intro arg_cong[where f = f]) auto + moreover have "h l = a_r r1" unfolding h_def a_r_def using \r1 + by (metis le_eq_less_or_eq less_le_trans) + ultimately have "col (h 0) = col (h l)" using pigeon_cols(4) by auto + have h_col: "col (h 0) = col (h i)" if "i\{..col (h 0) = col (h l)\ by auto + next + case False + then have "i{0.. {0..l}" + using that by auto + moreover have "(\j\r2-1. ?c j < l)" + using \i pigeon_cols(3) by force + ultimately have "col (f ?c) = + col (f (\i. if i \ r2-1 then 0 else ?c i))" + using rel_propE[of ?c "r2-1"] pigeon_cols by simp + then show ?thesis unfolding h_def f_def + by (smt (z3) Nat.lessE One_nat_def add_diff_cancel_left' + le_less less_Suc_eq_le multi_arith_prog_mono plus_1_eq_Suc) + qed + + define h_start where "h_start = start + l*(\i\{r2..i\{r1..xx = r2..x = r1..xx = (\xx=r1..x. int (if x < r1 then 0 else y) * int (steps x))"] + \r1 by auto + also have "\ = (\x=r1.. = int y * (\x=r1..x. int (if x < r1 then 0 else if x < r2 then y else l) + * int (steps x))" + have "(\xx=r1.. {r2..r1 < r2\ \r2 < k\ by auto + also have "(\x\\. aux_left x) = (\x=r1..x=r2..x=r1..x=r1..x=r2..x=r2..r1 < r2\ by (intro sum.cong) (auto simp: aux_left_def) + finally show ?thesis + by (simp add: aux_left_def sum_distrib_left) + qed + then show ?thesis + unfolding arith_prog_def h_start_def h_step_def h_def f_def + multi_arith_prog_def by (auto split:if_splits) + qed + + define j where "j = col (h 0)" + have case_j: "jcol (h 0) = col (h l)\ + \h l = a_r r1\ j_def pigeon_cols(1) by auto + have case_step: "h_step > 0" unfolding h_step_def + using pigeon_cols by (intro sum_pos prog(1)) auto + + have range_h: "h i \ {a..b}" if "i < l + 1" for i + unfolding h_def f_def by (rule is_multi_arith_prog_onD[OF prog(2)]) + (use that in auto) + + have case_on: "is_arith_prog_on (l+1) h_start h_step a b" + unfolding is_arith_prog_on_def h_arith_prog + using range_h[of 0] range_h[of l] + by (auto simp: Max_ge[of "{a..b}"] Min_le[of "{a..b}"] + h_arith_prog arith_prog_def) + + have case_col: "h ` {.. col -` {j} \ {a..b}" + using h_col range_h unfolding j_def by auto + + show ?case using case_j case_step case_on case_col + by (auto simp: h_arith_prog) +qed + +text \ Lastly, we assemble all lemmas to finally prove Van der Waerden's Theorem by induction on +$l$. The cases $l=1$ and the induction start $l=2$ are treated separately and have been shown +earlier.\ +theorem van_der_Waerden: assumes "l>0" "k>0" shows "\n. vdw k l n" +using assms proof (induction l arbitrary: k rule: less_induct) + case (less l) + consider "l=1" | "l=2" | "l>2" using less.prems by linarith + then show ?case + proof (cases) + assume "l=1" + then show ?thesis using vdw_1_right by auto + next + assume "l=2" + then show ?thesis using vdw_2_right by auto + next + assume "l > 2" + then have "2\l-1" by auto + from less.IH[of "l-1"] \l>2\ + have "\k'. k'>0 \ \n. vdw k' (l-1) n" by auto + with vdw_imp_vdw_lemma[of "l-1" k k] \l-1\2\ \k>0\ + obtain N where "vdw_lemma k k (l-1) N" by auto + then have "vdw k l N" using vdw_lemma_imp_vdw[of k "l-1" N] + by (simp add: less.prems(1)) + then show ?thesis by auto + qed +qed + +end diff --git a/thys/Van_der_Waerden/document/root.bib b/thys/Van_der_Waerden/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Van_der_Waerden/document/root.bib @@ -0,0 +1,6 @@ +@techreport{Swan, + address = {Department of Mathematics, University of Chicago}, + author = {Richard G. Swan}, + note = {Online at \url{http://www.math.uchicago.edu/~swan/expo/vdW.pdf}}, + title = {Van der {Waerden's} Theorem on Arithmetic Progressions}, + url = {http://www.math.uchicago.edu/~swan/expo/vdW.pdf}} diff --git a/thys/Van_der_Waerden/document/root.tex b/thys/Van_der_Waerden/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Van_der_Waerden/document/root.tex @@ -0,0 +1,44 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amsfonts, amsmath, amssymb} + +% this should be the last package used +\usepackage{pdfsetup} +\usepackage[shortcuts]{extdash} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{rm} + + +\begin{document} + +\title{Van der Waerden's Theorem} +\author{Katharina Kreuzer, Manuel Eberl} +\maketitle + +\begin{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 Swan~\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$. +\end{abstract} + + +\newpage +\tableofcontents + +\newpage +\parindent 0pt\parskip 0.5ex + +\input{session} + +\nocite{Swan} + +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + diff --git a/web/entries/Epistemic_Logic.html b/web/entries/Epistemic_Logic.html --- a/web/entries/Epistemic_Logic.html +++ b/web/entries/Epistemic_Logic.html @@ -1,215 +1,217 @@ Epistemic Logic: Completeness of Modal Logics - Archive of Formal Proofs

 

 

 

 

 

 

Epistemic Logic: Completeness of Modal Logics

 

- + + +
Title: Epistemic Logic: Completeness of Modal Logics
Author: Asta Halkjær From
Submission date: 2018-10-29
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).
Change history: [2021-04-15]: Added completeness of modal logics T, KB, K4, S4 and S5.
BibTeX:
@article{Epistemic_Logic-AFP,
   author  = {Asta Halkjær From},
   title   = {Epistemic Logic: Completeness of Modal Logics},
   journal = {Archive of Formal Proofs},
   month   = oct,
   year    = 2018,
   note    = {\url{https://isa-afp.org/entries/Epistemic_Logic.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Used by:Public_Announcement_Logic

\ No newline at end of file diff --git a/web/entries/IMP_Compiler.html b/web/entries/IMP_Compiler.html new file mode 100644 --- /dev/null +++ b/web/entries/IMP_Compiler.html @@ -0,0 +1,204 @@ + + + + +A Shorter Compiler Correctness Proof for Language IMP - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

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

 

+

 

+
+
+

 

+

A + + Shorter + + Compiler + + Correctness + + Proof + + for + + Language + + IMP + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:A Shorter Compiler Correctness Proof for Language IMP
+ Author: + + Pasquale Noce (pasquale /dot/ noce /dot/ lavoro /at/ gmail /dot/ com) +
Submission date:2021-06-04
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.
BibTeX: +
@article{IMP_Compiler-AFP,
+  author  = {Pasquale Noce},
+  title   = {A Shorter Compiler Correctness Proof for Language IMP},
+  journal = {Archive of Formal Proofs},
+  month   = jun,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/IMP_Compiler.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/MiniSail.html b/web/entries/MiniSail.html new file mode 100644 --- /dev/null +++ b/web/entries/MiniSail.html @@ -0,0 +1,209 @@ + + + + +MiniSail - A kernel language for the ISA specification language SAIL - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

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

 

+

 

+
+
+

 

+

MiniSail + + - + + A + + kernel + + language + + for + + the + + ISA + + specification + + language + + SAIL + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:MiniSail - A kernel language for the ISA specification language SAIL
+ Author: + + Mark Wassell (mpwassell /at/ gmail /dot/ com) +
Submission date:2021-06-18
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.
BibTeX: +
@article{MiniSail-AFP,
+  author  = {Mark Wassell},
+  title   = {MiniSail - A kernel language for the ISA specification language SAIL},
+  journal = {Archive of Formal Proofs},
+  month   = jun,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/MiniSail.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
Depends on:Native_Word, Nominal2, Show
+ +

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

 

 

 

 

 

 

Native Word

 

- +
Title: Native Word
Author: Andreas Lochbihler
Contributor: Peter Lammich
Submission date: 2013-09-17
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.
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)
BibTeX:
@article{Native_Word-AFP,
   author  = {Andreas Lochbihler},
   title   = {Native Word},
   journal = {Archive of Formal Proofs},
   month   = sep,
   year    = 2013,
   note    = {\url{https://isa-afp.org/entries/Native_Word.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: Word_Lib
Used by:Collections, Datatype_Order_Generator, Iptables_Semantics, JinjaThreads, Mersenne_Primes, ROBDD, Separation_Logic_Imperative_HOL, WebAssembly
Collections, Datatype_Order_Generator, Iptables_Semantics, JinjaThreads, Mersenne_Primes, MiniSail, ROBDD, Separation_Logic_Imperative_HOL, WebAssembly

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

 

 

 

 

 

 

Nominal 2

 

- +
Title: Nominal 2
Authors: Christian Urban, Stefan Berghofer and Cezary Kaliszyk
Submission date: 2013-02-21
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.

BibTeX:
@article{Nominal2-AFP,
   author  = {Christian Urban and Stefan Berghofer and Cezary Kaliszyk},
   title   = {Nominal 2},
   journal = {Archive of Formal Proofs},
   month   = feb,
   year    = 2013,
   note    = {\url{https://isa-afp.org/entries/Nominal2.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: BSD License
Depends on: FinFun
Used by:Goedel_HFSet_Semanticless, Incompleteness, LambdaAuth, Launchbury, Modal_Logics_for_NTS, Rewriting_Z, Robinson_Arithmetic
Goedel_HFSet_Semanticless, Incompleteness, LambdaAuth, Launchbury, MiniSail, Modal_Logics_for_NTS, Rewriting_Z, Robinson_Arithmetic

\ No newline at end of file diff --git a/web/entries/Public_Announcement_Logic.html b/web/entries/Public_Announcement_Logic.html new file mode 100644 --- /dev/null +++ b/web/entries/Public_Announcement_Logic.html @@ -0,0 +1,187 @@ + + + + +Public Announcement Logic - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

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

 

+

 

+
+
+

 

+

Public + + Announcement + + Logic + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Public Announcement Logic
+ Author: + + Asta Halkjær From +
Submission date:2021-06-17
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.
BibTeX: +
@article{Public_Announcement_Logic-AFP,
+  author  = {Asta Halkjær From},
+  title   = {Public Announcement Logic},
+  journal = {Archive of Formal Proofs},
+  month   = jun,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/Public_Announcement_Logic.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
Depends on:Epistemic_Logic
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Show.html b/web/entries/Show.html --- a/web/entries/Show.html +++ b/web/entries/Show.html @@ -1,247 +1,247 @@ Haskell's Show Class in Isabelle/HOL - Archive of Formal Proofs

 

 

 

 

 

 

Haskell's Show Class in Isabelle/HOL

 

- +
Title: Haskell's Show Class in Isabelle/HOL
Authors: Christian Sternagel (c /dot/ sternagel /at/ gmail /dot/ com) and René Thiemann (rene /dot/ thiemann /at/ uibk /dot/ ac /dot/ at)
Submission date: 2014-07-29
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".
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".
BibTeX:
@article{Show-AFP,
   author  = {Christian Sternagel and René Thiemann},
   title   = {Haskell's Show Class in Isabelle/HOL},
   journal = {Archive of Formal Proofs},
   month   = jul,
   year    = 2014,
   note    = {\url{https://isa-afp.org/entries/Show.html},
             Formal proof development},
   ISSN    = {2150-914x},
 }
License: GNU Lesser General Public License (LGPL)
Depends on: Deriving
Used by:Affine_Arithmetic, AI_Planning_Languages_Semantics, CakeML, CakeML_Codegen, Certification_Monads, Dict_Construction, Modular_arithmetic_LLL_and_HNF_algorithms, Monad_Memo_DP, Polynomial_Factorization, Polynomials, Real_Impl, XML
Affine_Arithmetic, AI_Planning_Languages_Semantics, CakeML, CakeML_Codegen, Certification_Monads, Dict_Construction, MiniSail, Modular_arithmetic_LLL_and_HNF_algorithms, Monad_Memo_DP, Polynomial_Factorization, Polynomials, Real_Impl, XML

\ No newline at end of file diff --git a/web/entries/SpecCheck.html b/web/entries/SpecCheck.html new file mode 100644 --- /dev/null +++ b/web/entries/SpecCheck.html @@ -0,0 +1,199 @@ + + + + +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
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ No newline at end of file diff --git a/web/entries/Van_der_Waerden.html b/web/entries/Van_der_Waerden.html new file mode 100644 --- /dev/null +++ b/web/entries/Van_der_Waerden.html @@ -0,0 +1,194 @@ + + + + +Van der Waerden's Theorem - Archive of Formal Proofs + + + + + + + + + + + + + + + + + + + + + + + + +
+

 

+ + + +

 

+

 

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

 

+

 

+
+
+

 

+

Van + + der + + Waerden's + + Theorem + +

+

 

+ + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + + +
Title:Van der Waerden's Theorem
+ Authors: + + Katharina Kreuzer and + Manuel Eberl +
Submission date:2021-06-22
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$.
BibTeX: +
@article{Van_der_Waerden-AFP,
+  author  = {Katharina Kreuzer and Manuel Eberl},
+  title   = {Van der Waerden's Theorem},
+  journal = {Archive of Formal Proofs},
+  month   = jun,
+  year    = 2021,
+  note    = {\url{https://isa-afp.org/entries/Van_der_Waerden.html},
+            Formal proof development},
+  ISSN    = {2150-914x},
+}
+
License:BSD License
+ +

+ + + + + + + + + + + + + + + + + + +
+
+ + + + + + \ 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,5517 +1,5560 @@ 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-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
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 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
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 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
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 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 and Peter Lammich
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,569 +1,574 @@ 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. - 24 May 2021 00:00:00 +0000 + 01 Jul 2021 00:00:00 +0000 + + 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). Graph Lemma https://www.isa-afp.org/entries/Combinatorics_Words_Graph_Lemma.html https://www.isa-afp.org/entries/Combinatorics_Words_Graph_Lemma.html Štěpán Holub, Štěpán Starosta 24 May 2021 00:00:00 +0000 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. Combinatorics on Words Basics https://www.isa-afp.org/entries/Combinatorics_Words.html https://www.isa-afp.org/entries/Combinatorics_Words.html Štěpán Holub, Martin Raška, Štěpán Starosta 24 May 2021 00:00:00 +0000 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). Regression Test Selection https://www.isa-afp.org/entries/Regression_Test_Selection.html https://www.isa-afp.org/entries/Regression_Test_Selection.html Susannah Mansky 30 Apr 2021 00:00:00 +0000 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). Isabelle's Metalogic: Formalization and Proof Checker https://www.isa-afp.org/entries/Metalogic_ProofChecker.html https://www.isa-afp.org/entries/Metalogic_ProofChecker.html Tobias Nipkow, Simon Roßkopf 27 Apr 2021 00:00:00 +0000 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 <a href="https://arxiv.org/pdf/2104.12224.pdf">CADE 28 paper</a>. Lifting the Exponent https://www.isa-afp.org/entries/Lifting_the_Exponent.html https://www.isa-afp.org/entries/Lifting_the_Exponent.html Jakub Kądziołka 27 Apr 2021 00:00:00 +0000 We formalize the <i>Lifting the Exponent Lemma</i>, 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 <a href="https://s3.amazonaws.com/aops-cdn.artofproblemsolving.com/resources/articles/lifting-the-exponent.pdf">Amir Hossein Parvardi's</a>. The BKR Decision Procedure for Univariate Real Arithmetic https://www.isa-afp.org/entries/BenOr_Kozen_Reif.html https://www.isa-afp.org/entries/BenOr_Kozen_Reif.html Katherine Cordwell, Yong Kiam Tan, André Platzer 24 Apr 2021 00:00:00 +0000 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. Gale-Stewart Games https://www.isa-afp.org/entries/GaleStewart_Games.html https://www.isa-afp.org/entries/GaleStewart_Games.html Sebastiaan Joosten 23 Apr 2021 00:00:00 +0000 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. Formalization of Timely Dataflow's Progress Tracking Protocol https://www.isa-afp.org/entries/Progress_Tracking.html https://www.isa-afp.org/entries/Progress_Tracking.html Matthias Brun, Sára Decova, Andrea Lattuada, Dmitriy Traytel 13 Apr 2021 00:00:00 +0000 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 <a href="https://traytel.bitbucket.io/papers/itp21-progress_tracking/safe.pdf">ITP'21 paper</a>. Information Flow Control via Dependency Tracking https://www.isa-afp.org/entries/IFC_Tracking.html https://www.isa-afp.org/entries/IFC_Tracking.html Benedikt Nordhoff 01 Apr 2021 00:00:00 +0000 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's Schemes in Algebraic Geometry https://www.isa-afp.org/entries/Grothendieck_Schemes.html https://www.isa-afp.org/entries/Grothendieck_Schemes.html Anthony Bordg, Lawrence Paulson, Wenda Li 29 Mar 2021 00:00:00 +0000 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. Hensel's Lemma for the p-adic Integers https://www.isa-afp.org/entries/Padic_Ints.html https://www.isa-afp.org/entries/Padic_Ints.html Aaron Crighton 23 Mar 2021 00:00:00 +0000 We formalize the ring of <em>p</em>-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 <em>p</em>. 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 <em>p</em>-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. Constructive Cryptography in HOL: the Communication Modeling Aspect https://www.isa-afp.org/entries/Constructive_Cryptography_CM.html https://www.isa-afp.org/entries/Constructive_Cryptography_CM.html Andreas Lochbihler, S. Reza Sefidgar 17 Mar 2021 00:00:00 +0000 Constructive Cryptography (CC) [<a href="https://conference.iiis.tsinghua.edu.cn/ICS2011/content/papers/14.html">ICS 2011</a>, <a href="https://doi.org/10.1007/978-3-642-27375-9_3">TOSCA 2011</a>, <a href="https://doi.org/10.1007/978-3-662-53641-4_1">TCC 2016</a>] 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 [<a href="https://isa-afp.org/entries/Constructive_Cryptography.html">Constructive_Cryptography</a>, <a href="https://doi.org/10.1109/CSF.2019.00018">CSF 2019</a>] 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 [<a href="https://isa-afp.org/entries/CryptHOL.html">CryptHOL</a>, <a href="https://doi.org/10.1007/978-3-662-49498-1_20">ESOP 2016</a>, <a href="https://doi.org/10.1007/s00145-019-09341-z">J Cryptol 2020</a>]. This formalization is described in <a href="http://www.andreas-lochbihler.de/pub/basin2021.pdf">Abstract Modeling of Systems Communication in Constructive Cryptography using CryptHOL</a>. Two algorithms based on modular arithmetic: lattice basis reduction and Hermite normal form computation https://www.isa-afp.org/entries/Modular_arithmetic_LLL_and_HNF_algorithms.html https://www.isa-afp.org/entries/Modular_arithmetic_LLL_and_HNF_algorithms.html Ralph Bottesch, Jose Divasón, René Thiemann 12 Mar 2021 00:00:00 +0000 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. Quantum projective measurements and the CHSH inequality https://www.isa-afp.org/entries/Projective_Measurements.html https://www.isa-afp.org/entries/Projective_Measurements.html Mnacho Echenim 03 Mar 2021 00:00:00 +0000 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. The Hermite–Lindemann–Weierstraß Transcendence Theorem https://www.isa-afp.org/entries/Hermite_Lindemann.html https://www.isa-afp.org/entries/Hermite_Lindemann.html Manuel Eberl 03 Mar 2021 00:00:00 +0000 <p>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.</p> <p>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}$.</p> <p>Like the <a href="https://doi.org/10.1007/978-3-319-66107-0_5">previous formalisation in Coq by Bernard</a>, I proceeded by formalising <a href="https://doi.org/10.1017/CBO9780511565977">Baker's version of the theorem and proof</a> 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.</p> <p>This has a number of direct corollaries, e.g.:</p> <ul> <li>$e$ and $\pi$ are transcendental</li> <li>$e^z$, $\sin z$, $\tan z$, etc. are transcendental for algebraic $z\in\mathbb{C}\setminus\{0\}$</li> <li>$\ln z$ is transcendental for algebraic $z\in\mathbb{C}\setminus\{0, 1\}$</li> </ul> Mereology https://www.isa-afp.org/entries/Mereology.html https://www.isa-afp.org/entries/Mereology.html Ben Blumson 01 Mar 2021 00:00:00 +0000 We use Isabelle/HOL to verify elementary theorems and alternative axiomatizations of classical extensional mereology. The Sunflower Lemma of Erdős and Rado https://www.isa-afp.org/entries/Sunflowers.html https://www.isa-afp.org/entries/Sunflowers.html René Thiemann 25 Feb 2021 00:00:00 +0000 We formally define sunflowers and provide a formalization of the sunflower lemma of Erd&odblac;s and Rado: whenever a set of size-<i>k</i>-sets has a larger cardinality than <i>(r - 1)<sup>k</sup> &middot; k!</i>, then it contains a sunflower of cardinality <i>r</i>. A Verified Imperative Implementation of B-Trees https://www.isa-afp.org/entries/BTree.html https://www.isa-afp.org/entries/BTree.html Niels Mündler 24 Feb 2021 00:00:00 +0000 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 <a href="https://www.isa-afp.org/entries/Refine_Imperative_HOL.html"> Isabelle Refinement Framework </a> . 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 <a href="https://mediatum.ub.tum.de/1596550">Bachelor's Thesis</a>. Formal Puiseux Series https://www.isa-afp.org/entries/Formal_Puiseux_Series.html https://www.isa-afp.org/entries/Formal_Puiseux_Series.html Manuel Eberl 17 Feb 2021 00:00:00 +0000 <p>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 <em>N</em> is an integer and <em>d</em> is a positive integer.</p> <p>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.</p> The Laws of Large Numbers https://www.isa-afp.org/entries/Laws_of_Large_Numbers.html https://www.isa-afp.org/entries/Laws_of_Large_Numbers.html Manuel Eberl 10 Feb 2021 00:00:00 +0000 <p>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]$.</p> <p> 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:</p> <ul> <li>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$.</li> <li>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$.</li> </ul> <p>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.</p> Tarski's Parallel Postulate implies the 5th Postulate of Euclid, the Postulate of Playfair and the original Parallel Postulate of Euclid https://www.isa-afp.org/entries/IsaGeoCoq.html https://www.isa-afp.org/entries/IsaGeoCoq.html Roland Coghetto 31 Jan 2021 00:00:00 +0000 <p>The <a href="https://geocoq.github.io/GeoCoq/">GeoCoq library</a> 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. </p> <p>It should be noted that T.J.M. Makarios has done <a href="https://www.isa-afp.org/entries/Tarskis_Geometry.html">some proofs in Tarski's Geometry</a>. It uses a definition that does not quite coincide with the definition used in Geocoq and here. Furthermore, corresponding definitions in the <a href="https://www.isa-afp.org/entries/Poincare_Disc.html">Poincaré Disc Model development</a> are not identical to those defined in GeoCoq. </p> <p>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. </p> Solution to the xkcd Blue Eyes puzzle https://www.isa-afp.org/entries/Blue_Eyes.html https://www.isa-afp.org/entries/Blue_Eyes.html Jakub Kądziołka 30 Jan 2021 00:00:00 +0000 In a <a href="https://xkcd.com/blue_eyes.html">puzzle published by Randall Munroe</a>, 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. Hood-Melville Queue https://www.isa-afp.org/entries/Hood_Melville_Queue.html https://www.isa-afp.org/entries/Hood_Melville_Queue.html Alejandro Gómez-Londoño 18 Jan 2021 00:00:00 +0000 This is a verified implementation of a constant time queue. The original design is due to <a href="https://doi.org/10.1016/0020-0190(81)90030-2">Hood and Melville</a>. This formalization follows the presentation in <em>Purely Functional Data Structures</em>by Okasaki. JinjaDCI: a Java semantics with dynamic class initialization https://www.isa-afp.org/entries/JinjaDCI.html https://www.isa-afp.org/entries/JinjaDCI.html Susannah Mansky 11 Jan 2021 00:00:00 +0000 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). - - Cofinality and the Delta System Lemma - https://www.isa-afp.org/entries/Delta_System_Lemma.html - https://www.isa-afp.org/entries/Delta_System_Lemma.html - Pedro Sánchez Terraf - 27 Dec 2020 00:00:00 +0000 - -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. - - - Topological semantics for paraconsistent and paracomplete logics - https://www.isa-afp.org/entries/Topological_Semantics.html - https://www.isa-afp.org/entries/Topological_Semantics.html - David Fuenmayor - 17 Dec 2020 00:00:00 +0000 - -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 (<a -href="http://dx.doi.org/10.1007/s11787-012-0052-y">SSE</a>) -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. - - - Relational Minimum Spanning Tree Algorithms - https://www.isa-afp.org/entries/Relational_Minimum_Spanning_Trees.html - https://www.isa-afp.org/entries/Relational_Minimum_Spanning_Trees.html - Walter Guttmann, Nicolas Robinson-O'Brien - 08 Dec 2020 00:00:00 +0000 - -We verify the correctness of Prim's, Kruskal's and -Borůvka's minimum spanning tree algorithms based on algebras for -aggregation and minimisation. - - - Inline Caching and Unboxing Optimization for Interpreters - https://www.isa-afp.org/entries/Interpreter_Optimizations.html - https://www.isa-afp.org/entries/Interpreter_Optimizations.html - Martin Desharnais - 07 Dec 2020 00:00:00 +0000 - -This Isabelle/HOL formalization builds on the -<em>VeriComp</em> entry of the <em>Archive of Formal -Proofs</em> to provide the following contributions: <ul> -<li>an operational semantics for a realistic virtual machine -(Std) for dynamically typed programming languages;</li> -<li>the formalization of an inline caching optimization (Inca), -a proof of bisimulation with (Std), and a compilation -function;</li> <li>the formalization of an unboxing -optimization (Ubx), a proof of bisimulation with (Inca), and a simple -compilation function.</li> </ul> This formalization was -described in the CPP 2021 paper <em>Towards Efficient and -Verified Virtual Machines for Dynamic Languages</em> - - - The Relational Method with Message Anonymity for the Verification of Cryptographic Protocols - https://www.isa-afp.org/entries/Relational_Method.html - https://www.isa-afp.org/entries/Relational_Method.html - Pasquale Noce - 05 Dec 2020 00:00:00 +0000 - -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. - diff --git a/web/topics.html b/web/topics.html --- a/web/topics.html +++ b/web/topics.html @@ -1,963 +1,968 @@ 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   Optimization: Simplex   Quantum computing: Isabelle_Marries_Dirac   Projective_Measurements  

Concurrency

Data structures

Functional programming

Hardware

SPARCv8  

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   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   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