Isabelle Build Manager

Diff: presentation/320

back to build

Isabelle:

o  changeset:   82687:97b9ac57751e
|  tag:         tip
|  user:        haftmann
|  date:        Wed Jun 04 19:43:13 2025 +0000
|  summary:     some more lemmas
|
o  changeset:   82686:e6f299c5dec6
|  user:        haftmann
|  date:        Wed Jun 04 19:43:13 2025 +0000
|  summary:     tuned syntax
|
o  changeset:   82685:8575092e21fa
|  user:        nipkow
|  date:        Wed Jun 04 09:52:40 2025 +0200
|  summary:     latex error
|
diff -r a6cfe84d0ddd -r 97b9ac57751e src/HOL/Library/Disjoint_Sets.thy
--- src/HOL/Library/Disjoint_Sets.thy
+++ src/HOL/Library/Disjoint_Sets.thy
@@ -472,7 +472,7 @@
   If $h$ is an involution on $X$ with no fixed points in $X$ and
   $f(h(x)) = -f(x)$ then $\sum_{x\in X} f(x) = 0$.
   
-  This is easy to show in a ring with characteristic not equal to $2", since then we can do
+  This is easy to show in a ring with characteristic not equal to $2$, since then we can do
   \[\sum_{x\in X} f(x) = \sum_{x\in X} f(h(x)) = -\sum_{x\in X} f(x)\]
   and therefore $2 \sum_{x\in X} f(x) = 0$.
 
diff -r a6cfe84d0ddd -r 97b9ac57751e src/HOL/Library/Word.thy
--- src/HOL/Library/Word.thy
+++ src/HOL/Library/Word.thy
@@ -3180,6 +3180,40 @@
   for i m :: "'a::len word"
   by uint_arith
 
+lemma less_imp_less_eq_dec:
+  \<open>v \<le> w - 1\<close> if \<open>v < w\<close> for v w :: \<open>'a::len word\<close>
+using that proof transfer
+  show \<open>take_bit LENGTH('a) k \<le> take_bit LENGTH('a) (l - 1)\<close>
+    if \<open>take_bit LENGTH('a) k < take_bit LENGTH('a) l\<close>
+    for k l :: int
+    using that by (cases \<open>take_bit LENGTH('a) l = 0\<close>)
+      (auto simp add: take_bit_decr_eq)
+qed
+
+lemma inc_less_eq_triv_imp:
+  \<open>w = - 1\<close> if \<open>w + 1 \<le> w\<close> for w :: \<open>'a::len word\<close>
+proof (rule ccontr)
+  assume \<open>w \<noteq> - 1\<close>
+  with that show False
+    by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)
+qed
+
+lemma less_eq_dec_triv_imp:
+  \<open>w = 0\<close> if \<open>w \<le> w - 1\<close> for w :: \<open>'a::len word\<close>
+proof (rule ccontr)
+  assume \<open>w \<noteq> 0\<close>
+  with that show False
+    by transfer (auto simp add: take_bit_eq_mask_iff dest: take_bit_decr_eq)
+qed
+
+lemma inc_less_eq_iff:
+  \<open>v + 1 \<le> w \<longleftrightarrow> v = - 1 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>
+  by (auto intro: inc_less_eq_triv_imp inc_le)
+
+lemma less_eq_dec_iff:
+  \<open>v \<le> w - 1 \<longleftrightarrow> w = 0 \<or> v < w\<close> for v w :: \<open>'a::len word\<close>
+  by (auto intro: less_eq_dec_triv_imp less_imp_less_eq_dec)
+
 lemma inc_i: "1 \<le> i \<Longrightarrow> i < m \<Longrightarrow> 1 \<le> i + 1 \<and> i + 1 \<le> m"
   for i m :: "'a::len word"
   by uint_arith
diff -r a6cfe84d0ddd -r 97b9ac57751e src/HOL/List.thy
--- src/HOL/List.thy
+++ src/HOL/List.thy
@@ -8279,8 +8279,8 @@
   by simp
 
 lemma subset_code [code]:
-  "set xs \<le> B \<longleftrightarrow> (\<forall>x\<in>set xs. x \<in> B)"
-  "A \<le> List.coset ys \<longleftrightarrow> (\<forall>y\<in>set ys. y \<notin> A)"
+  "set xs \<subseteq> B \<longleftrightarrow> (\<forall>x\<in>set xs. x \<in> B)"
+  "A \<subseteq> List.coset ys \<longleftrightarrow> (\<forall>y\<in>set ys. y \<notin> A)"
   "List.coset [] \<subseteq> set [] \<longleftrightarrow> False"
   by auto
 

AFP:

o    changeset:   15773:29846bc89de2
|\   tag:         tip
| |  parent:      15771:8c1f704fc41d
| |  parent:      15772:87ce4ad4ee2c
| |  user:        nipkow
| |  date:        Thu Jun 05 18:20:43 2025 +0200
| |  summary:     merged
| |
| o  changeset:   15772:87ce4ad4ee2c
| |  parent:      15770:df2f317f3654
| |  user:        nipkow
| |  date:        Thu Jun 05 18:20:20 2025 +0200
| |  summary:     generalized Finite_Automata_HF (partly) from hf to 's;
| |
o |  changeset:   15771:8c1f704fc41d
|/   user:        paulson <lp15@cam.ac.uk>
|    date:        Thu Jun 05 12:18:57 2025 +0100
|    summary:     Removal of library material that had already been added to the distribution
|
o  changeset:   15770:df2f317f3654
|  user:        nipkow
|  date:        Wed Jun 04 17:19:41 2025 +0200
|  summary:     adapted to generalized Finite_Automata_HF
|
o  changeset:   15769:e78b2e0058e2
|  user:        nipkow
|  date:        Wed Jun 04 17:07:34 2025 +0200
|  summary:     adapted to devel
|
o  changeset:   15768:044c5bc0cc1b
|  user:        nipkow
|  date:        Wed Jun 04 12:17:25 2025 +0200
|  summary:     Adapted to generalization of hf dfa to 's dfa.
|
o  changeset:   15767:9a8cee107210
|  user:        nipkow
|  date:        Wed Jun 04 11:53:48 2025 +0200
|  summary:     adapted to devel
|
o  changeset:   15766:9e266deb69b4
|  user:        nipkow
|  date:        Wed Jun 04 11:47:44 2025 +0200
|  summary:     generalized parts of the theory: from hf dfa to 's dfa.
|
o  changeset:   15765:8a6131d91c0f
|  user:        nipkow
|  date:        Wed Jun 04 10:36:01 2025 +0200
|  summary:     got rid of smt proof (that caused a timeout)
|
o  changeset:   15764:1fc2c36d4f4b
|  user:        nipkow
|  date:        Wed Jun 04 10:18:25 2025 +0200
|  summary:     adapted to devel
|
o  changeset:   15763:a1c1276b248e
|  user:        nipkow
|  date:        Wed Jun 04 10:13:05 2025 +0200
|  summary:     adapted to change in devel
|
o    changeset:   15762:5d7573bb1dc0
|\   parent:      15740:848f2c00e57f
| |  parent:      15761:b08f728e28b0
| |  user:        nipkow
| |  date:        Wed Jun 04 08:44:18 2025 +0200
| |  summary:     merge from AFP 2025
| |
diff -r 848f2c00e57f -r 29846bc89de2 doc/editors/new-entry-checkin.md
--- doc/editors/new-entry-checkin.md
+++ doc/editors/new-entry-checkin.md
@@ -71,7 +71,7 @@
 10. finally, when you are happy with everything, `hg push` all changes
     to Heptapod. The publish script will refuse to publish if the
     changes aren't pushed.
-11. to publish the changes to the web, run
+11. to publish the changes to the web, run (with **all sessions** of the entry)
 
          ../admin/publish <sessions...>
 
diff -r 848f2c00e57f -r 29846bc89de2 metadata/authors.toml
--- metadata/authors.toml
+++ metadata/authors.toml
@@ -288,6 +288,7 @@
 
 [ballenghien]
 name = "Benoît Ballenghien"
+orcid = "0009-0000-4941-187X"
 
 [ballenghien.emails]
 
@@ -1374,7 +1375,8 @@
 
 [coghetto.emails.coghetto_email]
 user = [
-  "roland.coghetto",
+  "roland",
+  "coghetto",
 ]
 host = [
   "cafr-msa2p",
@@ -2528,6 +2530,24 @@
 [grov.homepages]
 grov_homepage = "http://homepages.inf.ed.ac.uk/ggrov"
 
+[gschossmann]
+name = "Markus Gschoßmann"
+orcid = "0009-0005-5058-5806"
+
+[gschossmann.emails]
+
+[gschossmann.emails.gschossmann_email]
+user = [
+  "markus",
+  "gschossmann",
+]
+host = [
+  "tum",
+  "de",
+]
+
+[gschossmann.homepages]
+
 [guerraoui]
 name = "Rachid Guerraoui"
 
@@ -3788,6 +3808,23 @@
 [krauss.homepages]
 krauss_homepage = "http://www.in.tum.de/~krauss"
 
+[krayer]
+name = "Felix Krayer"
+
+[krayer.emails]
+
+[krayer.emails.krayer_email]
+user = [
+  "felix",
+  "krayer",
+]
+host = [
+  "tum",
+  "de",
+]
+
+[krayer.homepages]
+
 [kreuzer]
 name = "Katharina Kreuzer"
 
@@ -4062,6 +4099,23 @@
 
 [leek.homepages]
 
+[lehr]
+name = "Fabian Lehr"
+
+[lehr.emails]
+
+[lehr.emails.lehr_email]
+user = [
+  "fabian",
+  "lehr",
+]
+host = [
+  "tum",
+  "de",
+]
+
+[lehr.homepages]
+
 [leustean]
 name = "Laurentiu Leustean"
 
@@ -5556,6 +5610,7 @@
 
 [paulson]
 name = "Lawrence C. Paulson"
+orcid = "0000-0003-0288-4279"
 
 [paulson.emails]
 
@@ -5615,6 +5670,22 @@
 [petrovic.homepages]
 petrovic_homepage = "http://www.matf.bg.ac.rs/~danijela"
 
+[philipp]
+name = "Bruno Philipp"
+
+[philipp.emails]
+
+[philipp.emails.philipp_email]
+user = [
+  "ge95fek",
+]
+host = [
+  "mytum",
+  "de",
+]
+
+[philipp.homepages]
+
 [pierzchalski]
 name = "Edward Pierzchalski"
 
@@ -5799,6 +5870,7 @@
 
 [puyobro]
 name = "Benjamin Puyobro"
+orcid = "0009-0006-0294-9043"
 
 [puyobro.emails]
 
@@ -7138,6 +7210,13 @@
 [stevens.homepages]
 stevens_homepage = "https://www21.in.tum.de/team/stevensl"
 
+[stimpfle]
+name = "August Martin Stimpfle"
+
+[stimpfle.emails]
+
+[stimpfle.homepages]
+
 [stock]
 name = "Benedikt Stock"
 
@@ -7473,6 +7552,22 @@
 
 [tasch.homepages]
 
+[taskin]
+name = "Kaan Taskin"
+
+[taskin.emails]
+
+[taskin.emails.taskin_email]
+user = [
+  "kaantaskin2022",
+]
+host = [
+  "gmail",
+  "com",
+]
+
+[taskin.homepages]
+
 [taylor]
 name = "Ramsay G. Taylor"
 
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Context_Free_Grammar.toml
--- /dev/null
+++ metadata/entries/Context_Free_Grammar.toml
@@ -0,0 +1,49 @@
+title = "Context-Free Grammars and Languages"
+date = 2025-05-21
+topics = [
+  "Computer science/Automata and formal languages",
+]
+abstract = """
+This is a basic library of definitions and results about context-free grammars and languages.
+It includes context-free grammars and languages, parse trees, Chomsky normal form,
+pumping lemmas and the relationship of right-linear grammars to finite automata."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.nipkow]
+email = "nipkow_email"
+
+[authors.gschossmann]
+email = "gschossmann_email"
+
+[authors.krayer]
+email = "krayer_email"
+
+[authors.lehr]
+email = "lehr_email"
+
+[authors.philipp]
+email = "philipp_email"
+
+[authors.stimpfle]
+
+[authors.taskin]
+email = "taskin_email"
+
+[authors.yamada]
+email = "yamada_email"
+
+[contributors]
+
+[notify]
+nipkow = "nipkow_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Elementary_Ultrametric_Spaces.toml
--- /dev/null
+++ metadata/entries/Elementary_Ultrametric_Spaces.toml
@@ -0,0 +1,46 @@
+title = "Definition and Elementary Properties of Ultrametric Spaces"
+date = 2025-04-10
+topics = [
+  "Mathematics/Topology",
+]
+abstract = """
+An ultrametric space is a metric space in which the triangle inequality is strengthened
+by using the maximum instead of the sum. More formally, such a space is equipped with
+a real-valued application $dist$, called distance, verifying the four following conditions.
+\\begin{align*}
+    & dist \\ x \\ y \\ge 0\\\\
+    & dist \\ x \\ y = dist \\ y \\ x\\\\
+    & dist \\ x \\ y = 0 \\longleftrightarrow x = y\\\\
+    & dist \\ x \\ z \\le max \\ (dist \\ x \\ y) \\ (dist \\ y \\ z)
+\\end{align*}
+In this entry, we present an elementary formalization of these spaces relying on axiomatic type classes.
+The connection with standard metric spaces is obtained through a subclass relationship,
+and fundamental properties of ultrametric spaces are formally established."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[authors.puyobro]
+email = "puyobro_email"
+
+[authors.wolff]
+email = "wolff_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+puyobro = "puyobro_email"
+wolff = "wolff_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Farey_Sequences.toml
--- /dev/null
+++ metadata/entries/Farey_Sequences.toml
@@ -0,0 +1,39 @@
+title = "Farey Sequences and Ford Circles"
+date = 2025-05-15
+topics = [
+  "Mathematics/Number theory",
+]
+abstract = """
+The sequence $F_n$ of Farey fractions of order $n$ 
+has the form 
+$$\\frac{0}{1}, \\frac{1}{n}, \\frac{1}{n-1}, \\ldots, \\frac{n-1}{n}, \\frac{1}{1}$$
+where the fractions appear in numerical order and have denominators at most $n$.
+The transformation from $F_n$ to $F_{n+1}$ can be effected by combining adjacent elements of 
+the sequence~$F_n$, using an operation called the mediant.
+Adjacent (reduced) fractions $(a/b) < (c/d)$ satisfy the unimodular
+relation $bc - ad = 1$ and their mediant is $\\frac{a+c}{b+d}$.
+A Ford circle is specified by a rational number, and interesting consequences follow
+in the case of Ford circles obtained from some Farey sequence~$F_n$.
+The formalised material is drawn from Apostol's Modular Functions and Dirichlet Series in Number Theory."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.paulson]
+email = "paulson_email"
+
+[contributors]
+
+[notify]
+paulson = "paulson_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = [
+  "Modular Functions and Dirichlet Series in Number Theory",
+]
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/HOL-CSP_RS.toml
--- /dev/null
+++ metadata/entries/HOL-CSP_RS.toml
@@ -0,0 +1,31 @@
+title = "CSP Semantics over Restriction Spaces"
+date = 2025-05-07
+topics = [
+  "Computer science/Semantics and reasoning",
+  "Computer science/Concurrency/Process calculi",
+]
+abstract = "We use the Restriction_Spaces library as a semantic foundation for the process algebra framework HOL-CSP, offering a complementary backend to the existing HOLCF infrastructure. The type of processes is instantiated as a restriction space, and we prove that it is complete in this setting. This enables the construction of fixed points for recursive process definitions without having to rely exclusively on a pointed complete partial order. Notably, some operators are constructive without being Scott-continuous, and vice versa, illustrating the genuine complementarity between the two approaches. We also show that key CSP operators are either constructive or non-destructive, and verify the admissibility of several predicates, thereby supporting automated reasoning over recursive specifications."
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[authors.wolff]
+email = "wolff_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+wolff = "wolff_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Munta_Model_Checker.toml
--- /dev/null
+++ metadata/entries/Munta_Model_Checker.toml
@@ -0,0 +1,32 @@
+title = "Munta: A Verified Model Checker for Timed Automata"
+date = 2025-05-22
+topics = [
+  "Computer science/Automata and formal languages",
+  "Tools",
+]
+abstract = "Munta is a verified model checker for timed automata. It has been described in detail in a PhD thesis [3] and conference publications [4, 2]. This entry supersedes earlier versions of Munta hosted on GitHub."
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.wimmer]
+email = "wimmer_email"
+
+[contributors]
+
+[notify]
+wimmer = "wimmer_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = [
+  "10.1007/978-3-319-89960-2_4",
+  "10.1007/978-3-030-29662-9_14",
+]
+pubs = [
+  "https://nbn-resolving.org/urn:nbn:de:bvb:91-diss-20201209-1576100-1-0",
+]
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Restriction_Spaces-Examples.toml
--- /dev/null
+++ metadata/entries/Restriction_Spaces-Examples.toml
@@ -0,0 +1,42 @@
+title = "Examples of Restriction Spaces"
+date = 2025-05-07
+topics = [
+  "Mathematics/Analysis",
+  "Mathematics/Topology",
+  "Mathematics/Algebra",
+  "Computer science/Concurrency/Process calculi",
+]
+abstract = """
+In this session, a number of examples are provided to illustrate how the Restriction_Spaces library works.
+The simple cases are, of course, covered: trivial restriction, booleans, integers, option type, and so on.
+But we also explore more elaborate constructions, such as formal power series and a trace model of the CSP process algebra.
+Additionally, we provide a lightweight integration with HOLCF,
+equipping restriction spaces with the inherited partial order structure when needed."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[authors.puyobro]
+email = "puyobro_email"
+
+[authors.wolff]
+email = "wolff_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+puyobro = "puyobro_email"
+wolff = "wolff_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Restriction_Spaces-Ultrametric.toml
--- /dev/null
+++ metadata/entries/Restriction_Spaces-Ultrametric.toml
@@ -0,0 +1,47 @@
+title = "Ultrametric Structure for Restriction Spaces"
+date = 2025-05-07
+topics = [
+  "Mathematics/Analysis",
+  "Mathematics/Topology",
+  "Computer science/Semantics and reasoning",
+]
+abstract = """
+In this entry, we explore the relationship
+  between restriction spaces and usual metric
+  structures by instantiating the former as ultrametric spaces.
+  This is classically captured by defining the distance as
+  $$\\mathrm{dist} \\ x \\ y = \\inf_{x \\downarrow n \\ = \\ y \\downarrow n} \\left(\\frac{1}{2}\\right)^ n$$
+  but we actually generalize this perspective by introducing a hierarchy of increasingly
+  refined type classes to systematically relate ultrametric and restriction-based notions.
+  This layered approach enables a precise comparison of structural and topological properties.
+  In the end, our main result establishes that completeness in the sense of restriction spaces
+  coincides with standard metric completeness, thus bridging the gap between
+  our framework and Banach's fixed-point theorem established in HOL-Analysis."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[authors.puyobro]
+email = "puyobro_email"
+
+[authors.wolff]
+email = "wolff_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+puyobro = "puyobro_email"
+wolff = "wolff_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Restriction_Spaces.toml
--- /dev/null
+++ metadata/entries/Restriction_Spaces.toml
@@ -0,0 +1,58 @@
+title = "Restriction Spaces: a Fixed-Point Theory"
+date = 2025-05-07
+topics = [
+  "Tools",
+  "Mathematics/Analysis",
+  "Mathematics/Topology",
+  "Computer science/Semantics and reasoning",
+]
+abstract = """
+Fixed-point constructions are fundamental to defining recursive and co-recursive functions.
+  However, a general axiom $Y f = f (Y f)$ leads to inconsistency, and definitions
+  must therefore be based on theories guaranteeing existence under suitable conditions.
+  In Isabelle/HOL, these constructions are typically based on sets, well-founded orders
+  or domain-theoretic models such as for example HOLCF.
+  In this submission we introduce a formalization of restriction spaces i.e. spaces
+  equipped with a so-called restriction, denoted by $\\downarrow$,
+  satifying three properties:
+  \\begin{align*}
+    & x \\downarrow 0 = y \\downarrow 0\\\\
+    & x \\downarrow n \\downarrow m = x \\downarrow \\mathrm{min} \\ n \\ m\\\\
+    & x \\not = y \\Longrightarrow \\exists n. \\ x \\downarrow n \\not = y \\downarrow n
+  \\end{align*}
+  They turn out to be cartesian closed and admit natural notions of constructiveness and
+  completeness, enabling the definition of a fixed-point operator under verifiable side-conditions.
+  This is achieved in our entry, from topological definitions to induction principles.
+  Additionally, we configure the simplifier so that it can automatically solve both
+  constructiveness and admissibility subgoals, as long as users write higher-order rules
+  for their operators.
+  Since our implementation relies on axiomatic type classes, the resulting library
+  is a fully abstract, flexible and reusable framework."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[authors.puyobro]
+email = "puyobro_email"
+
+[authors.wolff]
+email = "wolff_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+puyobro = "puyobro_email"
+wolff = "wolff_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Shallow_Expressions.toml
--- /dev/null
+++ metadata/entries/Shallow_Expressions.toml
@@ -0,0 +1,30 @@
+title = "Shallow Expressions"
+date = 2025-05-16
+topics = [
+  "Computer science/Semantics and reasoning",
+  "Computer science/Programming languages/Logics",
+]
+abstract = "Most verification techniques use expressions, for example when assigning to variables or forming assertions over the state. Deep embeddings provide such expressions using a datatype, which allows queries over the syntax, such as calculating the free variables, and performing substitutions. Shallow embeddings, in contrast, model expressions as query functions over the state type, and are more amenable to automating verification tasks. The goal of this library is provide an intuitive implementation of shallow expressions, which nevertheless provides many of the benefits of a deep embedding. We harness the Optics library to provide an algebraic semantics for state variables, and use syntax translations to provide an intuitive lifted expression syntax. Furthermore, we provide a variety of meta-logic-style queries on expressions, such as dependencies on a state variable, and substitution of a variable for an expression. We also provide proof methods, based on the simplifier, to automate the associated proof tasks."
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.fosters]
+email = "fosters_email"
+
+[contributors]
+
+[notify]
+fosters = "fosters_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = [
+  "10.1007/s10817-024-09709-2",
+  "10.1016/j.scico.2020.102510",
+]
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 metadata/entries/Sophie_Germain.toml
--- /dev/null
+++ metadata/entries/Sophie_Germain.toml
@@ -0,0 +1,30 @@
+title = "Sophie Germain’s Theorem"
+date = 2025-04-09
+topics = [
+  "Mathematics/Number theory",
+]
+abstract = """
+Sophie Germain's theorem states that if $p$ is a prime number such that $2 p + 1$ is also prime (with $p$ then called a Sophie Germain prime), a weak version of Fermat’s Last Theorem (FLT) holds for the exponent $p$. That is, there exist no nontrivial integer solutions $x, y, z \\in \\mathbb{Z}$ to the equation
+$$x ^ p + y ^ p = z ^ p$$
+such that $p$ divides neither $x$, $y$ nor $z$.
+After some preliminary results, mainly on sufficient criteria for FLT, we propose a formalization of the classical version of Sophie Germain's theorem as well as a variant involving the concept of auxiliary primes that generalizes Sophie Germain primes."""
+license = "bsd"
+note = ""
+
+[authors]
+
+[authors.ballenghien]
+email = "ballenghien_email"
+
+[contributors]
+
+[notify]
+ballenghien = "ballenghien_email"
+
+[history]
+
+[extra]
+
+[related]
+dois = []
+pubs = []
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/AnBnCn_not_CFL.thy
--- /dev/null
+++ thys/Context_Free_Grammar/AnBnCn_not_CFL.thy
@@ -0,0 +1,248 @@
+(* 
+Author:     Bruno Philipp, TU München
+Author:     Tobias Nipkow
+*)
+
+section \<open>\<open>anbncn\<close> is Not Context-Free\<close>
+
+theory AnBnCn_not_CFL
+imports Pumping_Lemma_CFG
+begin                           
+
+text \<open>This theory proves that the language @{term "\<Union>n. {[a]^^n @ [b]^^n @ [c]^^n}"}
+is not context-free using the Pumping lemma.
+
+The formal proof follows the textbook proof (e.g.\ cite\<open>HopcroftMU\<close>) closely.
+The Isabelle proof is about 10\% of the length of the Coq proof by Ramos \emph{et al.} cite\<open>RamosAMQ and RamosGithub\<close>.
+The latter proof suffers from excessive case analyses.\<close>
+
+declare count_list_pow_list[simp]
+
+context
+  fixes a b c
+  assumes neq: "a\<noteq>b" "b\<noteq>c" "c\<noteq>a"
+begin
+
+lemma  c_greater_count0:
+  assumes "x@y = [a]^^n @ [b]^^n @ [c]^^n" "length y\<ge>n"
+  shows "count_list x c = 0"
+  using assms proof -
+   have "drop (2*n) (x@y) = [c]^^n" using assms
+     by simp
+   then have count_c_end: "count_list (drop (2*n) (x@y)) c = n"
+     by (simp)
+   have "count_list (x@y) c= n" using assms neq
+     by (simp)
+   then have count_c_front: "count_list (take (2*n) (x@y)) c = 0"
+     using count_c_end by (metis add_cancel_left_left append_take_drop_id count_list_append)
+   have  "\<exists>i. length y = n+i" using assms
+     by presburger
+   then obtain i where i: "length y= n+i"
+     by blast
+   then have "x = take (3*n-n-i) (x@y)"
+   proof -
+     have "x = take (2 * n - i) x @ take (2 * n - (i + length x)) y"
+       using i by (metis add.commute add_diff_cancel_left' append_eq_conv_conj assms(1) diff_diff_left
+         length_append length_pow_list_single mult_2 take_append)
+     then show ?thesis
+       by (simp)
+   qed
+   then have "x = take (3*n-n-i) (take (3*n-n) (x@y))"
+     by (metis diff_le_self min_def take_take)
+   then have "x = take (3*n-n-i) (take (2*n) (x@y))" 
+     by fastforce  
+  then show ?thesis using count_c_front count_list_0_iff in_set_takeD
+    by metis
+qed
+
+lemma  a_greater_count0:
+  assumes "x@y = [a]^^n @ [b]^^n @ [c]^^n" "length x\<ge>n"
+  shows "count_list y a = 0"
+  text \<open>this prof is easier than @{thm c_greater_count0} since a is at the start of the word rather than at the end \<close>
+proof -
+  have count_whole: "count_list (x@y) a = n"
+    using assms neq by auto
+  have take_n: "take n (x@y) = [a]^^n"
+    using assms by simp
+  then have count_take_n: "count_list (take n (x@y)) a = n"
+    by (simp)
+  have "\<exists> z. x = take n (x@y) @ z" 
+    by (metis append_eq_conv_conj assms(2) nat_le_iff_add take_add)
+  then have  count_a_x:"count_list x a = n" using count_take_n take_n count_whole 
+    by (metis add_diff_cancel_left' append.right_neutral count_list_append diff_add_zero)
+  have "count_list (x@y) a = n"
+    using assms neq by simp
+   then have "count_list y a = 0"
+    using count_a_x by simp
+  then show ?thesis
+    by presburger
+qed
+
+lemma a_or_b_zero:
+  assumes "u@w@y = [a]^^n @ [b]^^n @ [c]^^n" "length w \<le> n" (* neq not used *)
+  shows "count_list w a = 0 \<or> count_list w c = 0"
+  text \<open>This lemma uses @{term "count_list w a = 0 \<or> count_list w c = 0"} similar to all following proofs, focusing on the number of \<open>a\<close> and \<open>c\<close> found in \<open>w\<close> rather than the concrete structure.
+        It is also the merge of the two previous lemmas to make the final proof shorter\<close>
+proof-
+  show ?thesis proof (cases "length u <n")
+    case True 
+    have "length (u@w@y) = 3*n"  using assms 
+      by simp  
+    then have "length u + length w + length y = 3*n" 
+      by simp
+    then have "length u + length y \<ge> 2*n" using assms 
+      by linarith
+    then have u_or_y: "length u \<ge> n \<or> length y \<ge> n" 
+      by linarith
+    then  have "length y\<ge>n" using True  
+      by simp
+    then show ?thesis using c_greater_count0[of "u@w" y n] neq assms 
+      by simp
+  next
+    case False
+    then have "length u \<ge> n" 
+      by simp
+    then show ?thesis using a_greater_count0[of u "w@y" n ] neq assms 
+      by auto
+  qed
+qed
+
+lemma count_vx_not_zero:
+  assumes "u@v@w@x@y = [a]^^n @ [b]^^n @ [c]^^n" "v@x \<noteq> []"
+  shows "count_list (v@x) a \<noteq> 0 \<or> count_list (v@x) b \<noteq> 0 \<or> count_list (v@x) c \<noteq> 0"
+proof -
+  have set: "set ([a]^^n @ [b]^^n @ [c]^^n) = {a,b,c}" using assms pow_list_single_Nil_iff
+    by (fastforce simp add: pow_list_single)
+  show ?thesis proof (cases  "v\<noteq>[]")
+    case True
+    then have "\<exists>d\<in>set([a]^^n @ [b]^^n @ [c]^^n). count_list v d \<noteq> 0"
+      using assms(1)
+      by (metis append_Cons count_list_0_iff in_set_conv_decomp list.exhaust list.set_intros(1))
+    then have "count_list v a \<noteq> 0 \<or> count_list v b \<noteq> 0 \<or> count_list v c \<noteq>0"
+      using set by simp
+    then show ?thesis 
+      by force
+  next
+    case False
+    then have "x\<noteq>[]" using assms 
+      by fast
+    then have "\<exists>d\<in>set ([a]^^n @ [b]^^n @ [c]^^n). count_list x d \<noteq> 0"
+    proof -
+      have "\<exists>d\<in>set ([a] ^^ n) \<union> (set ([b] ^^ n) \<union> set ([c] ^^ n)). ya \<noteq> d \<longrightarrow> 0 < count_list ys d"
+        if "u @ v @ w @ ya # ys @ y = [a] ^^ n @ [b] ^^ n @ [c] ^^ n"
+          and "x = ya # ys"
+        for ya :: 'a
+          and ys :: "'a list"
+        using that by (metis Un_iff in_set_conv_decomp set_append)
+      then show ?thesis
+        using assms(1) \<open>x \<noteq> []\<close> by (auto simp: neq_Nil_conv)
+    qed
+    then have "count_list x a \<noteq> 0 \<or> count_list x b \<noteq> 0 \<or> count_list x c \<noteq> 0"
+       using set by simp
+    then show ?thesis 
+       by force
+   qed
+qed
+
+lemma not_ex_y_count:
+  assumes "i\<noteq>k \<or> k\<noteq>j \<or> i\<noteq>j" "count_list w a = i" "count_list w b = k" "count_list w c = j"
+  shows "\<not>(EX y. w = [a]^^y @ [b]^^y @ [c]^^y)"
+ proof 
+   assume "EX y. w = [a]^^y @ [b]^^y @ [c]^^y"
+   then obtain y where y: "w = [a]^^y @ [b]^^y @ [c]^^y" 
+     by blast
+   then have "count_list w a = y" using neq 
+     by simp
+   then have i_eq_y: "i=y" using assms 
+     by argo
+   then have "count_list w b = y"
+     using neq assms(2) y by (auto)
+   then have k_eq_y: "k=y" using assms 
+     by argo
+   have "count_list w c = y" using neq y
+     by simp
+   then have j_eq_y: "j=y" using assms  
+     by argo
+   show False  using i_eq_y k_eq_y j_eq_y assms 
+      by presburger
+ qed
+
+lemma not_in_count:
+  assumes (* no neq *) "count_list w a \<noteq> count_list w b \<or> count_list w b \<noteq> count_list w c \<or> count_list w c \<noteq> count_list w a"
+  shows "w \<notin> {word. \<exists> n.  word = [a]^^n @ [b]^^n @ [c]^^n}"
+  text \<open>This definition of a word not in the language is useful as it allows us to prove a word is not in the language just by knowing the number of each letter in a word\<close>
+  using assms not_ex_y_count
+  by (smt (verit, del_insts) mem_Collect_eq)
+
+text \<open>This is the central and only case analysis, which follows the textbook proofs.
+The Coq proof by Ramos is considerably more involved and ends up with a total of 24 cases\<close>
+
+lemma pumping_application:
+  assumes "u@v@w@x@y = [a]^^n @ [b]^^n @ [c]^^n"
+  and "count_list (v@w@x) a = 0 \<or> count_list (v@w@x) c = 0" and "v@x\<noteq>[]"
+  shows "u@w@y \<notin> (\<Union>n. {[a]^^n @ [b]^^n @ [c]^^n})"
+  text \<open>In this lemma it is proven that a word @{term "u @ v^^0 @ w @ x^^0 @ y"}
+        is not in the language @{term "\<Union>n. {[a]^^n @ [b]^^n @ [c]^^n}"}
+        as this is the easiest counterexample useful for the Pumping lemma\<close>
+proof-
+  have count_word_a: "count_list (u@v@w@x@y) a = n"
+    using neq assms(1) by simp
+  have count_word_b: "count_list (u@v@w@x@y) b = n"
+    using neq assms(1) by simp
+  have count_word_c: "count_list (u@v@w@x@y) c = n"
+    using neq assms(1) by simp
+  have count_non_zero: "((count_list (v@x) a) \<noteq>0) \<or> (count_list (v@x) b\<noteq>0) \<or> (count_list (v@x) c \<noteq> 0)"
+    using count_vx_not_zero[of u v w x y n] assms(1,3) by simp  
+  from assms(2) show ?thesis proof
+    assume *: "count_list (v @ w @ x) a = 0"
+    then have vx_b_or_c_not0: "count_list (v@x) b \<noteq> 0 \<or> count_list (v@x) c \<noteq> 0" using count_non_zero
+      by simp
+    have uwy_count_a: "count_list (u@w@y) a = n" using * count_word_a 
+      by simp 
+    have "count_list (u@w@y) b \<noteq> n \<or> count_list (u@w@y) c \<noteq> n" using vx_b_or_c_not0 count_word_b count_word_c 
+      by auto
+    then have "(count_list (u@w@y) a \<noteq> count_list (u@w@y) b) \<or> (count_list (u@w@y) c \<noteq> count_list (u@w@y) a)" using uwy_count_a
+      by argo
+    then show ?thesis using not_in_count[of "u@w@y"] 
+      by blast
+  next
+    assume *: "count_list (v @ w @ x) c = 0"
+    then have vx_a_or_b_not0: "(count_list (v@x) a\<noteq>0) \<or> (count_list (v@x) b \<noteq> 0)" using count_non_zero
+      by fastforce
+    have uwy_count_c: "count_list (u@w@y) c=n" using * count_word_c 
+      by auto 
+    have "count_list (u@w@y) a \<noteq>n \<or> count_list (u@w@y) b \<noteq>n" using vx_a_or_b_not0 count_word_a count_word_b
+      by force
+    then have "(count_list (u@w@y) a \<noteq> count_list (u@w@y) b) \<or> (count_list (u@w@y) c \<noteq> count_list (u@w@y) a)" using uwy_count_c
+      by argo
+    then show ?thesis using not_in_count[of "u@w@y"] 
+      by blast  
+  qed
+qed
+
+theorem anbncn_not_cfl:
+  assumes "finite (P :: ('n::infinite,'a)Prods)"
+  shows "Lang P S \<noteq> (\<Union>n. {[a]^^n @ [b]^^n @ [c]^^n})" (is "\<not> ?E")
+proof
+  assume "?E"
+  from Pumping_Lemma[OF \<open>finite P\<close>, of S] obtain n where
+    pump: "\<forall>word \<in> Lang P S. length word \<ge> n \<longrightarrow>
+     (\<exists>u v w x y. word = u@v@w@x@y \<and> length (v@w@x) \<le> n \<and> length (v@x) \<ge> 1 \<and> (\<forall>i. u@(v^^i)@w@(x^^i)@y \<in> Lang P S))" 
+    by blast
+  let ?word = "[a]^^n @ [b]^^n @ [c]^^n"
+  have wInLg: "?word \<in> Lang P S"
+    using \<open>?E\<close> by blast
+  have "length ?word \<ge> n"
+    by simp
+  then obtain u v w x y where uvwxy: "?word = u@v@w@x@y \<and> length (v@w@x) \<le> n \<and> length (v@x) \<ge> 1 \<and> (\<forall>i. u@(v^^i)@w@(x^^i)@y \<in> Lang P S)"
+    using pump wInLg by metis
+  let ?vwx= "v@w@x"
+  have "(count_list ?vwx  a=0 ) \<or> (count_list ?vwx c=0)" using  uvwxy a_or_b_zero assms 
+    by (metis (no_types, lifting) append.assoc)
+  then show False using assms uvwxy pumping_application[of u v w x y n]
+    by (metis \<open>?E\<close> append_Nil length_0_conv not_one_le_zero pow_list.simps(1))
+qed
+
+end
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/AnBn_Not_Regular.thy
--- /dev/null
+++ thys/Context_Free_Grammar/AnBn_Not_Regular.thy
@@ -0,0 +1,79 @@
+(*
+Authors: Kaan Taskin, Tobias Nipkow
+*)
+
+section \<open>\<open>anbn\<close> is Not Regular\<close>
+
+theory AnBn_Not_Regular
+imports Pumping_Lemma_Regular
+begin
+
+text
+\<open>The following theorem proves that the language \<open>a^nb^n\<close> cannot be produced by a right linear production set, using the contrapositive form 
+ of the pumping lemma\<close>
+
+theorem not_rlin2_ab:
+  assumes "a \<noteq> b"
+      and "Lang P A = (\<Union>n. {[a]^^n @ [b]^^n})" (is "_ = ?AnBn")
+      and "finite P"
+    shows "\<not>rlin2 P"
+proof -
+  have "\<exists>w \<in> Lang P A. length w \<ge> n \<and> (\<forall>x y z. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> n \<longrightarrow> (\<exists>i. x@(y^^i)@z \<notin> Lang P A))" for n
+  proof -
+    let ?anbn = "[a]^^n @ [b]^^n"
+    show ?thesis
+    proof
+      have **: "(\<exists>i. x @ (y^^i) @ z \<notin> Lang P A)"
+        if asm: "?anbn = x @ y @ z \<and> 1 \<le> length y \<and> length (x @ y) \<le> n" for x y z
+      proof
+        from asm have asm1: "[a]^^n @ [b]^^n = x @ y @ z" by blast
+        from asm have asm2: "1 \<le> length y" by blast
+        from asm have asm3: "length (x @ y) \<le> n" by blast
+        let ?kx = "length x" let ?ky = "length y"
+        have splitted: "x = [a]^^?kx \<and> y = [a]^^?ky \<and> z = [a]^^(n - ?kx - ?ky) @ [b]^^n"
+        proof -
+          have "\<forall>i < n. ([a]^^n @ [b]^^n)!i = a"
+            by (simp add: nth_append nth_pow_list_single)
+          with asm1 have xyz_tma: "\<forall>i < n. (x@y@z)!i = a" by metis
+          with asm3 have xy_tma: "\<forall>i < length(x@y). (x@y)!i = a"
+            by (metis append_assoc nth_append order_less_le_trans)
+          from xy_tma have "\<forall>i < ?kx. x!i = a"
+            by (metis le_add1 length_append nth_append order_less_le_trans)
+          hence *: "x = [a]^^?kx"
+            by (simp add: list_eq_iff_nth_eq nth_pow_list_single)
+          from xy_tma have "\<forall>i < ?ky. y!i = a"
+            by (metis length_append nat_add_left_cancel_less nth_append_length_plus)
+          hence **: "y = [a]^^?ky"
+            by (simp add: nth_equalityI pow_list_single)
+          from * ** asm1 have "[a]^^n @ [b]^^n = [a]^^?kx @ [a]^^?ky @ z" by simp
+          hence z_rest: "[a]^^n @ [b]^^n = [a]^^(?kx+?ky) @ z"
+            by (simp add: pow_list_add)
+          from asm3 have ***: "z = [a]^^(n-?kx-?ky) @ [b]^^n" 
+            using pow_list_eq_appends_iff[THEN iffD1, OF _ z_rest] by simp
+          from * ** *** show ?thesis by blast
+        qed
+        from splitted have "x @ y^^2 @ z = [a]^^?kx @ ([a]^^?ky)^^2 @ [a]^^(n - ?kx - ?ky) @ [b]^^n" by simp
+        also have "... = [a]^^?kx @ [a]^^(?ky*2) @ [a]^^(n - ?kx - ?ky) @ [b]^^n"
+          by (simp add: pow_list_mult)
+        also have "... = [a]^^(?kx + ?ky*2 + (n - ?kx - ?ky)) @ [b]^^n"
+          by (simp add: pow_list_add)
+        also from asm3 have "... = [a]^^(n+?ky) @ [b]^^n"
+          by (simp add: add.commute)
+        finally have wit: "x @ y^^2 @ z = [a]^^(n+?ky) @ [b]^^n" .
+        from asm2 have "[a]^^(n + ?ky) @ [b]^^n \<notin> ?AnBn" 
+          using \<open>a\<noteq>b\<close> by auto
+        with wit have "x @ y^^2 @ z \<notin> ?AnBn" by simp
+        thus "x @ y^^2 @ z \<notin> Lang P A" using assms(2) by blast
+      qed
+      from ** show "n \<le> length ?anbn \<and> (\<forall>x y z. ?anbn = x @ y @ z \<and> 1 \<le> length y \<and> length (x @ y) \<le> n \<longrightarrow> (\<exists>i. x @ (y^^i) @ z \<notin> Lang P A))" by simp
+    next
+      have "?anbn \<in> ?AnBn" by blast
+      thus "?anbn \<in> Lang P A" 
+        by (simp add: assms(2)) 
+    qed
+  qed
+  thus ?thesis 
+    using pumping_lemma_regular_contr[OF assms(3)] by blast
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Binarize.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Binarize.thy
@@ -0,0 +1,331 @@
+(* Authors: Kaan Taskin *)
+
+section \<open>Transforming Long Productions Into a Binary Cascade\<close>
+
+theory Binarize
+imports Inlining1Prod
+begin
+
+(* FFPI *) 
+lemma funpow_fix: fixes f :: "'a \<Rightarrow> 'a" and m :: "'a \<Rightarrow> nat"
+assumes "\<And>x. m(f x) < m x \<or> f x = x"
+shows "f((f ^^ m x) x) = (f ^^ m x) x"
+proof -
+ have "n + m ((f^^n) x) \<le> m x \<or> f((f^^n) x) = (f^^n) x" for n
+ proof(induction n)
+   case 0
+   then show ?case by simp
+ next
+   case (Suc n)
+   then show ?case
+   proof
+     assume a1: "n + m ((f ^^ n) x) \<le> m x"
+     then show ?thesis
+     proof (cases "m ((f ^^ Suc n) x) < m ((f ^^ n) x)")
+       case True
+       then show ?thesis using a1 by auto
+     next
+       case False
+       with assms have "(f ^^ Suc n) x = (f ^^ n) x" by auto
+       then show ?thesis by simp
+     qed
+   next
+     assume "f ((f ^^ n) x) = (f ^^ n) x"
+     then show ?thesis by simp
+   qed
+ qed
+  from this[of "m x"] show ?thesis using assms[of "(f ^^ m x) x"] by auto
+qed
+
+text
+\<open>In a binary grammar, every right-hand side consists of at most two symbols. The \<open>binarize\<close> function should convert an
+ arbitrary production list into a binary production list, without changing the language of the grammar. For this we make use
+ of fixpoint iteration and define the function \<open>binarize1\<close> for splitting a production, whose right-hand side exceeds the
+ maximum number of symbols 2, into two productions. The step function is then defined as the auxiliary function \<open>binarize'\<close>.
+ We also define the count function \<open>count\<close> that counts the right-hand sides whose length is more than or equal to 2\<close>
+
+fun binarize1 :: "('n :: infinite, 't) prods \<Rightarrow> ('n, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "binarize1 ps' [] = []"
+| "binarize1 ps' ((A, []) # ps) = (A, []) # binarize1 ps' ps"
+| "binarize1 ps' ((A, s0 # u) # ps) =
+ (if length u \<le> 1 then (A, s0 # u) # binarize1 ps' ps
+  else let B = fresh (nts ps') in (A,[s0, Nt B]) # (B, u) # ps)"
+
+definition binarize' :: "('n::infinite, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "binarize' ps = binarize1 ps ps"
+
+fun count :: "('n::infinite, 't) prods \<Rightarrow> nat" where
+  "count [] = 0"
+| "count ((A,u) # ps) = (if length u \<le> 2 then count ps else length u + count ps)"
+
+definition binarize :: "('n::infinite, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "binarize ps = (binarize' ^^ (count ps)) ps"
+
+text
+\<open>Firstly we show that the \<open>binarize\<close> function transforms a production list into a binary production list\<close>
+
+lemma count_dec1:
+  assumes "binarize1 ps' ps \<noteq> ps" 
+  shows "count ps > count (binarize1 ps' ps)"
+using assms proof (induction ps' ps rule: binarize1.induct)
+  case (3 ps' A s0 u ps)
+  show ?case proof (cases "length u \<le> 1")
+    case True
+    with "3.prems" have "binarize1 ps' ps \<noteq> ps" by simp
+    with True have "count (binarize1 ps' ps) < count ps"
+      using "3.IH" by simp
+    with True show ?thesis by simp
+  next
+    case False
+    let ?B = "fresh(nts ps')"
+    from False have "count (binarize1 ps' ((A, s0 # u) # ps)) = count ((A,[s0, Nt ?B]) # (?B, u) # ps)"
+      by (metis binarize1.simps(3))
+    also have "... = count ((?B, u) # ps)" by simp
+    also from False have "... < count ((A, s0 # u) # ps)" by simp
+    finally have "count (binarize1 ps' ((A, s0 # u) # ps)) < count ((A, s0 # u) # ps)" by simp
+    thus ?thesis .
+  qed
+qed simp_all
+
+lemma count_dec':
+  assumes "binarize' ps \<noteq> ps" 
+  shows "count ps > count (binarize' ps)"
+ using binarize'_def assms count_dec1 by metis
+
+lemma binarize_ffpi:
+  "binarize'((binarize' ^^ count x) x) = (binarize' ^^ count x) x"
+proof -
+  have "\<And>x. count(binarize' x) < count x \<or> binarize' x = x"
+    using count_dec' by blast
+  thus ?thesis using funpow_fix by metis
+qed
+
+lemma binarize_binary1: 
+  assumes "ps = binarize1 ps' ps"
+  shows "(A,w) \<in> set(binarize1 ps' ps) \<Longrightarrow> length w \<le> 2"
+  using assms proof (induction ps' ps rule: binarize1.induct)
+  case (3 ps' C s0 u ps)
+  show ?case proof (cases "length u \<le> 1")
+    case True
+    with 3 show ?thesis by auto
+  next
+    case False
+    hence "(C, s0 # u) # ps \<noteq> binarize1 ps' ((C, s0 # u) # ps)"
+      by (metis binarize1.simps(3) list.sel(3) not_Cons_self2)
+    with "3.prems"(2) show ?thesis by blast
+  qed
+qed auto
+
+lemma binarize_binary':
+  assumes "ps = binarize' ps"
+  shows "(A,w) \<in> set(binarize' ps) \<Longrightarrow> length w \<le> 2"
+  using binarize'_def assms binarize_binary1 by metis
+
+lemma binarize_binary: "(A,w) \<in> set(binarize ps) \<Longrightarrow> length w \<le> 2"
+  unfolding binarize_def using binarize_ffpi binarize_binary' by metis
+
+text
+\<open>Now we prove the property of language preservation\<close>
+
+lemma binarize1_cases:
+  "binarize1 ps' ps = ps \<or> (\<exists>A ps'' B u s. set ps = {(A, s#u)} \<union> set ps'' \<and> set (binarize1 ps' ps) = {(A,[s,Nt B]),(B,u)} \<union> set ps'' \<and> Nt B \<notin> syms ps')"
+proof (induction ps' ps rule: binarize1.induct)
+  case (2 ps' C ps)
+  then show ?case
+  proof (cases "binarize1 ps' ps = ps")
+    case True
+    then show ?thesis by simp
+  next
+    case False
+    then obtain A ps'' B u s where defs: "set ps = {(A, s # u)} \<union> set ps'' \<and> set (binarize1 ps' ps) = {(A, [s, Nt B]), (B, u)} \<union> set ps'' \<and> Nt B \<notin> syms ps'"
+      using 2 by blast
+    from defs have wit: "set ((C, []) # ps) = {(A, s # u)} \<union> set ((C, []) # ps'')" by simp
+    from defs have wit2: "set (binarize1 ps' ((C, []) # ps)) = {(A, [s, Nt B]), (B, u)} \<union> set ((C, []) # ps'')" by simp
+    from defs have wit3: "Nt B \<notin> syms ps'" by simp
+    from wit wit2 wit3 show ?thesis by blast
+  qed
+next
+  case (3 ps' C s0 u ps)
+  show ?case proof (cases "length u \<le> 1")
+    case T1: True
+    then show ?thesis proof (cases "binarize1 ps' ps = ps")
+      case T2: True
+      with T1 show ?thesis by simp
+    next
+      case False
+      with T1 obtain A ps'' B v s where defs: "set ps = {(A, s # v)} \<union> set ps'' \<and> set (binarize1 ps' ps) = {(A, [s, Nt B]), (B, v)} \<union> set ps'' \<and> Nt B \<notin> syms ps'"
+        using 3 by blast
+      from defs have wit: "set ((C, s0 # u) # ps) = {(A, s # v)} \<union> set ((C, s0 # u) # ps'')" by simp
+      from defs T1 have wit2: "set (binarize1 ps' ((C, s0 # u) # ps)) = {(A, [s, Nt B]), (B, v)} \<union> set ((C, s0 # u) # ps'')" by simp
+      from defs have wit3: "Nt B \<notin> syms ps'" by simp
+      from wit wit2 wit3 show ?thesis by blast
+    qed
+   next
+    case False
+    then show ?thesis
+      using fresh_nts in_Nts_iff_in_Syms[of "fresh (nts ps')" "set ps'"]
+      by (fastforce simp add: Let_def)
+  qed
+qed simp
+
+text
+\<open>We show that a list of terminals \<open>map Tm x\<close> can be derived from the original production set \<open>ps\<close> if and only if \<open>map Tm x\<close>
+ can be derived after the transformation of the step function \<open>binarize'\<close>, under the assumption that the starting symbol \<open>A\<close>
+ occurs in a left-hand side of at least one production in \<open>ps\<close>. We can then extend this property to the \<open>binarize\<close> function\<close>
+
+lemma binarize_der':
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (binarize' ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+  unfolding binarize'_def proof (cases "binarize1 ps ps = ps")
+  case False
+  then obtain C ps'' B u s where defs: "set ps = {(C, s # u)} \<union> set ps'' \<and> set (binarize1 ps ps) = {(C, [s, Nt B]), (B, u)} \<union> set ps'' \<and> Nt B \<notin> syms ps"
+    by (meson binarize1_cases)
+  from defs have a_not_b: "C \<noteq> B" unfolding Syms_def by fast
+  from defs assms have a1: "A \<noteq> B" unfolding Lhss_def Syms_def by auto
+  from defs have a2: "Nt B \<notin> set (map Tm x)" by auto
+  from defs have a3: "Nt B \<notin> set u" unfolding Syms_def by fastforce
+  from defs have "set ps = set ((C, s # u) # ps'')" by simp
+  with defs a_not_b have a4: "B \<notin> lhss ((C, [s, Nt B]) # ps'')" unfolding Lhss_def Syms_def by auto
+  from defs have notB: "Nt B \<notin> syms ps''" by fastforce
+  then have 1: "set ps = set (substP (Nt B) u ((C, [s, Nt B]) # ps''))" proof -
+    from defs have "set ps = {(C, s # u)} \<union> set ps''" by simp
+    also have "... = set ((C, s#u) # ps'')" by simp
+    also have "... = set ([(C, s#u)] @ ps'')" by simp
+    also from defs have "... = set ([(C,substs (Nt B) u [s, Nt B])] @ ps'')" unfolding Syms_def by fastforce
+    also have "... = set ((substP (Nt B) u [(C, [s, Nt B])]) @ ps'')" by (simp add: substP_def)
+    also have "... = set ((substP (Nt B) u [(C, [s, Nt B])]) @ substP (Nt B) u ps'')" using notB by (simp add: substP_skip2)
+    also have "... = set (substP (Nt B) u ((C, [s, Nt B]) # ps''))" by (simp add: substP_def)
+    finally show ?thesis .
+  qed
+  from defs have 2: "set (binarize1 ps ps) = set ((C, [s, Nt B]) # (B, u) # ps'')" by auto
+  with 1 2 a1 a2 a3 a4 show "(set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x) = (set (binarize1 ps ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x)"
+    by (simp add: derives_inlining insert_commute)
+qed simp
+
+lemma lhss_binarize1:
+  "lhss ps \<subseteq> lhss (binarize1 ps' ps)"
+proof (induction ps' ps rule: binarize1.induct)
+  case (3 ps' A s0 u ps)
+  then show ?case proof (cases "length u \<le> 1")
+    case True
+    with 3 show ?thesis by auto
+  next
+    case False
+    let ?B = "fresh(nts ps')"
+    have "lhss ((A, s0 # u) # ps) = {A} \<union> lhss ps" by simp
+    also have "... \<subseteq> {A} \<union> {?B} \<union> lhss ps" by blast
+    also have "... = lhss ((A,[s0, Nt ?B]) # (?B, u) # ps)" by simp
+    also from False have "... = lhss (binarize1 ps' ((A, s0 # u) # ps))"
+      by (metis binarize1.simps(3))
+    finally show ?thesis .
+  qed
+qed auto
+
+lemma lhss_binarize'n:
+  "lhss ps \<subseteq> lhss ((binarize'^^n) ps)"
+proof (induction n)
+  case (Suc n)
+  thus ?case unfolding binarize'_def using lhss_binarize1 by auto
+qed simp
+
+lemma binarize_der'n: 
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set ((binarize'^^n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+using assms proof (induction n)
+  case (Suc n)
+  hence "A \<in> lhss ((binarize' ^^ n) ps)"
+    using lhss_binarize'n by blast
+  hence "set ((binarize' ^^ n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (binarize' ((binarize' ^^ n) ps)) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+    using binarize_der' by blast
+  hence "set ((binarize' ^^ n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set ((binarize' ^^ Suc n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+    by simp
+  with Suc show ?case by blast
+qed simp
+
+lemma binarize_der: 
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (binarize ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+  unfolding binarize_def using binarize_der'n[OF assms] by simp
+
+lemma lang_binarize_lhss: 
+  assumes "A \<in> lhss ps"
+  shows "lang ps A = lang (binarize ps) A"
+  using binarize_der[OF assms] Lang_eqI_derives by metis
+
+text
+\<open>As a last step, we generalize the language preservation property to also include non-terminals which only occur at right-hand
+ sides of the production set\<close>
+
+lemma binarize_syms1:
+  assumes  "Nt A \<in> syms ps"
+    shows  "Nt A \<in> syms (binarize1 ps' ps)"
+using assms proof (induction ps' ps rule: binarize1.induct)
+  case (3 ps' A s0 u ps)
+  show ?case proof (cases "length u \<le> 1")
+    case True
+    with 3 show ?thesis by auto
+  next
+    case False
+    with 3 show ?thesis by (auto simp: Syms_def Let_def)
+  qed
+qed auto
+
+lemma binarize_lhss_nts1:
+  assumes "A \<notin> lhss ps"
+      and "A \<in> nts ps'"
+    shows "A \<notin> lhss (binarize1 ps' ps)"
+  using assms proof (induction ps' ps rule: binarize1.induct)
+  case (3 ps' A s0 u ps)
+  thus ?case proof (cases "length u \<le> 1")
+    case True
+    with 3 show ?thesis by auto
+  next
+    case False
+    with 3 show ?thesis by (auto simp add: Let_def fresh_nts)
+  qed
+qed simp_all
+
+lemma binarize_lhss_nts'n:
+  assumes "A \<notin> lhss ps"
+      and "A \<in> nts ps"
+    shows "A \<notin> lhss ((binarize'^^n) ps) \<and> A \<in> nts ((binarize'^^n) ps)"
+using assms proof (induction n)
+  case (Suc n)
+  thus ?case 
+    unfolding binarize'_def by (simp add: binarize_lhss_nts1 binarize_syms1 in_Nts_iff_in_Syms)
+qed simp
+
+lemma binarize_lhss_nts:
+   assumes "A \<notin> lhss ps"
+      and  "A \<in> nts ps"
+    shows "A \<notin> lhss (binarize ps) \<and> A \<in> nts (binarize ps)"
+  unfolding binarize_def using binarize_lhss_nts'n[OF assms] by simp
+
+lemma binarize_nts'n:
+  assumes "A \<in> nts ps"
+  shows   "A \<in> nts ((binarize' ^^ n) ps)"
+using assms proof (induction n)
+  case (Suc n)
+  thus ?case 
+    unfolding binarize'_def by (simp add: binarize_syms1 in_Nts_iff_in_Syms)
+qed simp
+
+lemma binarize_nts:
+  assumes "A \<in> nts ps"
+  shows   "A \<in> nts (binarize ps)"
+  unfolding binarize_def using assms binarize_nts'n by blast
+
+lemma lang_binarize: 
+  assumes "A \<in> nts ps"
+  shows "lang (binarize ps) A = lang ps A"
+proof (cases "A \<in> lhss ps")
+  case True
+  thus ?thesis
+    using lang_binarize_lhss by blast
+next
+  case False
+  thus ?thesis
+    using assms binarize_lhss_nts Lang_empty_if_notin_Lhss by fast
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/CFL_Not_Intersection_Closed.thy
--- /dev/null
+++ thys/Context_Free_Grammar/CFL_Not_Intersection_Closed.thy
@@ -0,0 +1,233 @@
+(*
+Authors: Felix Krayer, Tobias Nipkow
+*)
+
+section "CFLs Are Not Closed Under Intersection"
+
+theory CFL_Not_Intersection_Closed
+imports
+  "List_Power.List_Power"
+  Context_Free_Language
+  Pumping_Lemma_CFG
+  AnBnCn_not_CFL
+begin
+
+text \<open>The probably first formal proof was given by Ramos \emph{et al.} cite\<open>RamosAMQ and RamosGithub\<close>.
+The proof below is much shorter.\<close>
+
+text "Some lemmas:"
+
+lemma Lang_concat:
+  "L1 @@ L2 = {word. \<exists>w1 \<in> L1. \<exists>w2 \<in> L2. word = w1 @ w2}"
+  unfolding conc_def by blast
+
+lemma deriven_same_repl:
+  assumes "(A, u' @ [Nt A] @ v') \<in> P"
+  shows "P \<turnstile> u @ [Nt A] @ v \<Rightarrow>(n) u @ (u'^^n) @ [Nt A] @ (v'^^n) @ v"
+proof (induction n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  have "P \<turnstile> u @ (u'^^n) @ [Nt A] @ (v'^^n) @ v \<Rightarrow> u @ (u'^^n) @ u' @ [Nt A] @ v' @ (v'^^n) @ v" 
+    using assms derive.intros[of _ _ _ "u @ (u'^^n)" "(v'^^n) @ v"] by fastforce 
+  then have "P \<turnstile> u @ (u'^^n) @ [Nt A] @ (v'^^n) @ v \<Rightarrow> u @ (u'^^(Suc n)) @ [Nt A] @ (v'^^(Suc n)) @ v"
+    by (metis append.assoc pow_list_Suc2 pow_list_comm) 
+  then show ?case using Suc by auto
+qed
+
+
+text "Now the main proof."
+
+lemma an_CFL: "CFL TYPE('n) (\<Union>n. {[a]^^n})" (is "CFL _ ?L")
+proof  -
+  obtain P X where P_def: "(P::('n, 'a) Prods) = {(X, [Tm a, Nt X]), (X, [])}" by simp
+  have "Lang P X = ?L" 
+  proof
+    show "Lang P X \<subseteq> ?L" 
+    proof
+      fix w
+      assume "w \<in> Lang P X"
+      then have "P \<turnstile> [Nt X] \<Rightarrow>* map Tm w" using Lang_def by fastforce
+      then have "\<exists>n. map Tm w = ([Tm a]^^n) @ [Nt X] \<or> (map Tm w::('n, 'a)syms) = ([Tm a]^^n)"
+      proof(induction rule: derives_induct)
+        case base
+        then show ?case by (auto simp: pow_list_single_Nil_iff)
+      next
+        case (step u A v w')
+        then have "A=X" using P_def by auto
+        have "P \<turnstile> u @ [Nt X] @ v \<Rightarrow> u @ w' @ v" using \<open>A=X\<close> derive.intros step by fast
+        obtain n where n_src: "u @ [Nt X] @ v = ([Tm a]^^n) @ [Nt X] \<or> u @ [Nt X] @ v = ([Tm a]^^n)" 
+          using \<open>A=X\<close> step by auto 
+        have notin: "Nt X \<notin> set ([Tm a]^^n)" by (simp add: pow_list_single)
+        then have "u @ [Nt X] @ v = ([Tm a]^^n) @ [Nt X]" 
+          using n_src append_Cons in_set_conv_decomp by metis
+        then have uv_eq: "u = ([Tm a]^^n) \<and> v = []" 
+          using notin n_src Cons_eq_appendI append_Cons_eq_iff append_Nil in_set_insert insert_Nil snoc_eq_iff_butlast by metis
+        have "w' = [Tm a, Nt X] \<or> w' = []" using step(2) P_def by auto
+        then show ?case
+        proof
+          assume "w' = [Tm a, Nt X]"
+          then have "u @ w' @ v = ([Tm a]^^(Suc n)) @ [Nt X]" using uv_eq by (simp add: pow_list_comm)
+          then show ?case by blast
+        next
+          assume "w' = []"
+          then show ?case using uv_eq by blast
+        qed
+      qed
+      then obtain n' where n'_src: "(map Tm w) = ([Tm a]^^n') @ [Nt X] \<or> ((map Tm w)::('n, 'a)syms) = ([Tm a]^^n')" by auto
+      have "Nt X \<notin> set (map Tm w)" by auto
+      then have "((map Tm w)::('n, 'a)syms) = ([Tm a]^^n')" using n'_src by fastforce
+      have "map Tm ([a]^^n') = ([Tm a]^^n')" by (simp add: map_concat)
+      then have "w = [a]^^n'" using \<open>map Tm w = ([Tm a]^^n')\<close> by (metis list.inj_map_strong sym.inject(2))
+      then show "w \<in> ?L" by auto
+    qed
+  next
+    show "?L \<subseteq> Lang P X" 
+    proof
+      fix w
+      assume "w \<in> ?L"
+      then obtain n where n_src: "w = [a]^^n" by blast
+      (*"X \<Rightarrow>* a^n"*)
+      have Xa: "P \<turnstile> [Nt X] \<Rightarrow>(n) ([Tm a]^^n) @ [Nt X]"
+        using P_def deriven_same_repl[of "X" "[Tm a]" "[]" _ _ "[]" ] by (simp add: pow_list_Nil)
+      have X\<epsilon>: "P \<turnstile> ([Tm a]^^n) @ [Nt X] \<Rightarrow> ([Tm a]^^n)" using P_def derive.intros[of "X" "[]" _ "[Tm a]^^n" "[]"] by auto
+      have "([Tm a]^^n) = map Tm w" using n_src by auto
+      then have "P \<turnstile> [Nt X] \<Rightarrow>* map Tm w" using Xa X\<epsilon> relpowp_imp_rtranclp
+        by (smt (verit, best) rtranclp.simps) 
+      then show "w \<in> Lang P X" using Lang_def by fastforce
+    qed
+  qed
+  then show ?thesis unfolding CFL_def P_def by blast
+qed
+
+lemma anbn_CFL: "CFL TYPE('n) (\<Union>n. {[a]^^n @ [b]^^n})" (is "CFL _ ?L")
+proof  -
+  obtain P X where P_def: "(P::('n, 'a) Prods) = {(X, [Tm a, Nt X, Tm b]), (X, [])}" by simp
+  let ?G = "Cfg P X"
+  have "Lang P X = ?L" 
+  proof
+    show "Lang P X \<subseteq> ?L" proof
+      fix w
+      assume "w \<in> Lang P X"
+      then have "P \<turnstile> [Nt X] \<Rightarrow>* map Tm w" using Lang_def by fastforce
+      then have "\<exists>n. map Tm w = ([Tm a]^^n) @ [Nt X] @ ([Tm b]^^n) \<or> (map Tm w::('n, 'a)syms) = ([Tm a]^^n) @ ([Tm b]^^n)"
+      proof(induction rule: derives_induct)
+        case base
+        have "[Nt X] = ([Tm a]^^0) @ [Nt X] @ ([Tm b]^^0)" by auto
+        then show ?case by fast
+      next
+        case (step u A v w')
+        then have "A=X" using P_def by auto
+        have "P \<turnstile> u @ [Nt X] @ v \<Rightarrow> u @ w' @ v" using \<open>A=X\<close> derive.intros step by fast
+        obtain n where n_src: "u @ [Nt X] @ v = ([Tm a]^^n) @ [Nt X] @ ([Tm b]^^n) \<or> u @ [Nt X] @ v = ([Tm a]^^n) @ ([Tm b]^^n)" 
+          using \<open>A=X\<close> step by auto
+        have notin2: "Nt X \<notin> set ([Tm a]^^n) \<and> Nt X \<notin> set ([Tm b]^^n)"
+          by (simp add: pow_list_single)
+        then have notin: "Nt X \<notin> set (([Tm a]^^n) @ ([Tm b]^^n))" by simp
+        then have uv_split: "u @ [Nt X] @ v = ([Tm a]^^n) @ [Nt X] @ ([Tm b]^^n)" 
+          by (metis n_src append_Cons in_set_conv_decomp)
+        have u_eq: "u = ([Tm a]^^n)"
+          by (metis (no_types, lifting) uv_split notin2 Cons_eq_appendI append_Cons_eq_iff eq_Nil_appendI) 
+        then have v_eq: "v = ([Tm b]^^n)" 
+           by (metis (no_types, lifting) uv_split notin2 Cons_eq_appendI append_Cons_eq_iff eq_Nil_appendI)
+        have "w' = [Tm a, Nt X, Tm b] \<or> w' = []" using step(2) P_def by auto
+        then show ?case
+        proof
+          assume "w' = [Tm a, Nt X, Tm b]"
+          then have "u @ w' @ v = ([Tm a]^^n) @ [Tm a, Nt X, Tm b] @ ([Tm b]^^n)" using u_eq v_eq by simp
+          then have "u @ w' @ v = ([Tm a]^^(Suc n)) @ [Nt X] @ ([Tm b]^^(Suc n)) "
+            by (simp add: pow_list_comm)
+          then show ?case by blast
+        next
+          assume "w' = []"
+          then show ?case using u_eq v_eq by blast
+        qed
+      qed
+      then obtain n' where n'_src: "(map Tm w) = ([Tm a]^^n') @ [Nt X] @ ([Tm b]^^n') \<or> ((map Tm w)::('n, 'a)syms) = ([Tm a]^^n') @ ([Tm b]^^n')" by auto
+      have "Nt X \<notin> set (map Tm w)" by auto
+      then have w_ab: "((map Tm w)::('n, 'a)syms) = ([Tm a]^^n') @ ([Tm b]^^n')" using n'_src by fastforce
+      have map_a: "([Tm a]^^n') = map Tm ([a]^^n')" by (simp add: map_concat)
+      have map_b: "([Tm b]^^n') = map Tm ([b]^^n')" by (simp add: map_concat)
+      from w_ab map_a map_b have "((map Tm w)::('n, 'a)syms) = map Tm ([a]^^n') @  map Tm ([b]^^n')" by metis
+      then have "((map Tm w)::('n, 'a)syms) = map Tm (([a]^^n') @ ([b]^^n'))" by simp
+      then have "w = [a]^^n' @ [b]^^n'" by (metis list.inj_map_strong sym.inject(2))
+      then show "w \<in> ?L" by auto
+    qed
+  next
+    show "?L \<subseteq> Lang P X" 
+    proof
+      fix w
+      assume "w \<in> ?L"
+      then obtain n where n_src: "w = [a]^^n @ [b]^^n" by blast
+      (*"X \<Rightarrow>* a^nb^n"*)
+      have Xa: "P \<turnstile> [Nt X] \<Rightarrow>(n) ([Tm a]^^n) @ [Nt X] @ ([Tm b]^^n)"
+        using P_def deriven_same_repl[of "X" "[Tm a]" "[Tm b]" _ _ "[]" "[]"] by simp
+      have X\<epsilon>: "P \<turnstile> ([Tm a]^^n) @ [Nt X] @ ([Tm b]^^n) \<Rightarrow> ([Tm a]^^n) @ ([Tm b]^^n)" 
+        using P_def derive.intros[of "X" "[]" _ "[Tm a]^^n" "[Tm b]^^n"] by auto
+      have "[Tm a]^^n @ [Tm b]^^n = map Tm ([a]^^n) @ (map Tm ([b]^^n))" by simp
+      then have "([Tm a]^^n) @ ([Tm b]^^n) = map Tm w" using n_src by simp
+      then have "P \<turnstile> [Nt X] \<Rightarrow>* map Tm w" using Xa X\<epsilon> relpowp_imp_rtranclp
+        by (smt (verit, best) rtranclp.simps) 
+      then show "w \<in> Lang P X" using Lang_def by fastforce
+    qed
+  qed
+  then show ?thesis unfolding CFL_def P_def by blast
+qed
+
+lemma intersection_anbncn: 
+  assumes "a\<noteq>b" "b\<noteq>c" "c\<noteq>a"
+  and "\<exists>x y z. w = [a]^^x@[b]^^y@[c]^^z \<and> x = y"
+  and "\<exists>x y z. w = [a]^^x@[b]^^y@[c]^^z \<and> y = z"
+  shows "\<exists>x y z. w = [a]^^x@[b]^^y@[c]^^z \<and> x = y \<and> y = z"
+proof -
+  obtain x1 y1 z1 where src1: "w = [a]^^x1@[b]^^y1@[c]^^z1 \<and> x1 = y1" using assms by blast
+  obtain x2 y2 z2 where src2: "w = [a]^^x2@[b]^^y2@[c]^^z2 \<and> y2 = z2" using assms by blast
+  have "[a]^^x1@[b]^^y1@[c]^^z1 = [a]^^x2@[b]^^y2@[c]^^z2" using src1 src2 by simp
+  have cx1: "count_list w a = x1" using src1 assms by (simp)
+  have cx2: "count_list w a = x2" using src2 assms by (simp)
+  from cx1 cx2 have eqx: "x1 = x2" by simp
+
+  have cy1: "count_list w b = y1" using assms src1 by (simp)
+  have cy2: "count_list w b = y2" using assms src2 by simp
+  from cy1 cy2 have eqy: "y1 = y2" by simp
+
+  have cz1: "count_list w c = z1" using assms src1 by simp
+  have cz2: "count_list w c = z2" using assms src2 by simp
+  from cz1 cz2 have eqz: "z1 = z2" by simp
+
+  have "w = [a]^^x1@[b]^^y1@[c]^^z1 \<and> x1 = y1 \<and> y1 = z1" using eqx eqy eqz src1 src2 by blast
+  then show ?thesis by blast
+qed
+
+lemma CFL_intersection_not_closed:
+  fixes a b c :: 'a
+  assumes "a\<noteq>b" "b\<noteq>c" "c\<noteq>a"
+  shows "\<exists>L1 L2 :: 'a list set.
+    CFL TYPE(('n1 + 'n1) option) L1 \<and> CFL TYPE(('n2 + 'n2) option) L2
+ \<and> (\<nexists>(P::('x::infinite,'a)Prods) S. Lang P S = L1 \<inter> L2 \<and> finite P)"
+proof -
+  let ?anbn = "\<Union>n. {[a]^^n @ [b]^^n}"
+  let ?cm = "\<Union>n. {[c]^^n}"
+  let ?an = "\<Union>n. {[a]^^n}"
+  let ?bmcm = "\<Union>n. {[b]^^n @ [c]^^n}"
+  let ?anbncm = "\<Union>n. \<Union>m. {[a]^^n @ [b]^^n @ [c]^^m}"
+  let ?anbmcm = "\<Union>n. \<Union>m. {[a]^^n @ [b]^^m @ [c]^^m}"
+  have anbn: "CFL TYPE('n1) ?anbn" by(rule anbn_CFL)
+  have cm: "CFL TYPE('n1) ?cm" by(rule an_CFL)
+  have an: "CFL TYPE('n2) ?an" by(rule an_CFL)
+  have bmcm: "CFL TYPE('n2) ?bmcm" by(rule anbn_CFL)
+  have "?anbncm = ?anbn @@ ?cm" unfolding Lang_concat by auto
+  then have anbncm: "CFL TYPE(('n1 + 'n1) option) ?anbncm" using anbn cm CFL_concat_closed by auto
+  have "?anbmcm = ?an @@ ?bmcm" unfolding Lang_concat by auto
+  then have anbmcm: "CFL TYPE(('n2 + 'n2) option) ?anbmcm" using an bmcm CFL_concat_closed by auto
+  have "?anbncm \<inter> ?anbmcm = (\<Union> n. {[a]^^n@[b]^^n@[c]^^n})" 
+    using intersection_anbncn[OF assms] by auto
+  then have "CFL TYPE(('n1 + 'n1) option) ?anbncm \<and> 
+        CFL TYPE(('n2 + 'n2) option) ?anbmcm \<and> 
+        (\<nexists>P::('x,'a)Prods.\<exists>S. Lang P S = ?anbncm \<inter> ?anbmcm \<and> finite P)" 
+    using  anbncn_not_cfl[OF assms, where 'n = 'x] anbncm anbmcm by auto
+  then show ?thesis by auto
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Chomsky_Normal_Form.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Chomsky_Normal_Form.thy
@@ -0,0 +1,1067 @@
+(*
+Authors: August Martin Stimpfle, Tobias Nipkow
+Based on HOL4 theories by Aditi Barthwal
+*)
+
+section \<open>Conversion to Chomsky Normal Form\<close>
+
+theory Chomsky_Normal_Form
+imports Unit_Elimination Epsilon_Elimination
+begin
+
+definition CNF :: "('n, 't) Prods \<Rightarrow> bool" where
+"CNF P \<equiv> (\<forall>(A,\<alpha>) \<in> P. (\<exists>B C. \<alpha> = [Nt B, Nt C]) \<or> (\<exists>t. \<alpha> = [Tm t]))"
+
+lemma Nts_correct: "A \<notin> Nts P \<Longrightarrow> (\<nexists>S \<alpha>. (S, \<alpha>) \<in> P \<and> (Nt A \<in> {Nt S} \<union> set \<alpha>))"
+unfolding Nts_def nts_syms_def by auto
+
+(* Chomsky Normal Form *)
+
+definition uniformize :: "'n::infinite \<Rightarrow> 't \<Rightarrow> 'n \<Rightarrow> ('n,'t)prods \<Rightarrow> ('n,'t) prods \<Rightarrow> bool" where 
+      "uniformize A t S ps ps' \<equiv> (
+    \<exists> l r p s. (l,r) \<in> set ps \<and> (r = p@[Tm t]@s) 
+    \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> A = fresh(nts ps \<union> {S})
+    \<and> ps' = ((removeAll (l,r) ps) @ [(A,[Tm t]), (l, p@[Nt A]@s)]))"
+
+lemma uniformize_Eps_free:
+  assumes "Eps_free (set ps)"
+    and "uniformize A t S ps ps'"
+  shows "Eps_free (set ps')"
+  using assms unfolding uniformize_def Eps_free_def by force
+
+lemma uniformize_Unit_free:
+  assumes "Unit_free (set ps)"
+    and "uniformize A t S ps ps'"
+  shows "Unit_free (set ps')"
+proof -
+  have 1: "(\<nexists>l A. (l,[Nt A]) \<in> (set ps))"
+    using assms(1) unfolding Unit_free_def by simp
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Tm t]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) 
+      \<and> set ps' = ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using assms(2) set_removeAll unfolding uniformize_def by force
+  hence "\<nexists>l' A'. (l,[Nt A']) \<in> {(A,[Tm t]), (l, p@[Nt A]@s)}" 
+    using Cons_eq_append_conv by fastforce
+  hence "\<nexists>l' A'. (l',[Nt A']) \<in> ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using 1 by simp
+  moreover have "set ps' = ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using lrps by simp
+  ultimately show ?thesis unfolding Unit_free_def by simp
+qed
+
+definition prodTms :: "('n,'t) prod \<Rightarrow> nat" where
+"prodTms p \<equiv> (if length (snd p) \<le> 1 then 0 else length (filter (isTm) (snd p)))"
+
+definition prodNts :: "('n,'t) prod \<Rightarrow> nat" where
+"prodNts p \<equiv> (if length (snd p) \<le> 2 then 0 else length (filter (isNt) (snd p)))"
+
+(* This definition can be reduced to badTmsCount ps \<equiv> fold (+) (map prodTms ps) 0. 
+   However, the proofs turned out to be much nicer with this version  *)
+fun badTmsCount :: "('n,'t) prods \<Rightarrow> nat" where
+  "badTmsCount ps = sum_list(map prodTms ps)"
+
+lemma badTmsCountSet: "(\<forall>p \<in> set ps. prodTms p = 0) \<longleftrightarrow> badTmsCount ps = 0"
+by auto
+
+fun badNtsCount :: "('n,'t) prods \<Rightarrow> nat" where
+  "badNtsCount ps = sum_list(map prodNts ps)"
+
+lemma badNtsCountSet: "(\<forall>p \<in> set ps. prodNts p = 0) \<longleftrightarrow> badNtsCount ps = 0"
+  by auto
+
+definition uniform :: "('n, 't) Prods \<Rightarrow> bool" where
+  "uniform P \<equiv> \<forall>(A, \<alpha>) \<in> P. (\<nexists>t. Tm t \<in> set \<alpha>) \<or> (\<exists>t. \<alpha> = [Tm t])"
+
+lemma uniform_badTmsCount: 
+  "uniform (set ps) \<longleftrightarrow> badTmsCount ps = 0"
+proof 
+  assume assm: "uniform (set ps)"
+  have "\<forall>p \<in> set ps. prodTms p = 0"
+  proof 
+    fix p assume "p \<in> set ps"
+    hence "(\<nexists>t. Tm t \<in> set (snd p)) \<or> (\<exists>t. snd p = [Tm t])"
+      using assm unfolding uniform_def by auto
+    hence "length (snd p) \<le> 1 \<or> (\<nexists>t. Tm t \<in> set (snd p))"
+      by auto
+    hence "length (snd p) \<le> 1 \<or> length (filter (isTm) (snd p)) = 0"
+      unfolding isTm_def by (auto simp: filter_empty_conv)
+    thus "prodTms p = 0"
+      unfolding prodTms_def by argo
+   qed
+   thus "badTmsCount ps = 0"
+     using badTmsCountSet by blast
+next 
+  assume assm: "badTmsCount ps = 0"
+  have "\<forall>p \<in> set ps. ((\<nexists>t. Tm t \<in> set (snd p)) \<or> (\<exists>t. snd p = [Tm t]))"
+  proof 
+    fix p assume "p \<in> set ps"
+    hence "prodTms p = 0"
+      using assm badTmsCountSet by blast
+    hence "length (snd p) \<le> 1 \<or> length (filter (isTm) (snd p)) = 0"
+      unfolding prodTms_def by argo
+    hence "length (snd p) \<le> 1 \<or> (\<nexists>t. Tm t \<in> set (snd p))"
+      by (auto simp: isTm_def filter_empty_conv)
+    hence "length (snd p) = 0 \<or> length (snd p) = 1 \<or> (\<nexists>t. Tm t \<in> set (snd p))"
+      using order_neq_le_trans by blast
+    thus "(\<nexists>t. Tm t \<in> set (snd p)) \<or> (\<exists>t. snd p = [Tm t])"
+      by (auto simp: length_Suc_conv)
+  qed
+  thus "uniform (set ps)"
+    unfolding uniform_def by auto
+qed
+
+definition binary :: "('n, 't) Prods \<Rightarrow> bool" where
+  "binary P \<equiv> \<forall>(A, \<alpha>) \<in> P. length \<alpha> \<le> 2"
+
+lemma binary_badNtsCount:
+  assumes "uniform (set ps)" "badNtsCount ps = 0"
+  shows "binary (set ps)"
+proof -
+  have "\<forall>p \<in> set ps. length (snd p) \<le> 2"
+  proof 
+    fix p assume assm: "p \<in> set ps"
+    obtain A \<alpha> where "(A, \<alpha>) = p"
+      using prod.collapse by blast
+    hence "((\<nexists>t. Tm t \<in> set \<alpha>) \<or> (\<exists>t. \<alpha> = [Tm t])) \<and> (prodNts (A, \<alpha>) = 0)"
+      using assms badNtsCountSet assm unfolding uniform_def by auto
+    hence "((\<nexists>t. Tm t \<in> set \<alpha>) \<or> (\<exists>t. \<alpha> = [Tm t])) \<and> (length \<alpha> \<le> 2 \<or> length (filter (isNt) \<alpha>) = 0)"
+      unfolding prodNts_def by force
+    hence "((\<nexists>t. Tm t \<in> set \<alpha>) \<or> (length \<alpha> \<le> 1)) \<and> (length \<alpha> \<le> 2 \<or> (\<nexists>N. Nt N \<in> set \<alpha>))"
+      by (auto simp: filter_empty_conv[of isNt \<alpha>] isNt_def)
+    hence "length \<alpha> \<le> 2"
+      by (metis Suc_1 Suc_le_eq in_set_conv_nth le_Suc_eq nat_le_linear sym.exhaust)
+    thus "length (snd p) \<le> 2"
+      using \<open>(A, \<alpha>) = p\<close> by auto
+  qed
+  thus ?thesis
+    by (auto simp: binary_def)
+qed
+
+lemma count_bin_un: "(binary (set ps) \<and> uniform (set ps)) \<longleftrightarrow> (badTmsCount ps = 0 \<and> badNtsCount ps = 0)"
+proof 
+  assume "binary (set ps) \<and> uniform (set ps)"
+  hence "badTmsCount ps = 0 \<and> (\<forall>(A, \<alpha>) \<in> set ps. length \<alpha> \<le> 2)"
+    unfolding binary_def using uniform_badTmsCount by blast
+  thus "badTmsCount ps = 0 \<and> badNtsCount ps = 0"
+    by (metis badNtsCountSet case_prodE prod.sel(2) prodNts_def)
+next
+  assume "badTmsCount ps = 0 \<and> badNtsCount ps = 0"
+  thus "binary (set ps) \<and> uniform (set ps)"
+    using binary_badNtsCount uniform_badTmsCount by blast 
+qed
+
+
+definition binarizeNt :: "'n::infinite \<Rightarrow> 'n \<Rightarrow> 'n \<Rightarrow> 'n \<Rightarrow> ('n,'t)prods \<Rightarrow> ('n,'t)prods \<Rightarrow> bool" where
+"binarizeNt A B1 B2 S ps ps' \<equiv> (
+    \<exists>l r p s. (l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s)
+    \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A = fresh(nts ps \<union> {S}))
+    \<and> ps' = ((removeAll (l,r) ps) @ [(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)]))"
+
+lemma binarizeNt_Eps_free:
+  assumes "Eps_free (set ps)"
+    and "binarizeNt A B1 B2 S ps ps'"
+  shows "Eps_free (set ps')"
+  using assms unfolding binarizeNt_def Eps_free_def by force
+
+lemma binarizeNt_Unit_free:
+  assumes "Unit_free (set ps)"
+    and "binarizeNt A B1 B2 S ps ps'"
+  shows "Unit_free (set ps')"
+  proof -
+  have 1: "(\<nexists>l A. (l,[Nt A]) \<in> (set ps))"
+    using assms(1) unfolding Unit_free_def by simp
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) 
+      \<and> (set ps' = ((set ps - {(l,r)}) \<union> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)}))"
+    using assms(2) set_removeAll unfolding binarizeNt_def by force
+  hence "\<nexists>l' A'. (l,[Nt A']) \<in> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)}" 
+    using Cons_eq_append_conv by fastforce
+  hence "\<nexists>l' A'. (l',[Nt A']) \<in> ((set ps - {(l,r)}) \<union> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)})"
+    using 1 by simp
+  moreover have "set ps' = ((set ps - {(l,r)}) \<union> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)})"
+    using lrps by simp
+  ultimately show ?thesis unfolding Unit_free_def by simp
+qed
+
+lemma fresh_nts_single: "fresh(nts ps \<union> {S}) \<notin> nts ps \<union> {S}"
+by(rule fresh_finite) (simp add: finite_nts)
+
+lemma binarizeNt_aux1:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+  shows "A \<noteq> B1 \<and> A \<noteq> B2"
+using assms fresh_nts_single unfolding binarizeNt_def Nts_def nts_syms_def by fastforce
+
+lemma derives_sub:
+  assumes "P \<turnstile> [Nt A] \<Rightarrow> u" and "P \<turnstile> xs \<Rightarrow> p @ [Nt A] @ s"
+  shows "P \<turnstile> xs \<Rightarrow>* p @ u @ s"
+proof -
+  have "P \<turnstile> p @ [Nt A] @ s \<Rightarrow>* p @ u @ s"
+    using assms derive_append derive_prepend by blast
+  thus ?thesis
+    using assms(2) by simp
+qed
+
+lemma cnf_r1Tm: 
+  assumes "uniformize A t S ps ps'"
+    and "set ps \<turnstile> lhs \<Rightarrow> rhs"
+  shows "set ps' \<turnstile> lhs \<Rightarrow>* rhs"
+proof -
+  obtain p' s' B v where Bv: "lhs = p'@[Nt B]@s' \<and> rhs = p'@v@s' \<and> (B,v) \<in> set ps"
+    using derive.cases[OF assms(2)] by fastforce
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Tm t]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps))
+      \<and> set ps' = ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using assms(1) set_removeAll fresh_nts_single[of ps S] unfolding uniformize_def by fastforce
+  thus ?thesis
+  proof (cases "(B, v) \<in> set ps'")
+    case True
+    then show ?thesis
+      using derive.intros[of B v] Bv by blast
+  next
+    case False
+    hence "B = l \<and> v = p@[Tm t]@s"
+      by (simp add: lrps Bv) 
+    have 1: "set ps' \<turnstile> [Nt l] \<Rightarrow> p@[Nt A]@s"
+      using lrps by (simp add: derive_singleton)
+    have "set ps' \<turnstile> [Nt A] \<Rightarrow> [Tm t]"
+      using lrps by (simp add: derive_singleton)
+    hence "set ps' \<turnstile> [Nt l] \<Rightarrow>* p@[Tm t]@s"
+      using 1 derives_sub[of \<open>set ps'\<close>] by blast
+    then show ?thesis 
+      using False \<open>B = l \<and> v = p@[Tm t]@s\<close> Bv derives_append derives_prepend by blast
+  qed
+qed
+
+lemma cnf_r1Nt:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+    and "set ps \<turnstile> lhs \<Rightarrow> rhs"
+  shows "set ps' \<turnstile> lhs \<Rightarrow>* rhs"
+proof -
+  obtain p' s' C v where Cv: "lhs = p'@[Nt C]@s' \<and> rhs = p'@v@s' \<and> (C,v) \<in> set ps"
+    using derive.cases[OF assms(2)] by fastforce
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps))
+    \<and> (set ps' = ((set ps - {(l,r)}) \<union> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)}))"
+    using assms(1) set_removeAll fresh_nts_single[of ps S] unfolding binarizeNt_def by fastforce
+  thus ?thesis
+  proof (cases "(C, v) \<in> set ps'")
+    case True
+    then show ?thesis
+      using derive.intros[of C v] Cv by blast
+  next
+    case False
+    hence "C = l \<and> v = p@[Nt B1,Nt B2]@s"
+      by (simp add: lrps Cv)
+    have 1: "set ps' \<turnstile> [Nt l] \<Rightarrow> p@[Nt A]@s"
+      using lrps by (simp add: derive_singleton)
+    have "set ps' \<turnstile> [Nt A] \<Rightarrow> [Nt B1,Nt B2]"
+      using lrps by (simp add: derive_singleton)
+    hence "set ps' \<turnstile> [Nt l] \<Rightarrow>* p@[Nt B1,Nt B2]@s" 
+      using 1 derives_sub[of \<open>set ps'\<close>] by blast
+    thus ?thesis 
+      using False \<open>C = l \<and> v = p@[Nt B1,Nt B2]@s\<close> Cv derives_append derives_prepend by blast
+  qed
+qed
+
+lemma slemma1_1: 
+  assumes "uniformize A t S ps ps'"
+    and "(A, \<alpha>) \<in> set ps'"
+  shows "\<alpha> = [Tm t]"
+proof -
+  have "A \<notin> Nts (set ps)"
+    using assms(1) fresh_nts_single unfolding uniformize_def by blast
+  hence "\<nexists>\<alpha>. (A, \<alpha>) \<in> set ps"
+    unfolding Nts_def by auto
+  hence "\<nexists>\<alpha>. \<alpha> \<noteq> [Tm t] \<and> (A, \<alpha>) \<in> set ps'"
+    using assms(1) unfolding uniformize_def by auto
+  thus ?thesis 
+    using assms(2) by blast
+qed
+
+lemma slemma1_1Nt:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+    and "(A, \<alpha>) \<in> set ps'"
+  shows "\<alpha> = [Nt B1,Nt B2]"
+proof -
+  have "A \<notin> Nts (set ps)"
+    using assms(1) fresh_nts_single unfolding binarizeNt_def by blast
+  hence "\<nexists>\<alpha>. (A, \<alpha>) \<in> set ps"
+    unfolding Nts_def  by auto
+  hence "\<nexists>\<alpha>. \<alpha> \<noteq> [Nt B1,Nt B2] \<and> (A, \<alpha>) \<in> set ps'"
+    using assms(1) unfolding binarizeNt_def by auto
+  thus ?thesis 
+    using assms(2) by blast
+qed
+
+lemma slemma4_1:
+  assumes "Nt A \<notin> set rhs"
+  shows "\<forall>\<alpha>. rhs = substsNt A \<alpha> rhs"
+  using assms by (simp add: substs_skip)
+
+lemma slemma4_3_1:
+  assumes "lhs = A"
+  shows "\<alpha> = substsNt A \<alpha> [Nt lhs]"
+  using assms by simp
+
+lemma slemma4_4:
+  assumes "uniformize A t S ps ps'"
+    and "(l,r) \<in> set ps"
+  shows "Nt A \<notin> set r"
+proof -
+  have "A \<notin> Nts (set ps)"
+    using assms(1) fresh_nts_single unfolding uniformize_def by blast
+  hence "\<nexists>S \<alpha>. (S, \<alpha>) \<in> set ps \<and> (Nt A \<in> {Nt S} \<union> set \<alpha>)"
+    using Nts_correct[of A \<open>set ps\<close>] by blast
+  thus ?thesis 
+    using assms(2) by blast
+qed
+
+lemma slemma4_4Nt:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+    and "(l,r) \<in> set ps"
+  shows "(Nt A) \<notin> set r"
+proof -
+  have "A \<notin> Nts (set ps)"
+    using assms(1) fresh_nts_single unfolding binarizeNt_def by blast
+  hence "\<nexists>S \<alpha>. (S, \<alpha>) \<in> set ps \<and> (Nt A \<in> {Nt S} \<union> set \<alpha>)"
+    using Nts_correct[of A \<open>set ps\<close>] by blast
+  thus ?thesis 
+    using assms(2) by blast
+qed
+
+
+lemma lemma1:
+  assumes "uniformize A t S ps ps'"
+    and "set ps' \<turnstile> lhs \<Rightarrow> rhs"
+  shows "substsNt A [Tm t] lhs = substsNt A [Tm t] rhs
+    \<or> set ps \<turnstile> substsNt A [Tm t] lhs \<Rightarrow> substsNt A [Tm t] rhs"
+proof -
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Tm t]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps)) 
+      \<and> set ps' = ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using assms(1) set_removeAll fresh_nts_single[of ps S] unfolding uniformize_def by fastforce
+  obtain p' s' u v where uv: "lhs = p'@[Nt u]@s' \<and> rhs = p'@v@s' \<and> (u,v) \<in> set ps'"
+    using derive.cases[OF assms(2)] by fastforce
+  thus ?thesis
+  proof (cases "u = A")
+    case True
+    then show ?thesis 
+    proof (cases "v = [Tm t]")
+      case True
+      have "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ substsNt A [Tm t] ([Nt A]@s')"
+        using uv \<open>u = A\<close> by simp
+      hence "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ [Tm t] @ substsNt A [Tm t] s'"
+        by simp
+      then show ?thesis
+        by (simp add: True uv) 
+    next
+      case False
+      then show ?thesis 
+        using True uv assms(1) slemma1_1 by fastforce 
+    qed
+  next
+    case False
+    then show ?thesis 
+    proof (cases "(Nt A) \<in> set v")
+      case True
+      hence 1: "v = p@[Nt A]@s \<and> Nt A \<notin> set p \<and> Nt A \<notin> set s" 
+        using lrps uv assms slemma4_4 by fastforce
+      hence "substsNt A [Tm t] v = substsNt A [Tm t] p @ substsNt A [Tm t] ([Nt A]@s)"
+        by simp
+      hence "substsNt A [Tm t] v = p @ [Tm t] @ s"
+        using 1 substs_append slemma4_1 slemma4_3_1 by metis
+      hence 2: "(u, substsNt A [Tm t] v) \<in> set ps" using lrps
+        using True uv assms(1) slemma4_4 by fastforce
+      have "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ substsNt A [Tm t] ([Nt u]@s')"
+        using uv by simp
+      hence 3: "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ [Nt u] @ substsNt A [Tm t] s'" 
+        using \<open>u \<noteq> A\<close> by simp
+      have "substsNt A [Tm t] rhs = substsNt A [Tm t] p' @ substsNt A [Tm t] (v@s')"
+        using uv by simp
+      hence "substsNt A [Tm t] rhs = substsNt A [Tm t] p' @ substsNt A [Tm t] v @ substsNt A [Tm t] s'"
+        by simp
+      then show ?thesis 
+        using 2 3 assms(2) uv derive.simps by fast
+    next
+      case False
+      hence 1: "(u, v) \<in> set ps" 
+        using assms(1) uv \<open>u \<noteq> A\<close> lrps by (simp add: in_set_conv_decomp)
+       have "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ substsNt A [Tm t] ([Nt u]@s')"
+         using uv by simp
+       hence 2: "substsNt A [Tm t] lhs = substsNt A [Tm t] p' @ [Nt u] @ substsNt A [Tm t] s'"
+         using \<open>u \<noteq> A\<close> by simp
+       have "substsNt A [Tm t] rhs = substsNt A [Tm t] p' @ substsNt A [Tm t] (v@s')"
+         using uv by simp
+       hence "substsNt A [Tm t] rhs = substsNt A [Tm t] p' @ substsNt A [Tm t] v @ substsNt A [Tm t] s'"
+         by simp
+       hence "substsNt A [Tm t] rhs = substsNt A [Tm t] p' @ v @ substsNt A [Tm t] s'"
+         using False slemma4_1 by fastforce
+       thus ?thesis 
+         using 1 2 assms(2) uv derive.simps by fast
+    qed
+  qed
+qed
+
+lemma lemma1Nt: 
+  assumes "binarizeNt A B1 B2 S ps ps'"
+    and "set ps' \<turnstile> lhs \<Rightarrow> rhs"
+  shows "(substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] rhs) 
+          \<or> ((set ps) \<turnstile> (substsNt A [Nt B1,Nt B2] lhs) \<Rightarrow> substsNt A [Nt B1,Nt B2] rhs)"
+proof -
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps))
+    \<and> (set ps' = ((set ps - {(l,r)}) \<union> {(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)}))"
+    using assms(1) set_removeAll fresh_nts_single[of ps S] unfolding binarizeNt_def by fastforce
+  obtain p' s' u v where uv: "lhs = p'@[Nt u]@s' \<and> rhs = p'@v@s' \<and> (u,v) \<in> set ps'"
+    using derive.cases[OF assms(2)] by fastforce
+  thus ?thesis
+  proof (cases "u = A")
+    case True
+    then show ?thesis 
+    proof (cases "v = [Nt B1,Nt B2]")
+      case True
+      have "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] ([Nt A]@s')"
+        using uv \<open>u = A\<close> by simp
+      hence 1: "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ [Nt B1,Nt B2] @ substsNt A [Nt B1,Nt B2] s'"
+        by simp
+      have "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] ([Nt B1,Nt B2]@s')"
+        using uv \<open>u = A\<close> True by simp
+      hence "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ [Nt B1,Nt B2] @ substsNt A [Nt B1,Nt B2] s'"
+        using assms(1) binarizeNt_aux1[of A B1 B2 S ps ps'] by auto 
+      then show ?thesis
+        using 1 by simp
+    next
+      case False
+      then show ?thesis
+        using True uv assms(1) slemma1_1Nt by fastforce
+    qed
+  next
+    case False
+    then show ?thesis 
+    proof (cases "(Nt A) \<in> set v")
+      case True
+      have "Nt A \<notin> set p \<and> Nt A \<notin> set s" 
+        using lrps assms(1) by (metis UnI1 UnI2 set_append slemma4_4Nt)
+      hence 1: "v = p@[Nt A]@s \<and> Nt A \<notin> set p \<and> Nt A \<notin> set s" 
+        using True lrps uv assms slemma4_4Nt[of A B1 B2 S ps ps'] binarizeNt_aux1[of A B1 B2 S ps ps'] by auto
+      hence "substsNt A [Nt B1,Nt B2] v = substsNt A [Nt B1,Nt B2] p @ substsNt A [Nt B1,Nt B2] ([Nt A]@s)"
+        by simp
+      hence "substsNt A [Nt B1,Nt B2] v = p @ [Nt B1,Nt B2] @ s"
+        using 1 substs_append slemma4_1 slemma4_3_1 by metis
+      hence 2: "(u, substsNt A [Nt B1,Nt B2] v) \<in> set ps" 
+        using True lrps uv assms(1) slemma4_4Nt by fastforce
+      have "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] ([Nt u]@s')"
+        using uv by simp
+      hence 3: "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ [Nt u] @ substsNt A [Nt B1,Nt B2] s'" 
+        using \<open>u \<noteq> A\<close> by simp
+      have "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] (v@s')"
+        using uv by simp
+      hence "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] v @ substsNt A [Nt B1,Nt B2] s'"
+        by simp
+      then show ?thesis 
+        using 2 3 assms(2) uv derive.simps by fast
+    next
+      case False
+      hence 1: "(u, v) \<in> set ps" 
+        using assms(1) uv \<open>u \<noteq> A\<close> lrps by (simp add: in_set_conv_decomp)
+       have "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] ([Nt u]@s')"
+         using uv by simp
+       hence 2: "substsNt A [Nt B1,Nt B2] lhs = substsNt A [Nt B1,Nt B2] p' @ [Nt u] @ substsNt A [Nt B1,Nt B2] s'"
+         using \<open>u \<noteq> A\<close> by simp
+       have "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] (v@s')"
+         using uv by simp
+       hence "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ substsNt A [Nt B1,Nt B2] v @ substsNt A [Nt B1,Nt B2] s'"
+         by simp
+       hence "substsNt A [Nt B1,Nt B2] rhs = substsNt A [Nt B1,Nt B2] p' @ v @ substsNt A [Nt B1,Nt B2] s'"
+         using False slemma4_1 by fastforce
+       thus ?thesis 
+         using 1 2 assms(2) uv derive.simps by fast
+    qed
+  qed
+qed
+
+lemma lemma3:
+  assumes "set ps' \<turnstile> lhs \<Rightarrow>* rhs"
+    and "uniformize A t S ps ps'"
+  shows "set ps \<turnstile> substsNt A [Tm t] lhs \<Rightarrow>* substsNt A [Tm t] rhs"
+  using assms
+proof (induction rhs rule: rtranclp_induct)
+  case (step y z)
+  then show ?case 
+    using lemma1[of A t S ps ps' y z] by auto 
+qed simp
+
+lemma lemma3Nt:
+  assumes "set ps' \<turnstile> lhs \<Rightarrow>* rhs"
+    and "binarizeNt A B1 B2 S ps ps'"
+  shows "set ps \<turnstile> substsNt A [Nt B1, Nt B2] lhs \<Rightarrow>* substsNt A [Nt B1, Nt B2] rhs"
+  using assms 
+proof (induction rhs rule: rtranclp_induct)
+  case (step y z)
+  then show ?case 
+    using lemma1Nt[of A B1 B2 S ps ps' y z] by auto
+qed simp
+
+lemma lemma4:
+  assumes "uniformize A t S ps ps'" 
+  shows "lang ps' S \<subseteq> lang ps S"
+proof 
+  fix w
+  assume "w \<in> lang ps' S"
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    unfolding Lang_def by simp
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    using assms unfolding uniformize_def by auto
+  hence "set ps \<turnstile> substsNt A [Tm t] [Nt S] \<Rightarrow>* substsNt A [Tm t] (map Tm w)"
+    using assms lemma3[of ps' \<open>[Nt S]\<close> \<open>map Tm w\<close>] by blast
+  moreover have "substsNt A [Tm t] [Nt S] = [Nt S]"
+    using assms fresh_nts_single[of ps S] unfolding uniformize_def by auto
+  moreover have "substsNt A [Tm t] (map Tm w) = map Tm w" 
+    by simp
+  ultimately show "w \<in> lang ps S" 
+    by (simp add: Lang_def)
+qed
+
+lemma lemma4Nt:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+  shows "lang ps' S \<subseteq> lang ps S"
+proof
+  fix w
+  assume "w \<in> lang ps' S"
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    by (simp add: Lang_def)
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    using assms unfolding binarizeNt_def by auto
+  hence "set ps \<turnstile> substsNt A [Nt B1, Nt B2] [Nt S] \<Rightarrow>*  substsNt A [Nt B1, Nt B2] (map Tm w)"
+    using assms lemma3Nt[of ps' \<open>[Nt S]\<close> \<open>map Tm w\<close>] by blast
+  moreover have "substsNt A [Nt B1, Nt B2] [Nt S] = [Nt S]"
+    using assms fresh_nts_single[of ps S] unfolding binarizeNt_def by auto
+  moreover have "substsNt A [Nt B1, Nt B2] (map Tm w) = map Tm w" by simp
+  ultimately show "w \<in> lang ps S" using Lang_def
+    by (metis (no_types, lifting) mem_Collect_eq)
+qed
+
+lemma slemma5_1:
+  assumes "set ps \<turnstile> u \<Rightarrow>* v"
+    and "uniformize A t S ps ps'"
+  shows "set ps' \<turnstile> u \<Rightarrow>* v"
+  using assms by (induction v rule: rtranclp_induct) (auto simp: cnf_r1Tm rtranclp_trans)
+
+lemma slemma5_1Nt:
+  assumes "set ps \<turnstile> u \<Rightarrow>* v"
+    and "binarizeNt A B1 B2 S ps ps'"
+  shows "set ps' \<turnstile> u \<Rightarrow>* v"
+  using assms by (induction v rule: rtranclp_induct) (auto simp: cnf_r1Nt rtranclp_trans)
+
+lemma lemma5: 
+  assumes "uniformize A t S ps ps'"
+  shows "lang ps S \<subseteq> lang ps' S"
+proof 
+  fix w
+  assume "w \<in> lang ps S"
+  hence "set ps \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    using assms unfolding Lang_def uniformize_def by auto 
+  thus "w \<in> lang ps' S" 
+    using assms slemma5_1 Lang_def by fastforce
+qed 
+
+lemma lemma5Nt: 
+  assumes "binarizeNt A B1 B2 S ps ps'"
+  shows "lang ps S \<subseteq> lang ps' S"
+proof 
+  fix w
+  assume "w \<in> lang ps S"
+  hence "set ps \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    using assms unfolding Lang_def binarizeNt_def by auto 
+  thus "w \<in> lang ps' S" 
+    using assms slemma5_1Nt Lang_def by fast
+qed 
+
+lemma cnf_lemma1: "uniformize A t S ps ps' \<Longrightarrow> lang ps S = lang ps' S"
+  using lemma4 lemma5 by fast
+
+lemma cnf_lemma1Nt: "binarizeNt A B1 B2 S ps ps' \<Longrightarrow> lang ps S = lang ps' S"
+  using lemma4Nt lemma5Nt by fast
+
+lemma uniformizeRtc_Eps_free: 
+  assumes "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps'"
+    and "Eps_free (set ps)"
+  shows "Eps_free (set ps')"
+  using assms by (induction rule: rtranclp_induct) (auto simp: uniformize_Eps_free)
+
+lemma binarizeNtRtc_Eps_free:
+  assumes "(\<lambda>x y. \<exists>A t B1 B2. binarizeNt A B1 B2 S x y)^** ps ps'"
+    and "Eps_free (set ps)"
+  shows "Eps_free (set ps')"
+  using assms by (induction rule: rtranclp_induct) (auto simp: binarizeNt_Eps_free)
+
+lemma uniformizeRtc_Unit_free: 
+  assumes "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps'"
+    and "Unit_free (set ps)"
+  shows "Unit_free (set ps')"
+  using assms by (induction rule: rtranclp_induct) (auto simp: uniformize_Unit_free)
+
+lemma binarizeNtRtc_Unit_free:
+  assumes "(\<lambda>x y. \<exists>A t B1 B2. binarizeNt A B1 B2 S x y)^** ps ps'"
+    and "Unit_free (set ps)"
+  shows "Unit_free (set ps')"
+  using assms by (induction rule: rtranclp_induct) (auto simp: binarizeNt_Unit_free)
+
+(* proofs about Nts *)
+
+lemma uniformize_Nts: 
+  assumes "uniformize A t S ps ps'" "S \<in> Nts (set ps)"
+  shows "S \<in> Nts (set ps')"
+proof -
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Tm t]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps)) 
+      \<and> set ps' = ((set ps - {(l,r)}) \<union> {(A,[Tm t]), (l, p@[Nt A]@s)})"
+    using assms(1) set_removeAll fresh_nts_single[of ps S] unfolding uniformize_def by fastforce
+  thus ?thesis
+  proof (cases "S \<in> Nts {(l,r)}")
+    case True
+    hence "S \<in> Nts {(A,[Tm t]), (l, p@[Nt A]@s)}"
+      unfolding Nts_def nts_syms_def using lrps by auto
+    then show ?thesis using  lrps Nts_Un by (metis UnCI)
+  next
+    case False
+    hence "S \<in> Nts (set ps - {(l,r)})"
+      unfolding Nts_def using lrps 
+      by (metis UnCI UnE Un_Diff_cancel2 assms(2) Nts_Un Nts_def)
+    then show ?thesis 
+      by (simp add: lrps Nts_def)
+  qed
+qed  
+
+lemma uniformizeRtc_Nts: 
+  assumes "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps'" "S \<in> Nts (set ps)"
+  shows "S \<in> Nts (set ps')"
+  using assms by (induction rule: rtranclp_induct) (auto simp: uniformize_Nts)
+
+(* Termination *)
+
+theorem cnf_lemma2: 
+  assumes "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps'"
+  shows "lang ps S = lang ps' S"
+  using assms by (induction rule: rtranclp_induct) (fastforce simp: cnf_lemma1)+ 
+
+theorem cnf_lemma2Nt: 
+  assumes "(\<lambda>x y. \<exists>A t B1 B2. binarizeNt A B1 B2 S x y)^** ps ps'"
+  shows "lang ps S = lang ps' S"
+  using assms by (induction rule: rtranclp_induct) (fastforce simp: cnf_lemma1Nt)+
+
+theorem cnf_lemma: 
+  assumes "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps'"
+    and "(\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)^** ps' ps''"
+  shows "lang ps S = lang ps'' S"
+  using assms cnf_lemma2 cnf_lemma2Nt uniformizeRtc_Nts by fastforce
+
+(* Part 2 *)
+lemma badTmsCount_append: "badTmsCount (ps@ps') = badTmsCount ps + badTmsCount ps'"
+by auto
+
+lemma badNtsCount_append: "badNtsCount (ps@ps') = badNtsCount ps + badNtsCount ps'"
+by auto
+
+lemma badTmsCount_removeAll: 
+  assumes "prodTms p > 0" "p \<in> set ps"
+  shows "badTmsCount (removeAll p ps) < badTmsCount ps"
+  using assms by (induction ps) fastforce+
+
+lemma badNtsCount_removeAll: 
+  assumes "prodNts p > 0" "p \<in> set ps"
+  shows "badNtsCount (removeAll p ps) < badNtsCount ps"
+  using assms by (induction ps) fastforce+
+
+lemma badTmsCount_removeAll2:
+  assumes "prodTms p > 0" "p \<in> set ps" "prodTms p' < prodTms p"
+  shows "badTmsCount (removeAll p ps) + prodTms p' < badTmsCount ps"
+  using assms by (induction ps) fastforce+
+
+lemma badNtsCount_removeAll2:
+  assumes "prodNts p > 0" "p \<in> set ps" "prodNts p' < prodNts p"
+  shows "badNtsCount (removeAll p ps) + prodNts p' < badNtsCount ps"
+  using assms by (induction ps) fastforce+
+
+lemma lemma6_a:
+  assumes "uniformize A t S ps ps'" shows "badTmsCount (ps') < badTmsCount ps"
+proof -
+  from assms obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Tm t]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps)) 
+      \<and> ps' = ((removeAll (l,r) ps) @ [(A,[Tm t]), (l, p@[Nt A]@s)])"
+    using fresh_nts_single[of ps S] unfolding uniformize_def by auto
+  hence "prodTms (l,p@[Tm t]@s) = length (filter (isTm) (p@[Tm t]@s))"
+    unfolding prodTms_def by auto
+  hence 1: "prodTms (l,p@[Tm t]@s) = Suc (length (filter (isTm) (p@s)))"
+    by (simp add: isTm_def)
+  have 2: "badTmsCount ps' = badTmsCount (removeAll (l,r) ps) + badTmsCount [(A,[Tm t])] + badTmsCount [(l, p@[Nt A]@s)]"
+    using lrps by (auto simp: badTmsCount_append)
+  have 3: "badTmsCount (removeAll (l,r) ps) < badTmsCount ps"
+    using 1 badTmsCount_removeAll lrps gr0_conv_Suc by blast
+  have "prodTms (l, p@[Nt A]@s) = (length (filter (isTm) (p@[Nt A]@s))) \<or> prodTms (l, p@[Nt A]@s) = 0"
+    unfolding prodTms_def using lrps by simp
+  thus ?thesis
+  proof 
+    assume "prodTms (l, p@[Nt A]@s) = (length (filter (isTm) (p@[Nt A]@s)))"
+    hence "badTmsCount ps' = badTmsCount (removeAll (l,r) ps) + prodTms (l, p@[Nt A]@s)"
+      using 2 by (simp add: prodTms_def)
+    moreover have "prodTms (l,p@[Nt A]@s) < prodTms (l,p@[Tm t]@s)"
+      using 1 \<open>prodTms (l, p @ [Nt A] @ s) = length (filter isTm (p @ [Nt A] @ s))\<close> isTm_def by force 
+    ultimately show "badTmsCount ps' < badTmsCount ps" 
+      using badTmsCount_removeAll2[of "(l,r)" ps "(l,p @[Nt A]@s)"] lrps 1 by auto
+  next 
+    assume "prodTms (l, p@[Nt A]@s) = 0"
+    hence "badTmsCount ps' = badTmsCount (removeAll (l,r) ps)"
+      using 2 by (simp add: prodTms_def)
+    thus "badTmsCount ps' < badTmsCount ps" 
+      using 3 by simp
+  qed
+qed
+
+lemma lemma6_b:
+  assumes "binarizeNt A B1 B2 S ps ps'" shows "badNtsCount ps' < badNtsCount ps"
+proof -
+  from assms obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps))
+    \<and> ps' = ((removeAll (l,r) ps) @ [(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)])"
+    using fresh_nts_single[of ps S] unfolding binarizeNt_def by auto
+  hence "prodNts (l,p@[Nt B1,Nt B2]@s) = length (filter (isNt) (p@[Nt B1,Nt B2]@s))"
+    unfolding prodNts_def by auto
+  hence 1: "prodNts (l,p@[Nt B1,Nt B2]@s) = Suc (Suc (length (filter (isNt) (p@s))))"
+    by (simp add: isNt_def)
+  have 2: "badNtsCount ps' = badNtsCount (removeAll (l,r) ps) + badNtsCount [(A, [Nt B1,Nt B2])] + badNtsCount [(l, (p@[Nt A]@s))]"
+    using lrps by (auto simp: badNtsCount_append prodNts_def)
+  have 3: "badNtsCount (removeAll (l,r) ps) < badNtsCount ps"
+    using lrps badNtsCount_removeAll 1 by force
+  have "prodNts (l, p@[Nt A]@s) = length (filter (isNt) (p@[Nt A]@s)) \<or> prodNts (l, p@[Nt A]@s) = 0"
+    unfolding prodNts_def using lrps by simp
+  thus ?thesis 
+  proof 
+    assume "prodNts (l, p@[Nt A]@s) = length (filter (isNt) (p@[Nt A]@s))"
+    hence "badNtsCount ps' = badNtsCount (removeAll (l,r) ps) + badNtsCount [(l, (p@[Nt A]@s))]"
+      using 2 by (simp add: prodNts_def)
+    moreover have "prodNts (l, p@[Nt A]@s) < prodNts (l,p@[Nt B1,Nt B2]@s)"
+      using 1 \<open>prodNts (l, p@[Nt A]@s) = length (filter (isNt) (p@[Nt A]@s))\<close> isNt_def by simp
+    ultimately show ?thesis 
+      using badNtsCount_removeAll2[of "(l,r)" ps "(l, (p@[Nt A]@s))"] 1 lrps by auto
+  next 
+    assume "prodNts (l, p@[Nt A]@s) = 0"
+    hence "badNtsCount ps' = badNtsCount (removeAll (l,r) ps)"
+      using 2 by (simp add: prodNts_def)
+    thus ?thesis 
+      using 3 by simp
+  qed
+qed
+
+lemma badTmsCount0_removeAll: "badTmsCount ps = 0 \<Longrightarrow> badTmsCount (removeAll (l,r) ps) = 0" 
+by auto 
+
+lemma slemma15_a:
+  assumes "binarizeNt A B1 B2 S ps ps'"
+    and "badTmsCount ps = 0"
+  shows "badTmsCount ps' = 0"
+proof -
+  obtain l r p s where lrps: "(l,r) \<in> set ps \<and> (r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> []) \<and> (A \<notin> Nts (set ps))
+    \<and> (ps' = ((removeAll (l,r) ps) @ [(A, [Nt B1,Nt B2]), (l, p@[Nt A]@s)]))"
+    using assms(1) fresh_nts_single[of ps S] unfolding binarizeNt_def by auto
+  hence "badTmsCount ps' = badTmsCount (removeAll (l,r) ps) + badTmsCount [(l, (p@[Nt A]@s))]"
+    by (auto simp: badTmsCount_append prodTms_def isTm_def)
+  moreover have "badTmsCount (removeAll (l,r) ps) = 0"
+    using badTmsCount0_removeAll[of ps l r] assms(2) by simp
+  moreover have "badTmsCount [(l, (p@[Nt A]@s))] = 0" 
+  proof -
+    have "prodTms (l,p@[Nt B1,Nt B2]@s) = 0"
+      using lrps assms(2) badTmsCountSet by auto
+    thus "badTmsCount [(l, (p@[Nt A]@s))] = 0"
+      by (auto simp: isTm_def prodTms_def)
+  qed
+  ultimately show ?thesis 
+    by simp
+qed
+
+lemma lemma15_a:
+  assumes "(\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)^** ps ps'"
+    and "badTmsCount ps = 0"
+  shows "badTmsCount ps' = 0"
+  using assms by (induction) (auto simp: slemma15_a simp del: badTmsCount.simps)
+
+lemma noTms_prodTms0:
+  assumes "prodTms (l,r) = 0"
+  shows "length r \<le> 1 \<or> (\<forall>a \<in> set r. isNt a)"
+proof -
+  have "length r \<le> 1 \<or> (\<nexists>a. a \<in> set r \<and> isTm a)"
+    using assms unfolding prodTms_def using empty_filter_conv by fastforce
+  thus ?thesis 
+    by (metis isNt_def isTm_def sym.exhaust)
+qed
+
+lemma badTmsCountNot0:
+  assumes "badTmsCount ps > 0"
+  shows "\<exists>l r t. (l,r) \<in> set ps \<and> length r \<ge> 2 \<and> Tm t \<in> set r"
+proof -
+  have "\<exists>p \<in> set ps. prodTms p > 0"
+    using assms badTmsCountSet not_gr0 by blast
+  from this obtain l r where lr: "(l, r) \<in> set ps \<and> prodTms (l,r) > 0"
+    by auto
+  hence 1: "length r \<ge> 2"
+    unfolding prodTms_def using not_le_imp_less by fastforce
+  hence "prodTms (l,r) = length (filter (isTm) r)"
+    unfolding prodTms_def by simp
+  hence "\<exists>t. Tm t \<in> set r"
+    by (metis lr empty_filter_conv isTm_def length_greater_0_conv)
+  thus ?thesis using lr 1 by blast
+qed
+
+lemma badNtsCountNot0: 
+  assumes "badNtsCount ps > 0" 
+  shows "\<exists>l r. (l, r) \<in> set ps \<and> length r \<ge> 3"
+proof -
+  have "\<exists>p \<in> set ps. prodNts p > 0"
+    using assms badNtsCountSet not_gr0 by blast
+  from this obtain l r where lr: "(l, r) \<in> set ps \<and> prodNts (l,r) > 0"
+    by auto
+  hence "length r \<ge> 3"
+    unfolding prodNts_def using not_le_imp_less by fastforce
+  thus ?thesis using lr by auto
+qed
+
+lemma list_longer2: "length l \<ge> 2 \<and> x \<in> set l \<Longrightarrow> (\<exists>hd tl . l = hd@[x]@tl \<and> (hd \<noteq> [] \<or> tl \<noteq> []))"
+  using split_list_last by fastforce 
+
+lemma list_longer3: "length l \<ge> 3 \<Longrightarrow> (\<exists>hd tl x y. l = hd@[x]@[y]@tl \<and> (hd \<noteq> [] \<or> tl \<noteq> []))"
+  by (metis Suc_le_length_iff append.left_neutral append_Cons neq_Nil_conv numeral_3_eq_3)
+
+lemma lemma8_a: "badTmsCount ps > 0 \<Longrightarrow> \<exists>ps' A t. uniformize A t S ps ps'"
+proof -
+  assume "badTmsCount ps > 0"
+  then obtain l r t where lr: "(l,r) \<in> set ps \<and> length r \<ge> 2 \<and> Tm t \<in> set r"
+    using badTmsCountNot0 by blast
+  from this obtain p s where ps: "r = p@[Tm t]@s \<and> (p \<noteq> [] \<or> s \<noteq> [])"
+    unfolding isTm_def using lr list_longer2[of r] by blast
+  from this obtain A ps' where "A = fresh(nts ps \<union> {S}) \<and> ps' = removeAll (l, r) ps @ [(A, [Tm t]), (l, p @ [Nt A] @ s)]" 
+    by auto
+  hence "uniformize A t S ps ps'"
+    unfolding uniformize_def using lr ps by auto
+  thus ?thesis by blast
+qed
+
+lemma lemma8_b:
+  assumes "badTmsCount ps = 0" and "badNtsCount ps > 0"
+  shows "\<exists>ps' A B1 B2. binarizeNt A B1 B2 S ps ps'"
+proof -
+  obtain l r where lr: "(l, r) \<in> set ps \<and> length r \<ge> 3"
+    using assms(2) badNtsCountNot0 by blast
+  obtain p s X Y where psXY: "r = p@[X]@[Y]@s \<and> (p \<noteq> [] \<or> s \<noteq> [])"
+    using lr list_longer3[of r] by blast
+  have "(\<forall>a \<in> set r. isNt a)"
+    using lr assms(1) badTmsCountSet[of ps] noTms_prodTms0[of l r] by fastforce
+  from this obtain B1 B2 where "X = Nt B1 \<and> Y = Nt B2"
+    using isNt_def psXY by fastforce
+  hence B: "(r = p@[Nt B1,Nt B2]@s) \<and> (p \<noteq> [] \<or> s \<noteq> [])"
+    using psXY by auto
+  from this obtain A ps' where "A = fresh(nts ps \<union> {S}) \<and> ps' = removeAll (l, r) ps @ [(A, [Nt B1, Nt B2]), (l, p @ [Nt A] @ s)]"
+    by simp
+  hence "binarizeNt A B1 B2 S ps ps'"
+    unfolding binarizeNt_def using lr B by auto
+  thus ?thesis by blast
+qed
+
+lemma uniformize_2: "\<exists>ps'. (\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps' \<and> (badTmsCount ps' = 0)"
+proof (induction "badTmsCount ps" arbitrary: ps rule: less_induct)
+  case less
+  then show ?case
+  proof (cases "badTmsCount ps = 0")
+    case False
+    from this obtain ps' A t where g': "uniformize A t S ps ps'"
+      using lemma8_a by blast
+    hence "badTmsCount ps' < badTmsCount ps"
+      using lemma6_a[of A t] by blast
+    from this obtain ps'' where "(\<lambda>x y. \<exists>A t. uniformize A t S x y)** ps' ps'' \<and> badTmsCount ps'' = 0"
+      using less by blast
+    thus ?thesis 
+      using g' converse_rtranclp_into_rtranclp[of "(\<lambda>x y. \<exists>A t. uniformize A t S x y)" ps ps' ps''] by blast
+  qed blast
+qed
+
+lemma binarizeNt_2: 
+  assumes "badTmsCount ps = 0"
+    shows "\<exists>ps'. (\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)^** ps ps' \<and> (badNtsCount ps' = 0)"
+using assms proof (induction "badNtsCount ps" arbitrary: ps rule: less_induct)
+  case less
+  then show ?case 
+  proof (cases "badNtsCount ps = 0")
+    case False
+    from this obtain ps' A B1 B2 where g': "binarizeNt A B1 B2 S ps ps'"
+      using assms lemma8_b less(2) by blast
+    hence "badNtsCount ps' < badNtsCount ps"
+      using lemma6_b by blast
+    from this obtain ps'' where "(\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)** ps' ps'' \<and> badNtsCount ps'' = 0"
+      using less slemma15_a[of A B1 B2 S ps ps'] g' by blast
+    then show ?thesis 
+      using g' converse_rtranclp_into_rtranclp[of "(\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)" ps ps' ps''] by blast
+  qed blast
+qed
+
+theorem cnf_noe_nou:  fixes ps :: "('n::infinite,'t)prods"
+  assumes "Eps_free (set ps)" and "Unit_free (set ps)"
+  shows "\<exists>ps'::('n,'t)prods. uniform (set ps') \<and> binary (set ps') \<and> lang ps S = lang ps' S \<and> Eps_free (set ps') \<and> Unit_free (set ps')"
+proof -
+  obtain ps' where g': "(\<lambda>x y. \<exists>A t. uniformize A t S x y)^** ps ps' \<and> badTmsCount ps' = 0 \<and> Eps_free (set ps') \<and> Unit_free (set ps')"
+    using assms uniformize_2 uniformizeRtc_Eps_free uniformizeRtc_Unit_free by blast
+  obtain ps'' where g'': "(\<lambda>x y. \<exists>A B1 B2. binarizeNt A B1 B2 S x y)^** ps' ps'' \<and> (badNtsCount ps'' = 0) \<and> (badTmsCount ps'' = 0)"
+    using g' binarizeNt_2 lemma15_a by blast
+  hence "uniform (set ps'') \<and> binary (set ps'') \<and> Eps_free (set ps'') \<and> Unit_free (set ps'')"
+    using g' count_bin_un binarizeNtRtc_Eps_free binarizeNtRtc_Unit_free by fastforce
+  moreover have "lang ps S = lang ps'' S"
+    using g' g'' cnf_lemma by blast
+  ultimately show ?thesis by blast
+qed
+
+text \<open>Alternative form more similar to the one Jana Hofmann used:\<close>
+
+lemma CNF_eq: "CNF P \<longleftrightarrow> (uniform P \<and> binary P \<and> Eps_free P \<and> Unit_free P)"
+proof 
+  assume "CNF P"
+  hence "Eps_free P"
+    unfolding CNF_def Eps_free_def by fastforce
+  moreover have "Unit_free P"
+    using \<open>CNF P\<close> unfolding CNF_def Unit_free_def isNt_def isTm_def by fastforce
+  moreover have "uniform P"
+  proof -
+    have "\<forall>(A,\<alpha>) \<in> P. (\<exists>B C. \<alpha> = [Nt B, Nt C]) \<or> (\<exists>t. \<alpha> = [Tm t])"
+      using \<open>CNF P\<close> unfolding CNF_def.
+    hence "\<forall>(A, \<alpha>) \<in> P. (\<forall>N \<in> set \<alpha>. isNt N) \<or> (\<exists>t. \<alpha> = [Tm t])"
+      unfolding isNt_def by fastforce
+    hence "\<forall>(A, \<alpha>) \<in> P. (\<nexists>t. Tm t \<in> set \<alpha>) \<or> (\<exists>t. \<alpha> = [Tm t])"
+      by (auto simp: isNt_def)
+    thus "uniform P"
+      unfolding uniform_def.
+  qed
+  moreover have "binary P"
+    using \<open>CNF P\<close> unfolding binary_def CNF_def by auto
+  ultimately show "uniform P \<and> binary P \<and> Eps_free P \<and> Unit_free P"
+    by blast
+next 
+  assume assm: "uniform P \<and> binary P \<and> Eps_free P \<and> Unit_free P"
+  have"\<forall>p \<in> P. (\<exists>B C. (snd p) = [Nt B, Nt C]) \<or> (\<exists>t. (snd p) = [Tm t])"
+  proof
+    fix p assume "p \<in> P"
+    obtain A \<alpha> where A\<alpha>: "(A, \<alpha>) = p"
+      by (metis prod.exhaust_sel)
+    hence "length \<alpha> = 1 \<or> length \<alpha> = 2"
+      using assm \<open>p \<in> P\<close> order_neq_le_trans unfolding binary_def Eps_free_def by fastforce
+    hence "(\<exists>B C. (snd p) = [Nt B, Nt C]) \<or> (\<exists>t. \<alpha> = [Tm t])"
+    proof 
+      assume "length \<alpha> = 1"
+      hence "\<exists>S. \<alpha> = [S]"
+        by (simp add: length_Suc_conv)
+      moreover have "\<nexists>N. \<alpha> = [Nt N]"
+        using assm A\<alpha> \<open>p \<in> P\<close> unfolding Unit_free_def by blast
+      ultimately show ?thesis by (metis sym.exhaust)
+    next
+      assume "length \<alpha> = 2"
+      obtain B C where BC: "\<alpha> = [B, C]"
+        using \<open>length \<alpha> = 2\<close> by (metis One_nat_def Suc_1 diff_Suc_1 length_0_conv length_Cons neq_Nil_conv)
+      have "(\<nexists>t. Tm t \<in> set \<alpha>)"
+        using \<open>length \<alpha> = 2\<close> assm A\<alpha> \<open>p \<in> P\<close> unfolding uniform_def by auto
+      hence "(\<forall>N \<in> set \<alpha>. \<exists>A. N = Nt A)"
+        by (metis sym.exhaust)
+      hence "\<exists>B' C'. B = Nt B' \<and> C = Nt C'" 
+        using BC by simp
+      thus ?thesis using A\<alpha> BC by auto
+    qed
+    thus "(\<exists>B C. (snd p) = [Nt B, Nt C]) \<or> (\<exists>t. (snd p) = [Tm t])" using A\<alpha> by auto
+  qed
+  thus "CNF P" by (auto simp: CNF_def)
+qed
+
+text \<open>Main Theorem: existence of CNF with the same language except for the empty word \<open>[]\<close>:\<close>
+
+theorem cnf_exists:
+  fixes ps :: "('n::infinite,'t) prods"
+  shows "\<exists>ps'::('n,'t)prods. CNF(set ps') \<and> lang ps' S = lang ps S - {[]}"
+proof -
+  obtain ps0 where ps0: "eps_elim_rel ps ps0"
+    using eps_elim_rel_exists by blast
+  obtain psu where psu: "unit_elim_rel ps0 psu"
+    using unit_elim_rel_exists by blast
+  hence 1: "Eps_free (set psu) \<and> Unit_free (set psu)"
+    using ps0 psu Eps_free_if_eps_elim_rel Unit_free_if_unit_elim_rel unit_elim_rel_Eps_free by fastforce
+  have 2: "lang psu S = lang ps S - {[]}"
+    using psu eps_elim_rel_lang_eq[OF ps0] unit_elim_rel_lang_eq[OF psu] by (simp add: eps_elim_rel_lang_eq)
+  obtain ps'::"('n,'t)prods" where g': "uniform (set ps') \<and> binary (set ps') \<and> lang psu S = lang ps' S \<and> Eps_free (set ps') \<and> Unit_free (set ps')"
+    using 1 cnf_noe_nou by blast
+  hence "CNF (set ps')" 
+    using g' CNF_eq by blast
+  moreover have "lang ps' S = lang ps S - {[]}"
+    using 2 g' by blast
+  ultimately show ?thesis by blast
+qed
+
+
+text \<open>Some helpful properties:\<close>
+
+lemma cnf_length_derive: 
+  assumes "CNF P" "P \<turnstile> [Nt S] \<Rightarrow>* \<alpha>"
+  shows "length \<alpha> \<ge> 1"
+  using assms CNF_eq Eps_free_derives_Nil length_greater_0_conv less_eq_Suc_le by auto
+
+lemma cnf_length_derive2: 
+  assumes "CNF P" "P \<turnstile> [Nt A, Nt B] \<Rightarrow>* \<alpha>"
+  shows "length \<alpha> \<ge> 2"
+proof -
+  obtain u v where uv: "P \<turnstile> [Nt A] \<Rightarrow>* u \<and> P \<turnstile> [Nt B] \<Rightarrow>* v \<and> \<alpha> = u @ v"
+    using assms(2) derives_append_decomp[of P \<open>[Nt A]\<close> \<open>[Nt B]\<close> \<alpha>] by auto
+  hence "length u \<ge> 1 \<and> length v \<ge> 1" 
+    using cnf_length_derive[OF assms(1)] by blast
+  thus ?thesis
+    using uv by simp
+qed
+
+lemma cnf_single_derive:
+  assumes "CNF P" "P \<turnstile> [Nt S] \<Rightarrow>* [Tm t]"
+  shows "(S, [Tm t]) \<in> P"
+proof -
+  obtain \<alpha> where \<alpha>: "P \<turnstile> [Nt S] \<Rightarrow> \<alpha> \<and> P \<turnstile> \<alpha> \<Rightarrow>* [Tm t]"
+    using converse_rtranclpE[OF assms(2)] by auto
+  hence 1: "(S, \<alpha>) \<in> P" 
+    by (simp add: derive_singleton)
+  have "\<nexists>A B. \<alpha> = [Nt A, Nt B]"
+  proof (rule ccontr)
+    assume "\<not> (\<nexists>A B. \<alpha> = [Nt A, Nt B])"
+    from this obtain A B where AB: "\<alpha> = [Nt A, Nt B]"
+      by blast
+    have "\<forall>w. P \<turnstile> [Nt A, Nt B] \<Rightarrow>* w \<longrightarrow> length w \<ge> 2"
+      using cnf_length_derive2[OF assms(1)] by simp
+    moreover have "length [Tm t] = 1"
+      by simp
+    ultimately show False
+      using \<alpha> AB by auto
+  qed
+  from this obtain a where "\<alpha> = [Tm a]"
+    using 1 assms(1) unfolding CNF_def by auto
+  hence "t = a"
+    using \<alpha> by (simp add: derives_Tm_Cons)
+  thus ?thesis 
+    using 1 \<open>\<alpha> = [Tm a]\<close> by blast
+qed
+
+lemma cnf_word:
+  assumes "CNF P" "P \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+    and "length w \<ge> 2"
+  shows "\<exists>A B u v. (S, [Nt A, Nt B]) \<in> P \<and> P \<turnstile> [Nt A] \<Rightarrow>* map Tm u \<and> P \<turnstile> [Nt B] \<Rightarrow>* map Tm v \<and> u@v = w \<and> u \<noteq> [] \<and> v \<noteq> []"
+proof -
+  have 1: "(S, map Tm w) \<notin> P"
+    using assms(1) assms(3) unfolding CNF_def by auto
+  have "\<exists>\<alpha>. P \<turnstile> [Nt S] \<Rightarrow> \<alpha> \<and> P \<turnstile> \<alpha> \<Rightarrow>* map Tm w"
+    using converse_rtranclpE[OF assms(2)] by auto
+  from this obtain \<alpha> where \<alpha>: "(S, \<alpha>) \<in> P \<and> P \<turnstile> \<alpha> \<Rightarrow>* map Tm w"
+    by (auto simp: derive_singleton)
+  hence "(\<nexists>t. \<alpha> = [Tm t])"
+    using 1 derives_Tm_Cons[of P] derives_from_empty by auto
+  hence "\<exists>A B. P \<turnstile> [Nt S] \<Rightarrow> [Nt A, Nt B] \<and> P \<turnstile> [Nt A, Nt B] \<Rightarrow>* map Tm w"
+    using assms(1) \<alpha> derive_singleton[of P \<open>Nt S\<close> \<alpha>] unfolding CNF_def by fast
+  from this obtain A B where AB: "(S, [Nt A, Nt B]) \<in> P \<and> P \<turnstile> [Nt A, Nt B] \<Rightarrow>* map Tm w"
+    using derive_singleton[of P \<open>Nt S\<close>] by blast
+  hence "\<not>(P \<turnstile> [Nt A] \<Rightarrow>* []) \<and> \<not>(P \<turnstile> [Nt B] \<Rightarrow>* [])"
+    using assms(1) CNF_eq Eps_free_derives_Nil by blast
+  from this obtain u v where uv: "P \<turnstile> [Nt A] \<Rightarrow>* u \<and> P \<turnstile> [Nt B] \<Rightarrow>* v \<and> u@v = map Tm w \<and> u \<noteq> [] \<and> v \<noteq> []"
+    using AB derives_append_decomp[of P \<open>[Nt A]\<close> \<open>[Nt B]\<close> \<open>map Tm w\<close>] by force
+  moreover have "\<exists>u' v'. u = map Tm u' \<and> v = map Tm v'"
+    using uv map_eq_append_conv[of Tm w u v] by auto
+  ultimately show ?thesis
+    using AB append_eq_map_conv[of u v Tm w] list.simps(8)[of Tm] by fastforce
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Context_Free_Grammar.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Context_Free_Grammar.thy
@@ -0,0 +1,1143 @@
+(*
+Authors: Tobias Nipkow, Akihisa Yamada
+*)
+
+section "Context-Free Grammars"
+
+theory Context_Free_Grammar
+imports "HOL-Library.Infinite_Typeclass"
+begin
+
+(* TODO: make function fresh executable *)
+definition fresh :: "('n::infinite) set \<Rightarrow> 'n" where
+"fresh A = (SOME x. x \<notin> A)"
+
+lemma fresh_finite: "finite A \<Longrightarrow> fresh A \<notin> A"
+unfolding fresh_def by (metis arb_element someI)
+
+declare relpowp.simps(2)[simp del]
+
+lemma bex_pair_conv: "(\<exists>(x,y) \<in> R. P x y) \<longleftrightarrow> (\<exists>x y. (x,y) \<in> R \<and> P x y)"
+  by auto
+
+lemma in_image_map_prod: "fgp \<in> map_prod f g ` R \<longleftrightarrow> (\<exists>(x,y)\<in>R. fgp = (f x,g y))"
+  by auto
+
+
+subsection "Symbols and Context-Free Grammars"
+
+text \<open>Most of the theory is based on arbitrary sets of productions.
+Finiteness of the set of productions is only required where necessary.
+Finiteness of the type of terminal symbols is only required where necessary.
+Whenever fresh nonterminals need to be invented, the type of nonterminals is assumed to be infinite.\<close>
+
+datatype ('n,'t) sym = Nt 'n | Tm 't
+
+type_synonym ('n,'t) syms = "('n,'t) sym list"
+
+type_synonym ('n,'t) prod = "'n \<times> ('n,'t) syms"
+
+type_synonym ('n,'t) prods = "('n,'t) prod list"
+type_synonym ('n,'t) Prods = "('n,'t) prod set"
+
+datatype ('n,'t) cfg = cfg (prods : "('n,'t) prods") (start : "'n")
+datatype ('n,'t) Cfg = Cfg (Prods : "('n,'t) Prods") (Start : "'n")
+
+definition isTm :: "('n, 't) sym \<Rightarrow> bool" where 
+"isTm S = (\<exists>a. S = Tm a)"
+
+definition isNt :: "('n, 't) sym \<Rightarrow> bool" where 
+"isNt S = (\<exists>A. S = Nt A)"
+
+fun destTm :: "('n, 't) sym  \<Rightarrow> 't" where 
+\<open>destTm (Tm a) = a\<close>
+
+lemma isTm_simps[simp]:
+  \<open>isTm (Nt A) = False\<close>
+  \<open>isTm (Tm a)\<close> 
+by (simp_all add: isTm_def)
+
+lemma filter_isTm_map_Tm[simp]: \<open>filter isTm (map Tm xs) = map Tm xs\<close>
+by(induction xs) auto
+
+lemma destTm_o_Tm[simp]: \<open>destTm \<circ> Tm = id\<close>
+by auto
+
+definition nts_syms :: "('n,'t)syms \<Rightarrow> 'n set" where
+"nts_syms w = {A. Nt A \<in> set w}"
+
+definition tms_syms :: "('n,'t)syms \<Rightarrow> 't set" where
+"tms_syms w = {a. Tm a \<in> set w}"
+
+definition Nts :: "('n,'t)Prods \<Rightarrow> 'n set" where
+  "Nts P = (\<Union>(A,w)\<in>P. {A} \<union> nts_syms w)"
+
+definition Tms :: "('n,'t)Prods \<Rightarrow> 't set" where 
+  "Tms P = (\<Union>(A,w)\<in>P. tms_syms w)"
+
+abbreviation nts :: "('n,'t) prods \<Rightarrow> 'n set" where
+  "nts P \<equiv> Nts (set P)"
+
+definition Syms :: "('n,'t)Prods \<Rightarrow> ('n,'t) sym set" where 
+  "Syms P = (\<Union>(A,w)\<in>P. {Nt A} \<union> set w)"
+
+abbreviation tms :: "('n,'t) prods \<Rightarrow> 't set" where
+  "tms P \<equiv> Tms (set P)"
+
+abbreviation syms :: "('n,'t) prods \<Rightarrow> ('n,'t) sym set" where
+  "syms P \<equiv> Syms (set P)"
+
+definition Lhss :: "('n, 't) Prods \<Rightarrow> 'n set" where
+"Lhss P = (\<Union>(A,w) \<in> P. {A})"
+
+abbreviation lhss :: "('n, 't) prods \<Rightarrow> 'n set" where
+"lhss ps \<equiv> Lhss(set ps)"
+
+definition Rhs_Nts :: "('n, 't) Prods \<Rightarrow> 'n set" where
+"Rhs_Nts P = (\<Union>(_,w)\<in>P. nts_syms w)"
+
+definition Rhss :: "('n \<times> 'a) set \<Rightarrow> 'n \<Rightarrow> 'a set" where
+"Rhss P A = {w. (A,w) \<in> P}"
+
+lemma inj_Nt: "inj Nt"
+by (simp add: inj_def)
+
+lemma map_Tm_inject_iff[simp]: "map Tm xs = map Tm ys \<longleftrightarrow> xs = ys"
+by (metis sym.inject(2) list.inj_map_strong)
+
+lemma map_Nt_eq_map_Nt_iff[simp]: "map Nt u = map Nt v \<longleftrightarrow> u = v"
+by(rule inj_map_eq_map[OF inj_Nt])
+
+lemma map_Nt_eq_map_Tm_iff[simp]: "map Nt u = map Tm v \<longleftrightarrow> u = [] \<and> v = []"
+by (cases u) auto
+
+lemmas map_Tm_eq_map_Nt_iff[simp] = eq_iff_swap[OF map_Nt_eq_map_Tm_iff]
+
+lemma nts_syms_Nil[simp]: "nts_syms [] = {}"
+unfolding nts_syms_def by auto
+
+lemma nts_syms_Cons[simp]: "nts_syms (a#v) = (case a of Nt A \<Rightarrow> {A} | _ \<Rightarrow> {}) \<union> nts_syms v"
+by (auto simp: nts_syms_def split: sym.split)
+
+lemma nts_syms_append[simp]: "nts_syms (u @ v) = nts_syms u \<union> nts_syms v"
+by (auto simp: nts_syms_def)
+
+lemma nts_syms_map_Nt[simp]: "nts_syms (map Nt w) = set w"
+unfolding nts_syms_def by auto
+
+lemma nts_syms_map_Tm[simp]: "nts_syms (map Tm w) = {}"
+unfolding nts_syms_def by auto
+
+lemma in_Nts_iff_in_Syms: "B \<in> Nts P \<longleftrightarrow> Nt B \<in> Syms P"
+unfolding Nts_def Syms_def nts_syms_def by (auto)
+
+lemma Nts_Un: "Nts (P1 \<union> P2) = Nts P1 \<union> Nts P2"
+by (simp add: Nts_def)
+
+lemma Nts_Lhss_Rhs_Nts: "Nts P = Lhss P \<union> Rhs_Nts P"
+unfolding Nts_def Lhss_def Rhs_Nts_def by auto
+
+lemma Nts_nts_syms: "w \<in> Rhss P A \<Longrightarrow> nts_syms w \<subseteq> Nts P"
+unfolding Rhss_def Nts_def by blast
+
+lemma Syms_simps[simp]:
+  "Syms {} = {}"
+  "Syms(insert (A,w) P) = {Nt A} \<union> set w \<union> Syms P"
+  "Syms(P \<union> P') = Syms P \<union> Syms P'"
+by(auto simp: Syms_def)
+
+lemma Lhss_simps[simp]:
+  "Lhss {} = {}"
+  "Lhss(insert (A,w) P) = {A} \<union> Lhss P"
+  "Lhss(P \<union> P') = Lhss P \<union> Lhss P'"
+by(auto simp: Lhss_def)
+
+
+subsubsection \<open>Finiteness Lemmas\<close>
+
+lemma finite_nts_syms: "finite (nts_syms w)"
+proof -
+  have "Nt ` {A. Nt A \<in> set w} \<subseteq> set w" by auto
+  from finite_inverse_image[OF _ inj_Nt]
+  show ?thesis unfolding nts_syms_def using finite_inverse_image[OF _ inj_Nt] by auto
+qed
+
+lemma finite_nts: "finite(nts ps)"
+unfolding Nts_def by (simp add: finite_nts_syms split_def)
+
+lemma fresh_nts: "fresh(nts ps) \<notin> nts ps"
+by(fact fresh_finite[OF finite_nts])
+
+lemma finite_nts_prods_start: "finite(nts(prods g) \<union> {start g})"
+unfolding Nts_def by (simp add: finite_nts_syms split_def)
+
+lemma fresh_nts_prods_start: "fresh(nts(prods g) \<union> {start g}) \<notin> nts(prods g) \<union> {start g}"
+by(fact fresh_finite[OF finite_nts_prods_start])
+
+lemma finite_Nts: "finite P \<Longrightarrow> finite (Nts P)"
+unfolding Nts_def by (simp add: case_prod_beta finite_nts_syms)
+
+lemma finite_Rhss: "finite P \<Longrightarrow> finite (Rhss P A)"
+unfolding Rhss_def by (metis Image_singleton finite_Image)
+
+
+subsection "Derivations"
+
+subsubsection \<open>The standard derivations \<open>\<Rightarrow>\<close>, \<open>\<Rightarrow>*\<close>, \<open>\<Rightarrow>(n)\<close>\<close>
+
+inductive derive :: "('n,'t) Prods \<Rightarrow> ('n,'t) syms \<Rightarrow> ('n,'t)syms \<Rightarrow> bool"
+  ("(2_ \<turnstile>/ (_ \<Rightarrow>/ _))" [50, 0, 50] 50) where
+"(A,\<alpha>) \<in> P \<Longrightarrow> P \<turnstile> u @ [Nt A] @ v \<Rightarrow> u @ \<alpha> @ v"
+
+abbreviation deriven ("(2_ \<turnstile>/ (_ /\<Rightarrow>'(_')/ _))" [50, 0, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>(n) v \<equiv> (derive P ^^ n) u v"
+
+abbreviation derives ("(2_ \<turnstile>/ (_/ \<Rightarrow>*/ _))" [50, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>* v \<equiv> ((derive P) ^**) u v"
+
+definition Ders :: "('n,'t)Prods \<Rightarrow> 'n \<Rightarrow> ('n,'t)syms set" where
+"Ders P A = {w. P \<turnstile> [Nt A] \<Rightarrow>* w}"
+
+abbreviation ders :: "('n,'t)prods \<Rightarrow> 'n \<Rightarrow> ('n,'t)syms set" where
+"ders ps \<equiv> Ders (set ps)"
+
+lemma DersI:
+  assumes "P \<turnstile> [Nt A] \<Rightarrow>* w" shows "w \<in> Ders P A"
+  using assms by (auto simp: Ders_def)
+
+lemma DersD:
+  assumes "w \<in> Ders P A" shows "P \<turnstile> [Nt A] \<Rightarrow>* w"
+  using assms by (auto simp: Ders_def)
+
+lemmas DersE = DersD[elim_format]
+
+definition Lang :: "('n,'t)Prods \<Rightarrow> 'n \<Rightarrow> 't list set" where
+"Lang P A = {w. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w}"
+
+abbreviation lang :: "('n,'t)prods \<Rightarrow> 'n \<Rightarrow> 't list set" where
+"lang ps A \<equiv> Lang (set ps) A"
+
+abbreviation LangS :: "('n,'t) Cfg \<Rightarrow> 't list set" where
+"LangS G \<equiv> Lang (Prods G) (Start G)"
+
+abbreviation langS :: "('n,'t) cfg \<Rightarrow> 't list set" where
+"langS g \<equiv> lang (prods g) (start g)"
+
+lemma Lang_Ders: "map Tm ` (Lang P A) \<subseteq> Ders P A"
+unfolding Lang_def Ders_def by auto
+
+lemma Lang_eqI_derives:
+  assumes "\<And>v. R \<turnstile> [Nt A] \<Rightarrow>* map Tm v \<longleftrightarrow> S \<turnstile> [Nt A] \<Rightarrow>* map Tm v"
+  shows "Lang R A = Lang S A"
+  by (auto simp: Lang_def assms)
+
+lemma derive_iff: "R \<turnstile> u \<Rightarrow> v \<longleftrightarrow> (\<exists> (A,w) \<in> R. \<exists>u1 u2. u = u1 @ Nt A # u2 \<and> v = u1 @ w @ u2)"
+  apply (rule iffI)
+   apply (induction rule: derive.induct)
+   apply (fastforce)
+  using derive.intros by fastforce 
+
+lemma not_derive_from_Tms: "\<not> P \<turnstile> map Tm as \<Rightarrow> w"
+by(auto simp add: derive_iff map_eq_append_conv)
+
+lemma deriven_from_TmsD: "P \<turnstile> map Tm as \<Rightarrow>(n) w \<Longrightarrow> w = map Tm as"
+by (metis not_derive_from_Tms relpowp_E2)
+ 
+lemma derives_from_Tms_iff: "P \<turnstile> map Tm as \<Rightarrow>* w \<longleftrightarrow> w = map Tm as"
+by (meson deriven_from_TmsD rtranclp.rtrancl_refl rtranclp_power)
+
+lemma Un_derive: "R \<union> S \<turnstile> y \<Rightarrow> z \<longleftrightarrow> R \<turnstile> y \<Rightarrow> z \<or> S \<turnstile> y \<Rightarrow> z"
+  by (fastforce simp: derive_iff)
+
+lemma derives_rule:
+  assumes 2: "(A,w) \<in> R" and 1: "R \<turnstile> x \<Rightarrow>* y @ Nt A # z" and 3: "R \<turnstile> y@w@z \<Rightarrow>* v"
+  shows "R \<turnstile> x \<Rightarrow>* v"
+proof-
+  note 1
+  also have "R \<turnstile> y @ Nt A # z \<Rightarrow> y @ w @ z"
+    using derive.intros[OF 2] by simp
+  also note 3
+  finally show ?thesis.
+qed
+
+lemma derives_Cons_rule:
+  assumes 1: "(A,w) \<in> R" and 2: "R \<turnstile> w @ u \<Rightarrow>* v" shows "R \<turnstile> Nt A # u \<Rightarrow>* v"
+  using derives_rule[OF 1, of "Nt A # u" "[]" u v] 2 by auto
+
+lemma deriven_mono: "P \<subseteq> P' \<Longrightarrow> P \<turnstile> u \<Rightarrow>(n) v \<Longrightarrow> P' \<turnstile> u \<Rightarrow>(n) v"
+by (metis Un_derive relpowp_mono subset_Un_eq)
+
+lemma derives_mono: "P \<subseteq> P' \<Longrightarrow> P \<turnstile> u \<Rightarrow>* v \<Longrightarrow> P' \<turnstile> u \<Rightarrow>* v"
+by (meson deriven_mono rtranclp_power)
+
+lemma Lang_mono: "P \<subseteq> P' \<Longrightarrow> Lang P A \<subseteq> Lang P' A"
+by (auto simp: Lang_def derives_mono)
+
+
+subsubsection "Customized Induction Principles"
+
+lemma deriven_induct[consumes 1, case_names 0 Suc]:
+  assumes "P \<turnstile> xs \<Rightarrow>(n) ys"
+  and "Q 0 xs"
+  and "\<And>n u A v w. \<lbrakk> P \<turnstile> xs \<Rightarrow>(n) u @ [Nt A] @ v; Q n (u @ [Nt A] @ v); (A,w) \<in> P \<rbrakk> \<Longrightarrow> Q (Suc n) (u @ w @ v)"
+  shows "Q n ys"
+using assms(1) proof (induction n arbitrary: ys)
+  case 0
+  thus ?case using assms(2) by auto
+next
+  case (Suc n)
+  from relpowp_Suc_E[OF Suc.prems]
+  obtain xs' where n: "P \<turnstile> xs \<Rightarrow>(n) xs'" and 1: "P \<turnstile> xs' \<Rightarrow> ys" by auto
+  from derive.cases[OF 1] obtain u A v w where "xs' = u @ [Nt A] @ v" "(A,w) \<in> P" "ys = u @ w @ v"
+    by metis
+  with Suc.IH[OF n] assms(3) n
+  show ?case by blast
+qed
+
+lemma derives_induct[consumes 1, case_names base step]:
+  assumes "P \<turnstile> xs \<Rightarrow>* ys"
+  and "Q xs"
+  and "\<And>u A v w. \<lbrakk> P \<turnstile> xs \<Rightarrow>* u @ [Nt A] @ v; Q (u @ [Nt A] @ v); (A,w) \<in> P \<rbrakk> \<Longrightarrow> Q (u @ w @ v)"
+  shows "Q ys"
+using assms
+proof (induction rule: rtranclp_induct)
+  case base
+  from this(1) show ?case .
+next
+  case step
+  from derive.cases[OF step(2)] step(1,3-) show ?case by metis
+qed
+
+lemma converse_derives_induct[consumes 1, case_names base step]:
+  assumes "P \<turnstile> xs \<Rightarrow>* ys"
+  and Base: "Q ys"
+  and Step: "\<And>u A v w. \<lbrakk> P \<turnstile> u @ [Nt A] @ v \<Rightarrow>* ys; Q (u @ w @ v); (A,w) \<in> P \<rbrakk> \<Longrightarrow> Q (u @ [Nt A] @ v)"
+  shows "Q xs"
+  using assms(1)
+  apply (induction rule: converse_rtranclp_induct)
+  by (auto elim!: derive.cases intro!: Base Step intro: derives_rule)
+
+
+subsubsection "(De)composing derivations"
+
+lemma derive_append:
+  "\<G> \<turnstile> u \<Rightarrow> v \<Longrightarrow> \<G> \<turnstile> u@w \<Rightarrow> v@w"
+apply(erule derive.cases)
+using derive.intros by fastforce
+
+lemma derive_prepend:
+  "\<G> \<turnstile> u \<Rightarrow> v \<Longrightarrow> \<G> \<turnstile> w@u \<Rightarrow> w@v"
+apply(erule derive.cases)
+by (metis append.assoc derive.intros)
+
+lemma deriven_append:
+  "P \<turnstile> u \<Rightarrow>(n) v \<Longrightarrow> P \<turnstile> u @ w \<Rightarrow>(n) v @ w"
+  apply (induction n arbitrary: v)
+   apply simp
+  using derive_append by (fastforce simp: relpowp_Suc_right)
+
+lemma deriven_prepend:
+  "P \<turnstile> u \<Rightarrow>(n) v \<Longrightarrow> P \<turnstile> w @ u \<Rightarrow>(n) w @ v"
+  apply (induction n arbitrary: v)
+   apply simp
+  using derive_prepend by (fastforce simp: relpowp_Suc_right)
+
+lemma derives_append:
+  "P \<turnstile> u \<Rightarrow>* v \<Longrightarrow> P \<turnstile> u@w \<Rightarrow>* v@w"
+  by (metis deriven_append rtranclp_power)
+
+lemma derives_prepend:
+  "P \<turnstile> u \<Rightarrow>* v \<Longrightarrow> P \<turnstile> w@u \<Rightarrow>* w@v"
+  by (metis deriven_prepend rtranclp_power)
+
+lemma derive_append_decomp:
+  "P \<turnstile> u@v \<Rightarrow> w \<longleftrightarrow>
+  (\<exists>u'. w = u'@v \<and> P \<turnstile> u \<Rightarrow> u') \<or> (\<exists>v'. w = u@v' \<and> P \<turnstile> v \<Rightarrow> v')"
+(is "?l \<longleftrightarrow> ?r")
+proof
+  assume ?l
+  then obtain A r u1 u2
+    where Ar: "(A,r) \<in> P"
+      and uv: "u@v = u1 @ Nt A # u2"
+      and w: "w = u1 @ r @ u2"
+    by (auto simp: derive_iff)
+  from uv have "(\<exists>s. u2 = s @ v \<and> u = u1 @ Nt A # s) \<or>
+  (\<exists>s. u1 = u@s \<and> v = s @ Nt A # u2)"
+    by (auto simp: append_eq_append_conv2 append_eq_Cons_conv)
+  with Ar w show "?r" by (fastforce simp: derive_iff)
+next
+  show "?r \<Longrightarrow> ?l"
+    by (auto simp add: derive_append derive_prepend)
+qed
+
+lemma deriven_append_decomp:
+  "P \<turnstile> u @ v \<Rightarrow>(n) w \<longleftrightarrow>
+  (\<exists>n1 n2 w1 w2. n = n1 + n2 \<and> w = w1 @ w2 \<and> P \<turnstile> u \<Rightarrow>(n1) w1 \<and> P \<turnstile> v \<Rightarrow>(n2) w2)"
+  (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?l \<Longrightarrow> ?r"
+  proof (induction n arbitrary: u v)
+    case 0
+    then show ?case by simp
+  next
+    case (Suc n)
+    from Suc.prems
+    obtain u' v'
+      where or: "P \<turnstile> u \<Rightarrow> u' \<and> v' = v \<or> u' = u \<and> P \<turnstile> v \<Rightarrow> v'"
+        and n: "P \<turnstile> u'@v' \<Rightarrow>(n) w"
+      by (fastforce simp: relpowp_Suc_left derive_append_decomp)
+    from Suc.IH[OF n] or
+    show ?case
+      apply (elim disjE)
+       apply (metis add_Suc relpowp_Suc_I2)
+      by (metis add_Suc_right relpowp_Suc_I2)
+  qed
+next
+  assume ?r
+  then obtain n1 n2 w1 w2
+    where [simp]: "n = n1 + n2" "w = w1 @ w2"
+      and u: "P \<turnstile> u \<Rightarrow>(n1) w1" and v: "P \<turnstile> v \<Rightarrow>(n2) w2"
+    by auto
+  from u deriven_append
+  have "P \<turnstile> u @ v \<Rightarrow>(n1) w1 @ v" by fastforce
+  also from v deriven_prepend
+  have "P \<turnstile> w1 @ v \<Rightarrow>(n2) w1 @ w2" by fastforce
+  finally show ?l by auto
+qed
+
+lemma derives_append_decomp:
+  "P \<turnstile> u @ v \<Rightarrow>* w \<longleftrightarrow> (\<exists>u' v'. P \<turnstile> u \<Rightarrow>* u' \<and> P \<turnstile> v \<Rightarrow>* v' \<and> w = u' @ v')"
+  by (auto simp: rtranclp_power deriven_append_decomp)
+
+lemma derives_concat:
+  "\<forall>i \<in> set is. P \<turnstile> f i \<Rightarrow>* g i \<Longrightarrow> P \<turnstile> concat(map f is) \<Rightarrow>* concat(map g is)"
+proof(induction "is")
+  case Nil
+  then show ?case by auto
+next
+  case Cons
+  thus ?case by(auto simp: derives_append_decomp less_Suc_eq)
+qed
+
+lemma derives_concat1:
+  "\<forall>i \<in> set is. P \<turnstile> [f i] \<Rightarrow>* g i \<Longrightarrow> P \<turnstile> map f is \<Rightarrow>* concat(map g is)"
+using derives_concat[where f = "\<lambda>i. [f i]"] by auto
+
+lemma derive_Cons:
+"P \<turnstile> u \<Rightarrow> v \<Longrightarrow> P \<turnstile> a#u \<Rightarrow> a#v"
+  using derive_prepend[of P u v "[a]"] by auto
+
+lemma derives_Cons:
+"R \<turnstile> u \<Rightarrow>* v \<Longrightarrow> R \<turnstile> a#u \<Rightarrow>* a#v"
+  using derives_prepend[of _ _ _ "[a]"] by simp
+
+lemma derive_from_empty[simp]:
+  "P \<turnstile> [] \<Rightarrow> w \<longleftrightarrow> False"
+  by (auto simp: derive_iff)
+
+lemma deriven_from_empty[simp]:
+  "P \<turnstile> [] \<Rightarrow>(n) w \<longleftrightarrow> n = 0 \<and> w = []"
+  by (induct n, auto simp: relpowp_Suc_left)
+
+lemma derives_from_empty[simp]:
+  "\<G> \<turnstile> [] \<Rightarrow>* w \<longleftrightarrow> w = []"
+  by (auto elim: converse_rtranclpE)
+
+lemma deriven_start1:
+  assumes "P \<turnstile> [Nt A] \<Rightarrow>(n) map Tm w"
+  shows "\<exists>\<alpha> m. n = Suc m \<and> P \<turnstile> \<alpha> \<Rightarrow>(m) (map Tm w) \<and> (A,\<alpha>) \<in> P"
+proof (cases n)
+  case 0
+  thus ?thesis
+    using assms by auto
+next
+  case (Suc m)
+  then obtain \<alpha> where *: "P \<turnstile> [Nt A] \<Rightarrow> \<alpha>" "P \<turnstile> \<alpha> \<Rightarrow>(m) map Tm w"
+    using assms by (meson relpowp_Suc_E2)
+  from  derive.cases[OF *(1)] have "(A, \<alpha>) \<in> P"
+    by (simp add: Cons_eq_append_conv) blast
+  thus ?thesis using *(2) Suc by auto
+qed
+
+lemma derives_start1: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w  \<Longrightarrow>  \<exists>\<alpha>. P \<turnstile> \<alpha> \<Rightarrow>* map Tm w \<and> (A,\<alpha>) \<in> P"
+using deriven_start1 by (metis rtranclp_power)
+
+lemma Lang_empty_if_notin_Lhss: "A \<notin> Lhss P \<Longrightarrow> Lang P A = {}"
+unfolding Lhss_def Lang_def
+using derives_start1 by fastforce
+
+lemma derive_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow> v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow> w)"
+  by (fastforce simp: derive_iff Cons_eq_append_conv)
+
+lemma deriven_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow>(n) v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow>(n) w)"
+proof (induction n arbitrary: u)
+  case 0
+  show ?case by auto
+next
+  case (Suc n)
+  then show ?case by (force simp: derive_Tm_Cons relpowp_Suc_left OO_def)
+qed
+
+lemma derives_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow>* v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow>* w)"
+  by (metis deriven_Tm_Cons rtranclp_power)
+
+lemma derives_Tm[simp]: "P \<turnstile> [Tm a] \<Rightarrow>* w \<longleftrightarrow> w = [Tm a]"
+by(simp add: derives_Tm_Cons)
+
+lemma derive_singleton: "P \<turnstile> [a] \<Rightarrow> u \<longleftrightarrow> (\<exists>A. (A,u) \<in> P \<and> a = Nt A)"
+  by (auto simp: derive_iff Cons_eq_append_conv)
+
+lemma deriven_singleton: "P \<turnstile> [a] \<Rightarrow>(n) u \<longleftrightarrow> (
+  case n of 0 \<Rightarrow> u = [a]
+   | Suc m \<Rightarrow> \<exists>(A,w) \<in> P. a = Nt A \<and> P \<turnstile> w \<Rightarrow>(m) u)"
+  (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?l \<Longrightarrow> ?r"
+  proof (induction n)
+    case 0
+    then show ?case by simp
+  next
+    case (Suc n)
+    then show ?case
+      by (smt (verit, ccfv_threshold) case_prod_conv derive_singleton nat.simps(5) relpowp_Suc_E2)
+  qed
+  show "?r \<Longrightarrow> ?l"
+    by (auto simp: derive_singleton relpowp_Suc_I2 split: nat.splits)
+qed
+
+lemma deriven_Cons_decomp:
+  "P \<turnstile> a # u \<Rightarrow>(n) v \<longleftrightarrow>
+  (\<exists>v2. v = a#v2 \<and> P \<turnstile> u \<Rightarrow>(n) v2) \<or>
+  (\<exists>n1 n2 A w v1 v2. n = Suc (n1 + n2) \<and> v = v1 @ v2 \<and> a = Nt A \<and>
+   (A,w) \<in> P \<and> P \<turnstile> w \<Rightarrow>(n1) v1 \<and> P \<turnstile> u \<Rightarrow>(n2) v2)"
+(is "?l = ?r")
+proof
+  assume ?l
+  then obtain n1 n2 v1 v2
+    where [simp]: "n = n1 + n2" "v = v1 @ v2"
+      and 1: "P \<turnstile> [a] \<Rightarrow>(n1) v1" and 2: "P \<turnstile> u \<Rightarrow>(n2) v2"
+    unfolding deriven_append_decomp[of n P "[a]" u v,simplified]
+    by auto
+  show ?r
+  proof (cases "n1")
+    case 0
+    with 1 2 show ?thesis by auto
+  next
+    case [simp]: (Suc m)
+    with 1 obtain A w
+      where [simp]: "a = Nt A" "(A,w) \<in> P" and w: "P \<turnstile> w \<Rightarrow>(m) v1"
+      by (auto simp: deriven_singleton)
+    with 2
+    have "n = Suc (m + n2) \<and> v = v1 @ v2 \<and> a = Nt A \<and>
+   (A,w) \<in> P \<and> P \<turnstile> w \<Rightarrow>(m) v1 \<and> P \<turnstile> u \<Rightarrow>(n2) v2"
+      by auto
+    then show ?thesis
+      by (auto simp: append_eq_Cons_conv)
+  qed
+next
+  assume "?r"
+  then
+  show "?l"
+  proof (elim disjE exE conjE)
+    fix v2
+    assume [simp]: "v = a # v2" and u: "P \<turnstile> u \<Rightarrow>(n) v2"
+    from deriven_prepend[OF u, of "[a]"]
+    show ?thesis
+      by auto
+  next
+    fix n1 n2 A w v1 v2
+    assume [simp]: "n = Suc (n1 + n2)" "v = v1 @ v2" "a = Nt A"
+      and Aw: "(A, w) \<in> P"
+      and w: "P \<turnstile> w \<Rightarrow>(n1) v1"
+      and u: "P \<turnstile> u \<Rightarrow>(n2) v2"
+    have "P \<turnstile> [a] \<Rightarrow> w"
+      by (simp add: Aw derive_singleton)
+    with w have "P \<turnstile> [a] \<Rightarrow>(Suc n1) v1"
+      by (metis relpowp_Suc_I2)
+    from deriven_append[OF this]
+    have 1: "P \<turnstile> a#u \<Rightarrow>(Suc n1) v1@u"
+      by auto
+    also have "P \<turnstile> \<dots> \<Rightarrow>(n2) v1@v2"
+      using deriven_prepend[OF u].
+    finally
+    show ?thesis by simp
+  qed
+qed
+
+lemma derives_Cons_decomp:
+  "P \<turnstile> s # u \<Rightarrow>* v \<longleftrightarrow>
+  (\<exists>v2. v = s#v2 \<and> P \<turnstile> u \<Rightarrow>* v2) \<or>
+  (\<exists>A w v1 v2. v = v1 @ v2 \<and> s = Nt A \<and> (A,w) \<in> P \<and> P \<turnstile> w \<Rightarrow>* v1 \<and> P \<turnstile> u \<Rightarrow>* v2)" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume ?L thus ?R using deriven_Cons_decomp[of _ P s u v] by (metis rtranclp_power)
+next
+  assume ?R thus ?L by (meson derives_Cons derives_Cons_rule derives_append_decomp)
+qed
+
+lemma deriven_Suc_decomp_left:
+  "P \<turnstile> u \<Rightarrow>(Suc n) v \<longleftrightarrow> (\<exists>p A u2 w v1 v2 n1 n2.
+  u = p @ Nt A # u2 \<and> v = p @ v1 @ v2 \<and> n = n1 + n2 \<and>
+  (A,w) \<in> P \<and> P \<turnstile> w \<Rightarrow>(n1) v1 \<and>
+  P \<turnstile> u2 \<Rightarrow>(n2) v2)" (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?r \<Longrightarrow> ?l"
+    by (auto intro!: deriven_prepend simp: deriven_Cons_decomp)
+  show "?l \<Longrightarrow> ?r"
+  proof(induction u arbitrary: v n)
+    case Nil
+    then show ?case by auto
+  next
+    case (Cons a u)
+    from Cons.prems[unfolded deriven_Cons_decomp]
+    show ?case
+    proof (elim disjE exE conjE, goal_cases)
+      case (1 v2)
+      with Cons.IH[OF this(2)] show ?thesis
+        by (metis append_Cons)
+    next
+      case (2 n1 n2 A w v1 v2)
+      then show ?thesis by (fastforce simp:Cons_eq_append_conv)
+    qed
+  qed
+qed
+
+lemma derives_NilD: "P \<turnstile> w \<Rightarrow>* [] \<Longrightarrow> s \<in> set w \<Longrightarrow> P \<turnstile> [s] \<Rightarrow>* []"
+proof(induction arbitrary: s rule: converse_derives_induct)
+  case base
+  then show ?case by simp
+next
+  case (step u A v w)
+  then show ?case using derives_append_decomp[where u="[Nt A]" and v=v]
+    by (auto simp: derives_append_decomp)
+qed
+
+lemma derive_decomp_Tm: "P \<turnstile> \<alpha> \<Rightarrow>(n) map Tm \<beta> \<Longrightarrow>
+  \<exists>\<beta>s ns. \<beta> = concat \<beta>s \<and> length \<alpha> = length \<beta>s \<and> length \<alpha> = length ns \<and> sum_list ns = n
+          \<and> (\<forall>i < length \<beta>s. P \<turnstile> [\<alpha> ! i] \<Rightarrow>(ns!i) map Tm (\<beta>s ! i))"
+  (is "_ \<Longrightarrow> \<exists>\<beta>s ns. ?G \<alpha> \<beta> n \<beta>s ns")
+proof (induction \<alpha> arbitrary: \<beta> n)
+  case (Cons s \<alpha>)
+  from deriven_Cons_decomp[THEN iffD1, OF Cons.prems]
+  show ?case
+  proof (elim disjE exE conjE)
+    fix \<gamma> assume as: "map Tm \<beta> = s # \<gamma>" "P \<turnstile> \<alpha> \<Rightarrow>(n) \<gamma>"
+    then obtain s' \<gamma>' where "\<beta> = s' # \<gamma>'"  "P \<turnstile> \<alpha> \<Rightarrow>(n) map Tm \<gamma>'" "s = Tm s'" by force
+    from Cons.IH[OF this(2)] obtain \<beta>s ns where *: "?G \<alpha> \<gamma>' n \<beta>s ns"
+      by blast
+    let ?\<beta>s = "[s']#\<beta>s"
+    let ?ns = "0#ns"
+    have "?G (s#\<alpha>) \<beta> n ?\<beta>s ?ns"
+      using \<open>\<beta> = _\<close> as * by (auto simp: nth_Cons')
+    then show ?thesis by blast
+  next
+    fix n1 n2 A w \<beta>1 \<beta>2
+    assume *: "n = Suc (n1 + n2)" "map Tm \<beta> = \<beta>1 @ \<beta>2" "s = Nt A" "(A, w) \<in> P" "P \<turnstile> w \<Rightarrow>(n1) \<beta>1" "P \<turnstile> \<alpha> \<Rightarrow>(n2) \<beta>2"
+    then obtain \<beta>1' \<beta>2' where **: "\<beta> = \<beta>1' @ \<beta>2'" "P \<turnstile> w \<Rightarrow>(n1) map Tm \<beta>1'" "P \<turnstile> \<alpha> \<Rightarrow>(n2) map Tm \<beta>2'"
+      by (metis (no_types, lifting) append_eq_map_conv)
+    from Cons.IH[OF this(3)] obtain \<beta>s ns
+      where ***: "?G \<alpha> \<beta>2' n2 \<beta>s ns"
+      by blast
+    let ?\<beta>s = "\<beta>1'#\<beta>s"
+    let ?ns = "Suc n1 # ns"
+    from * ** have "P \<turnstile> [(s#\<alpha>) ! 0] \<Rightarrow>(?ns ! 0) map Tm (?\<beta>s ! 0)"
+      by (metis derive_singleton nth_Cons_0 relpowp_Suc_I2)
+    then have "?G (s#\<alpha>) \<beta> n ?\<beta>s ?ns"
+      using * ** *** by (auto simp add: nth_Cons' derives_Cons_rule fold_plus_sum_list_rev)
+    then show ?thesis by blast
+  qed
+qed simp
+
+lemma derives_simul_rules:
+  assumes "\<And>A w. (A,w) \<in> P \<Longrightarrow> P' \<turnstile> [Nt A] \<Rightarrow>* w"
+  shows "P \<turnstile> w \<Rightarrow>* w' \<Longrightarrow> P' \<turnstile> w \<Rightarrow>* w'"
+proof(induction rule: derives_induct)
+  case base
+  then show ?case by simp
+next
+  case (step u A v w)
+  then show ?case
+    by (meson assms derives_append derives_prepend rtranclp_trans)
+qed
+
+lemma derives_set_subset:
+  "P \<turnstile> u \<Rightarrow>* v \<Longrightarrow> set v \<subseteq> set u \<union> Syms P"
+proof (induction rule: derives_induct)
+  case base
+  then show ?case by simp
+next
+  case (step u A v w)
+  then show ?case unfolding Syms_def by (auto)
+qed
+
+lemma derives_nts_syms_subset:
+  "P \<turnstile> u \<Rightarrow>* v \<Longrightarrow> nts_syms v \<subseteq> nts_syms u \<union> Nts P"
+proof (induction rule: derives_induct)
+  case base
+  then show ?case by simp
+next
+  case (step u A v w)
+  then show ?case unfolding Nts_def nts_syms_def by (auto)
+qed
+
+
+text \<open>Bottom-up definition of \<open>\<Rightarrow>*\<close>. Single definition yields more compact inductions.
+But \<open>derives_induct\<close> may already do the job.\<close>
+
+inductive derives_bu :: "('n, 't) Prods \<Rightarrow> ('n,'t) syms \<Rightarrow> ('n,'t) syms \<Rightarrow> bool"
+  ("(2_ \<turnstile>/ (_/ \<Rightarrow>bu/ _))" [50, 0, 50] 50) for P :: "('n, 't) Prods"
+  where
+bu_refl: "P \<turnstile> \<alpha> \<Rightarrow>bu \<alpha>" |
+bu_prod: "(A,\<alpha>) \<in> P \<Longrightarrow> P \<turnstile> [Nt A] \<Rightarrow>bu \<alpha>" |
+bu_embed: "\<lbrakk> P \<turnstile> \<alpha> \<Rightarrow>bu \<alpha>1 @ \<alpha>2 @ \<alpha>3; P \<turnstile> \<alpha>2 \<Rightarrow>bu \<beta> \<rbrakk> \<Longrightarrow> P \<turnstile> \<alpha> \<Rightarrow>bu \<alpha>1 @ \<beta> @ \<alpha>3"
+
+lemma derives_if_bu: "P \<turnstile> \<alpha> \<Rightarrow>bu \<beta> \<Longrightarrow> P \<turnstile> \<alpha> \<Rightarrow>* \<beta>"
+proof(induction rule: derives_bu.induct)
+  case (bu_refl) then show ?case by simp
+next
+  case (bu_prod A \<alpha>) then show ?case by (simp add: derives_Cons_rule)
+next
+  case (bu_embed \<alpha> \<alpha>1 \<alpha>2 \<alpha>3 \<beta>) then show ?case
+    by (meson derives_append derives_prepend rtranclp_trans)
+qed
+
+lemma derives_bu_if: "P \<turnstile> \<alpha> \<Rightarrow>* \<beta> \<Longrightarrow> P \<turnstile> \<alpha> \<Rightarrow>bu \<beta>"
+proof(induction rule: derives_induct)
+  case base
+  then show ?case by (simp add: bu_refl)
+next
+  case (step u A v w)
+  then show ?case
+    by (meson bu_embed bu_prod)
+qed
+
+lemma derives_bu_iff: "P \<turnstile> \<alpha> \<Rightarrow>bu \<beta> \<longleftrightarrow> P \<turnstile> \<alpha> \<Rightarrow>* \<beta>"
+by (meson derives_bu_if derives_if_bu)
+
+
+subsubsection "Leftmost/Rightmost Derivations"
+
+inductive derivel :: "('n,'t) Prods \<Rightarrow> ('n,'t) syms \<Rightarrow> ('n,'t)syms \<Rightarrow> bool"
+  ("(2_ \<turnstile>/ (_ \<Rightarrow>l/ _))" [50, 0, 50] 50) where
+"(A,\<alpha>) \<in> P \<Longrightarrow> P \<turnstile> map Tm u @ Nt A # v \<Rightarrow>l map Tm u @ \<alpha> @ v"
+
+abbreviation derivels ("(2_ \<turnstile>/ (_ \<Rightarrow>l*/ _))" [50, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>l* v \<equiv> ((derivel P) ^**) u v"
+
+abbreviation derivels1 ("(2_ \<turnstile>/ (_ \<Rightarrow>l+/ _))" [50, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>l+ v \<equiv> ((derivel P) ^++) u v"
+
+abbreviation deriveln ("(2_ \<turnstile>/ (_ \<Rightarrow>l'(_')/ _))" [50, 0, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>l(n) v \<equiv> ((derivel P) ^^ n) u v"
+
+inductive deriver :: "('n,'t) Prods \<Rightarrow> ('n,'t) syms \<Rightarrow> ('n,'t)syms \<Rightarrow> bool"
+  ("(2_ \<turnstile>/ (_ \<Rightarrow>r/ _))" [50, 0, 50] 50) where
+"(A,\<alpha>) \<in> P \<Longrightarrow> P \<turnstile> u @ Nt A # map Tm v \<Rightarrow>r u @ \<alpha> @ map Tm v"
+
+abbreviation derivers ("(2_ \<turnstile>/ (_ \<Rightarrow>r*/ _))" [50, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>r* v \<equiv> ((deriver P) ^**) u v"
+
+abbreviation derivers1 ("(2_ \<turnstile>/ (_ \<Rightarrow>r+/ _))" [50, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>r+ v \<equiv> ((deriver P) ^++) u v"
+
+abbreviation derivern ("(2_ \<turnstile>/ (_ \<Rightarrow>r'(_')/ _))" [50, 0, 0, 50] 50) where
+"P \<turnstile> u \<Rightarrow>r(n) v \<equiv> ((deriver P) ^^ n) u v"
+
+
+lemma derivel_iff: "R \<turnstile> u \<Rightarrow>l v \<longleftrightarrow>
+ (\<exists> (A,w) \<in> R. \<exists>u1 u2. u = map Tm u1 @ Nt A # u2 \<and> v = map Tm u1 @ w @ u2)"
+  by (auto simp: derivel.simps)
+
+lemma derivel_from_empty[simp]:
+  "P \<turnstile> [] \<Rightarrow>l w \<longleftrightarrow> False" by (auto simp: derivel_iff)
+
+lemma deriveln_from_empty[simp]:
+  "P \<turnstile> [] \<Rightarrow>l(n) w \<longleftrightarrow> n = 0 \<and> w = []"
+  by (induct n, auto simp: relpowp_Suc_left)
+
+lemma derivels_from_empty[simp]:
+  "\<G> \<turnstile> [] \<Rightarrow>l* w \<longleftrightarrow> w = []"
+  by (auto elim: converse_rtranclpE)
+
+lemma Un_derivel: "R \<union> S \<turnstile> y \<Rightarrow>l z \<longleftrightarrow> R \<turnstile> y \<Rightarrow>l z \<or> S \<turnstile> y \<Rightarrow>l z"
+  by (fastforce simp: derivel_iff)
+
+lemma derivel_append:
+  "P \<turnstile> u \<Rightarrow>l v \<Longrightarrow> P \<turnstile> u @ w \<Rightarrow>l v @ w"
+  by (force simp: derivel_iff)
+
+lemma deriveln_append:
+  "P \<turnstile> u \<Rightarrow>l(n) v \<Longrightarrow> P \<turnstile> u @ w \<Rightarrow>l(n) v @ w"
+proof (induction n arbitrary: u)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  then obtain y where uy: "P \<turnstile> u \<Rightarrow>l y" and yv: "P \<turnstile> y \<Rightarrow>l(n) v"
+    by (auto simp: relpowp_Suc_left)
+  have "P \<turnstile> u @ w \<Rightarrow>l y @ w"
+    using derivel_append[OF uy].
+  also from Suc.IH yv have "P \<turnstile> \<dots> \<Rightarrow>l(n) v @ w" by auto
+  finally show ?case.
+qed
+
+lemma derivels_append:
+  "P \<turnstile> u \<Rightarrow>l* v \<Longrightarrow> P \<turnstile> u @ w \<Rightarrow>l* v @ w"
+  by (metis deriveln_append rtranclp_power)
+
+lemma derivels1_append:
+  "P \<turnstile> u \<Rightarrow>l+ v \<Longrightarrow> P \<turnstile> u @ w \<Rightarrow>l+ v @ w"
+  by (metis deriveln_append tranclp_power)
+
+lemma derivel_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow>l v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow>l w)"
+apply (cases v)
+ apply (simp add: derivel_iff)
+apply (fastforce simp: derivel.simps Cons_eq_append_conv Cons_eq_map_conv)
+done
+
+lemma deriveln_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow>l(n) v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow>l(n) w)"
+  by (induction n arbitrary: u v;
+      fastforce simp: derivel_Tm_Cons relpowp_Suc_right OO_def)
+
+lemma derivels_Tm_Cons:
+  "P \<turnstile> Tm a # u \<Rightarrow>l* v \<longleftrightarrow> (\<exists>w. v = Tm a # w \<and> P \<turnstile> u \<Rightarrow>l* w)"
+  by (metis deriveln_Tm_Cons rtranclp_power)
+
+lemma derivel_Nt_Cons:
+  "P \<turnstile> Nt A # u \<Rightarrow>l v \<longleftrightarrow> (\<exists>w. (A,w) \<in> P \<and> v = w @ u)"
+  by (auto simp: derivel_iff Cons_eq_append_conv Cons_eq_map_conv)
+
+lemma derivels1_Nt_Cons:
+  "P \<turnstile> Nt A # u \<Rightarrow>l+ v \<longleftrightarrow> (\<exists>w. (A,w) \<in> P \<and> P \<turnstile> w @ u \<Rightarrow>l* v)"
+  by (auto simp: tranclp_unfold_left derivel_Nt_Cons OO_def)
+
+lemma derivels_Nt_Cons:
+  "P \<turnstile> Nt A # u \<Rightarrow>l* v \<longleftrightarrow> v = Nt A # u \<or> (\<exists>w. (A,w) \<in> P \<and> P \<turnstile> w @ u \<Rightarrow>l* v)"
+  by (auto simp: Nitpick.rtranclp_unfold derivels1_Nt_Cons)
+
+lemma deriveln_Nt_Cons:
+  "P \<turnstile> Nt A # u \<Rightarrow>l(n) v \<longleftrightarrow> (
+  case n of 0 \<Rightarrow> v = Nt A # u
+  | Suc m \<Rightarrow> \<exists>w. (A,w) \<in> P \<and> P \<turnstile> w @ u \<Rightarrow>l(m) v)"
+  by (cases n) (auto simp: derivel_Nt_Cons relpowp_Suc_left OO_def)
+
+lemma derivel_map_Tm_append:
+  "P \<turnstile> map Tm w @ u \<Rightarrow>l v \<longleftrightarrow> (\<exists>x. v = map Tm w @ x \<and> P \<turnstile> u \<Rightarrow>l x)"
+  apply (induction w arbitrary:v)
+  by (auto simp: derivel_Tm_Cons Cons_eq_append_conv)
+
+lemma deriveln_map_Tm_append:
+  "P \<turnstile> map Tm w @ u \<Rightarrow>l(n) v \<longleftrightarrow> (\<exists>x. v = map Tm w @ x \<and> P \<turnstile> u \<Rightarrow>l(n) x)"
+  by (induction n arbitrary: u;
+      force simp: derivel_map_Tm_append relpowp_Suc_left OO_def)
+
+lemma derivels_map_Tm_append:
+  "P \<turnstile> map Tm w @ u \<Rightarrow>l* v \<longleftrightarrow> (\<exists>x. v = map Tm w @ x \<and> P \<turnstile> u \<Rightarrow>l* x)"
+  by (metis deriveln_map_Tm_append rtranclp_power)
+
+lemma derivel_not_elim_Tm:
+  assumes "P \<turnstile> xs \<Rightarrow>l map Nt w"
+  shows "\<exists>v. xs = map Nt v"
+proof -
+  from assms obtain A \<alpha> u xs' where
+         A_w: "(A, \<alpha>)\<in>P"
+      and xs: "xs = map Tm u @ Nt A # xs'"
+      and ys: "map Nt w = map Tm u @ \<alpha> @ xs'"
+    unfolding derivel_iff by fast
+
+  from ys have u1: "u = []"
+    by (metis Nil_is_append_conv Nil_is_map_conv hd_append list.map_sel(1) sym.simps(4))
+  moreover from ys obtain u' where "xs' = map Nt u'"
+    by (metis append_eq_map_conv)
+
+  ultimately have "xs = map Nt (A # u')"
+    by (simp add: xs)
+  then show ?thesis by blast
+qed
+
+lemma deriveln_not_elim_Tm:
+  assumes "P \<turnstile> xs \<Rightarrow>l(n) map Nt w"
+  shows "\<exists>v. xs = map Nt v"
+using assms
+proof (induction n arbitrary: xs)
+  case 0
+  then show ?case by auto
+next
+  case (Suc n)
+  then obtain z where "P \<turnstile> xs \<Rightarrow>l z" and "P \<turnstile> z \<Rightarrow>l(n) map Nt w"
+    using relpowp_Suc_E2 by metis
+  with Suc show ?case using derivel_not_elim_Tm
+    by blast
+qed
+
+lemma decomp_derivel_map_Nts:
+  assumes "P \<turnstile> map Nt Xs \<Rightarrow>l map Nt Zs"
+  shows "\<exists>X Xs' Ys. Xs = X # Xs' \<and> P \<turnstile> [Nt X] \<Rightarrow>l map Nt Ys \<and> Zs = Ys @ Xs'"
+using assms unfolding derivel_iff
+by (fastforce simp: map_eq_append_conv split: prod.splits)
+
+lemma derivel_imp_derive: "P \<turnstile> u \<Rightarrow>l v \<Longrightarrow> P \<turnstile> u \<Rightarrow> v"
+  using derive.simps derivel.cases self_append_conv2 by fastforce
+
+lemma deriveln_imp_deriven:
+  "P \<turnstile> u \<Rightarrow>l(n) v \<Longrightarrow> P \<turnstile> u \<Rightarrow>(n) v"
+  using relpowp_mono derivel_imp_derive by metis
+
+lemma derivels_imp_derives:
+  "P \<turnstile> u \<Rightarrow>l* v \<Longrightarrow> P \<turnstile> u \<Rightarrow>* v"
+  by (metis derivel_imp_derive mono_rtranclp)
+
+lemma deriveln_iff_deriven:
+  "P \<turnstile> u \<Rightarrow>l(n) map Tm v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>(n) map Tm v"
+  (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?l \<Longrightarrow> ?r" using deriveln_imp_deriven.
+  assume ?r
+  then show "?l"
+  proof (induction n arbitrary: u v rule: less_induct)
+    case (less n)
+    from \<open>P \<turnstile> u \<Rightarrow>(n) map Tm v\<close>
+    show ?case
+    proof (induction u arbitrary: v)
+      case Nil
+      then show ?case by simp
+    next
+      case (Cons a u)
+      show ?case
+        using Cons.prems(1) Cons.IH less.IH
+        by (auto simp: deriven_Cons_decomp deriveln_Tm_Cons deriveln_Nt_Cons)
+           (metis deriven_append_decomp lessI)
+    qed
+  qed
+qed
+
+lemma derivels_iff_derives: "P \<turnstile> u \<Rightarrow>l* map Tm v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>* map Tm v"
+  using deriveln_iff_deriven
+  by (metis rtranclp_power)
+
+lemma deriver_iff: "R \<turnstile> u \<Rightarrow>r v \<longleftrightarrow>
+  (\<exists> (A,w) \<in> R. \<exists>u1 u2. u = u1 @ Nt A # map Tm u2 \<and> v = u1 @ w @ map Tm u2)"
+  by (auto simp: deriver.simps)
+
+lemma deriver_imp_derive: "R \<turnstile> u \<Rightarrow>r v \<Longrightarrow> R \<turnstile> u \<Rightarrow> v"
+  by (auto simp: deriver_iff derive_iff)
+
+lemma derivern_imp_deriven: "R \<turnstile> u \<Rightarrow>r(n) v \<Longrightarrow> R \<turnstile> u \<Rightarrow>(n) v"
+  by (metis (no_types, lifting) deriver_imp_derive relpowp_mono)
+
+lemma derivers_imp_derives: "R \<turnstile> u \<Rightarrow>r* v \<Longrightarrow> R \<turnstile> u \<Rightarrow>* v"
+  by (metis deriver_imp_derive mono_rtranclp)
+
+lemma deriver_iff_rev_derivel:
+  "P \<turnstile> u \<Rightarrow>r v \<longleftrightarrow> map_prod id rev ` P \<turnstile> rev u \<Rightarrow>l rev v" (is "?l \<longleftrightarrow> ?r")
+proof
+  assume ?l
+  then obtain A w u1 u2 where Aw: "(A,w) \<in> P"
+    and u: "u = u1 @ Nt A # map Tm u2"
+    and v: "v = u1 @ w @ map Tm u2" by (auto simp: deriver.simps)
+  from bexI[OF _ Aw] have "(A, rev w) \<in> map_prod id rev ` P" by (auto simp: image_def)
+  from derivel.intros[OF this, of "rev u2" "rev u1"] u v
+  show ?r by (simp add: rev_map)
+next
+  assume ?r
+  then obtain A w u1 u2 where Aw: "(A,w) \<in> P"
+    and u: "u = u1 @ Nt A # map Tm u2"
+    and v: "v = u1 @ w @ map Tm u2"
+    by (auto simp: derivel_iff rev_eq_append_conv rev_map)
+  then show ?l by (auto simp: deriver_iff)
+qed
+
+lemma rev_deriver_iff_derivel:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow>r v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow>l rev v"
+  by (simp add: deriver_iff_rev_derivel image_image prod.map_comp o_def)
+
+lemma derivern_iff_rev_deriveln:
+  "P \<turnstile> u \<Rightarrow>r(n) v \<longleftrightarrow> map_prod id rev ` P \<turnstile> rev u \<Rightarrow>l(n) rev v"
+proof (induction n arbitrary: u)
+  case 0
+  show ?case by simp
+next
+  case (Suc n)
+  show ?case
+    apply (unfold relpowp_Suc_left OO_def)
+    apply (unfold Suc deriver_iff_rev_derivel)
+    by (metis rev_rev_ident)
+qed
+
+lemma rev_derivern_iff_deriveln:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow>r(n) v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow>l(n) rev v"
+  by (simp add: derivern_iff_rev_deriveln image_image prod.map_comp o_def)
+
+lemma derivers_iff_rev_derivels:
+  "P \<turnstile> u \<Rightarrow>r* v \<longleftrightarrow> map_prod id rev ` P \<turnstile> rev u \<Rightarrow>l* rev v"
+  using derivern_iff_rev_deriveln
+  by (metis rtranclp_power)
+
+lemma rev_derivers_iff_derivels:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow>r* v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow>l* rev v"
+  by (simp add: derivers_iff_rev_derivels image_image prod.map_comp o_def)
+
+lemma rev_derive:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow> v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow> rev v"
+  by (force simp: derive_iff rev_eq_append_conv bex_pair_conv in_image_map_prod intro: exI[of _ "rev _"])
+
+lemma rev_deriven:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow>(n) v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow>(n) rev v"
+apply (induction n arbitrary: u)
+ apply simp
+by (auto simp: relpowp_Suc_left OO_def rev_derive intro: exI[of _ "rev _"])
+
+lemma rev_derives:
+  "map_prod id rev ` P \<turnstile> u \<Rightarrow>* v \<longleftrightarrow> P \<turnstile> rev u \<Rightarrow>* rev v"
+  using rev_deriven
+  by (metis rtranclp_power)
+
+lemma derivern_iff_deriven: "P \<turnstile> u \<Rightarrow>r(n) map Tm v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>(n) map Tm v"
+  by (auto simp: derivern_iff_rev_deriveln rev_map deriveln_iff_deriven rev_deriven)
+
+lemma derivers_iff_derives: "P \<turnstile> u \<Rightarrow>r* map Tm v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>* map Tm v"
+  by (simp add: derivern_iff_deriven rtranclp_power)
+
+lemma deriver_append_map_Tm:
+  "P \<turnstile> u @ map Tm w \<Rightarrow>r v \<longleftrightarrow> (\<exists>x. v = x @ map Tm w \<and> P \<turnstile> u \<Rightarrow>r x)"
+  by (fastforce simp: deriver_iff_rev_derivel rev_map derivel_map_Tm_append rev_eq_append_conv)
+
+lemma derivern_append_map_Tm:
+  "P \<turnstile> u @ map Tm w \<Rightarrow>r(n) v \<longleftrightarrow> (\<exists>x. v = x @ map Tm w \<and> P \<turnstile> u \<Rightarrow>r(n) x)"
+  by (fastforce simp: derivern_iff_rev_deriveln rev_map deriveln_map_Tm_append rev_eq_append_conv)
+
+lemma deriver_snoc_Nt:
+  "P \<turnstile> u @ [Nt A] \<Rightarrow>r v \<longleftrightarrow> (\<exists>w. (A,w) \<in> P \<and> v = u @ w)"
+  by (force simp: deriver_iff_rev_derivel derivel_Nt_Cons rev_eq_append_conv)
+
+lemma deriver_singleton:
+  "P \<turnstile> [Nt A] \<Rightarrow>r v \<longleftrightarrow> (A,v) \<in> P"
+  using deriver_snoc_Nt[of P "[]"] by auto
+
+lemma derivers1_snoc_Nt:
+  "P \<turnstile> u @ [Nt A] \<Rightarrow>r+ v \<longleftrightarrow> (\<exists>w. (A,w) \<in> P \<and> P \<turnstile> u @ w \<Rightarrow>r* v)"
+  by (auto simp: tranclp_unfold_left deriver_snoc_Nt OO_def)
+
+lemma derivers_snoc_Nt:
+  "P \<turnstile> u @ [Nt A] \<Rightarrow>r* v \<longleftrightarrow> v = u @ [Nt A] \<or> (\<exists>w. (A,w) \<in> P \<and> P \<turnstile> u @ w \<Rightarrow>r* v)"
+  by (auto simp: Nitpick.rtranclp_unfold derivers1_snoc_Nt)
+
+lemma derivern_snoc_Tm:
+  "P \<turnstile> u @ [Tm a] \<Rightarrow>r(n) v \<longleftrightarrow> (\<exists>w. v = w @ [Tm a] \<and> P \<turnstile> u \<Rightarrow>r(n) w)"
+  by (force simp: derivern_iff_rev_deriveln deriveln_Tm_Cons)
+
+lemma derivern_snoc_Nt:
+  "P \<turnstile> u @ [Nt A] \<Rightarrow>r(n) v \<longleftrightarrow> (
+  case n of 0 \<Rightarrow> v = u @ [Nt A]
+  | Suc m \<Rightarrow> \<exists>w. (A,w) \<in> P \<and> P \<turnstile> u @ w \<Rightarrow>r(m) v)"
+  by (cases n) (auto simp: relpowp_Suc_left deriver_snoc_Nt OO_def)
+
+lemma derivern_singleton:
+  "P \<turnstile> [Nt A] \<Rightarrow>r(n) v \<longleftrightarrow> (
+  case n of 0 \<Rightarrow> v = [Nt A]
+  | Suc m \<Rightarrow> \<exists>w. (A,w) \<in> P \<and> P \<turnstile> w \<Rightarrow>r(m) v)"
+  using derivern_snoc_Nt[of n P "[]" A v] by (cases n, auto)
+
+
+subsection \<open>Substitution in Lists\<close>
+
+text \<open>Function \<open>substs y ys xs\<close> replaces every occurrence of \<open>y\<close> in \<open>xs\<close> with the list \<open>ys\<close>\<close>
+
+fun substs :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+"substs y ys [] = []" |
+"substs y ys (x#xs) = (if x = y then ys @ substs y ys xs else x # substs y ys xs)"
+
+text \<open>Alternative definition, but apparently no simpler to use:
+@{prop "substs y ys xs = concat (map (\<lambda>x. if x=y then ys else [x]) xs)"}\<close>
+
+abbreviation "substsNt A \<equiv> substs (Nt A)"
+
+lemma substs_append[simp]: "substs y ys (xs @ xs') = substs y ys xs @ substs y ys xs'"
+by (induction xs) auto
+
+lemma substs_skip: "y \<notin> set xs \<Longrightarrow> substs y ys xs = xs"
+by (induction xs) auto
+
+lemma susbstsNT_map_Tm[simp]: "substsNt A \<alpha> (map Tm w) = map Tm w"
+by(rule substs_skip) auto
+
+lemma substs_len: "length (substs y [y'] xs) = length xs"
+by (induction xs) auto
+
+lemma substs_rev: "y' \<notin> set xs \<Longrightarrow> substs y' [y] (substs y [y'] xs) = xs"
+by (induction xs) auto
+
+lemma substs_der:
+  "(B,v) \<in> P \<Longrightarrow> P \<turnstile> u \<Rightarrow>* substs (Nt B) v u"
+proof (induction u)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons a u)
+  then show ?case
+    by (auto simp add: derives_Cons_rule derives_prepend derives_Cons)
+qed
+
+
+subsection \<open>Epsilon-Freeness\<close>
+
+definition Eps_free where "Eps_free R = (\<forall>(_,r) \<in> R. r \<noteq> [])"
+
+abbreviation "eps_free rs == Eps_free(set rs)"
+
+lemma Eps_freeI:
+  assumes "\<And>A r. (A,r) \<in> R \<Longrightarrow> r \<noteq> []" shows "Eps_free R"
+  using assms by (auto simp: Eps_free_def)
+
+lemma Eps_free_Nil: "Eps_free R \<Longrightarrow> (A,[]) \<notin> R"
+  by (auto simp: Eps_free_def)
+
+lemma Eps_freeE_Cons: "Eps_free R \<Longrightarrow> (A,w) \<in> R \<Longrightarrow> \<exists>a u. w = a#u"
+  by (cases w, auto simp: Eps_free_def)
+
+lemma Eps_free_derives_Nil:
+  assumes R: "Eps_free R" shows "R \<turnstile> l \<Rightarrow>* [] \<longleftrightarrow> l = []" (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?l \<Longrightarrow> ?r"
+  proof (induction rule: converse_derives_induct)
+    case base
+    show ?case by simp
+  next
+    case (step u A v w)
+    then show ?case by (auto simp: Eps_free_Nil[OF R])
+  qed
+qed auto
+
+lemma Eps_free_derivels_Nil: "Eps_free R \<Longrightarrow> R \<turnstile> l \<Rightarrow>l* [] \<longleftrightarrow> l = []"
+by (meson Eps_free_derives_Nil derivels_from_empty derivels_imp_derives)
+
+lemma Eps_free_deriveln_Nil: "Eps_free R \<Longrightarrow> R \<turnstile> l \<Rightarrow>l(n) [] \<Longrightarrow> l = []"
+by (metis Eps_free_derivels_Nil relpowp_imp_rtranclp)
+
+lemma decomp_deriveln_map_Nts:
+  assumes "Eps_free P"
+  shows "P \<turnstile> Nt X # map Nt Xs \<Rightarrow>l(n) map Nt Zs \<Longrightarrow>
+     \<exists>Ys'. Ys' @ Xs = Zs \<and> P \<turnstile> [Nt X] \<Rightarrow>l(n) map Nt Ys'"
+proof (induction n arbitrary: Zs)
+  case 0
+  then show ?case
+    by (auto)
+next
+  case (Suc n)
+  then obtain ys where n: "P \<turnstile> Nt X # map Nt Xs \<Rightarrow>l(n) ys" and "P \<turnstile> ys \<Rightarrow>l map Nt Zs"
+    using relpowp_Suc_E by metis
+  from \<open>P \<turnstile> ys \<Rightarrow>l map Nt Zs\<close> obtain Ys where "ys = map Nt Ys"
+    using derivel_not_elim_Tm by blast
+  from Suc.IH[of Ys] this n
+  obtain Ys' where "Ys = Ys' @ Xs \<and> P \<turnstile> [Nt X] \<Rightarrow>l(n) map Nt Ys'" by auto
+  moreover from \<open>ys = _\<close> \<open>P \<turnstile> ys \<Rightarrow>l map Nt Zs\<close> decomp_derivel_map_Nts[of P Ys Zs]
+  obtain Y Xs' Ysa where "Ys = Y # Xs' \<and> P \<turnstile> [Nt Y] \<Rightarrow>l map Nt Ysa \<and> Zs = Ysa @ Xs'" by auto
+  ultimately show ?case using Eps_free_deriveln_Nil[OF assms, of n "[Nt X]"]
+    by (auto simp: Cons_eq_append_conv derivel_Nt_Cons relpowp_Suc_I)
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Context_Free_Language.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Context_Free_Language.thy
@@ -0,0 +1,303 @@
+(*
+Authors: Tobias Nipkow, Fabian Lehr
+*)
+
+section "Context-Free Languages"
+
+theory Context_Free_Language
+imports
+  "Regular-Sets.Regular_Set"
+  Renaming_CFG
+  Disjoint_Union_CFG
+begin
+
+
+subsection \<open>Basic Definitions\<close>
+
+text \<open>This definition depends on the type of nonterminals of the grammar.\<close>
+
+definition CFL :: "'n itself \<Rightarrow> 't list set \<Rightarrow> bool" where
+"CFL (TYPE('n)) L = (\<exists>P S::'n. L = Lang P S \<and> finite P)"
+
+text \<open>Ideally one would existentially quantify over \<open>'n\<close> on the right-hand side, but we cannot
+quantify over types in HOL. But we can prove that the type is irrelevant because we can always
+use another type via renaming.\<close>
+
+text \<open>For hiding the infinite type of nonterminals:\<close>
+
+abbreviation cfl :: "'a lang \<Rightarrow> bool" where
+"cfl L \<equiv> CFL (TYPE(nat)) L"
+
+
+subsection \<open>Closure Properties\<close>
+
+lemma CFL_Un_closed:
+  assumes "CFL TYPE('n1) L1" "CFL TYPE('n2) L2"
+  shows "CFL TYPE(('n1+'n2)option) (L1 \<union> L2)"
+proof -
+  from assms obtain P1 P2 and S1 :: 'n1 and S2 :: 'n2
+    where L: "L1 = Lang P1 S1" "L2 = Lang P2 S2" and fin: "finite P1" "finite P2"
+    unfolding CFL_def by blast
+  let ?f1 = "Some o (Inl:: 'n1 \<Rightarrow> 'n1 + 'n2)"
+  let ?f2 = "Some o (Inr:: 'n2 \<Rightarrow> 'n1 + 'n2)"
+  let ?P1 = "rename_Prods ?f1 P1"
+  let ?P2 = "rename_Prods ?f2 P2"
+  let ?S1 = "?f1 S1"
+  let ?S2 = "?f2 S2"
+  let ?P = "{(None, [Nt ?S1]), (None, [Nt ?S2])} \<union> (?P1 \<union> ?P2)"
+  have "Lang ?P None = Lang ?P1 ?S1 \<union> Lang ?P2 ?S2"
+    by (rule Lang_disj_Un2)(auto simp: Nts_Un in_Nts_rename_Prods)
+  moreover have "\<dots> = Lang P1 S1 \<union> Lang P2 S2"
+  proof -
+    have "Lang ?P1 ?S1 = L1" unfolding \<open>L1 = _\<close>
+      by (meson Lang_rename_Prods comp_inj_on inj_Inl inj_Some)
+    moreover have "Lang ?P2 ?S2 = L2" unfolding \<open>L2 = _\<close>
+      by (meson Lang_rename_Prods comp_inj_on inj_Inr inj_Some)
+    ultimately show ?thesis using L by argo
+  qed
+  moreover have "finite ?P" using fin by auto
+  ultimately show ?thesis
+    unfolding CFL_def using L by blast 
+qed
+
+lemma CFL_concat_closed: 
+assumes "CFL TYPE('n1) L1" and "CFL TYPE('n2) L2"
+shows "CFL TYPE(('n1 + 'n2) option) (L1 @@ L2)"
+proof -
+  obtain P1 S1 where L1_def: "L1 = Lang P1 (S1::'n1)" "finite P1"
+    using assms(1) CFL_def[of L1] by auto 
+  obtain P2 S2 where L2_def: "L2 = Lang P2 (S2::'n2)" "finite P2"
+    using assms(2) CFL_def[of L2] by auto
+  let ?f1 = "Some o (Inl:: 'n1 \<Rightarrow> 'n1 + 'n2)"
+  let ?f2 = "Some o (Inr:: 'n2 \<Rightarrow> 'n1 + 'n2)"
+  let ?P1 = "rename_Prods ?f1 P1"
+  let ?P2 = "rename_Prods ?f2 P2"
+  let ?S1 = "?f1 S1"
+  let ?S2 = "?f2 S2"
+  let ?S = "None :: ('n1+'n2)option"
+  let ?P = "{(None, [Nt ?S1, Nt ?S2])} \<union> (?P1 \<union> ?P2)"
+  have "inj ?f1" by (simp add: inj_on_def) 
+  then have L1r_def: "L1 = Lang ?P1 ?S1" 
+    using L1_def Lang_rename_Prods[of ?f1 P1 S1] inj_on_def by force
+  have "inj ?f2" by (simp add: inj_on_def) 
+  then have L2r_def: "L2 = Lang ?P2 ?S2" 
+    using L2_def Lang_rename_Prods[of ?f2 P2 S2] inj_on_def by force
+  have "Lang ?P ?S = L1 @@ L2" unfolding L1r_def L2r_def  
+    by(rule Lang_concat_disj) (auto simp add: disjoint_iff in_Nts_rename_Prods)
+  moreover have "finite ?P" using \<open>finite P1\<close> \<open>finite P2\<close> by auto
+  ultimately show ?thesis unfolding CFL_def by blast
+qed
+
+
+subsection \<open>CFG as an Equation System\<close>
+
+text \<open>A CFG can be viewed as a system of equations. The least solution is denoted by \<open>Lang_lfp\<close>.\<close>
+
+definition inst_sym :: "('n \<Rightarrow> 't lang) \<Rightarrow> ('n, 't) sym \<Rightarrow> 't lang" where
+"inst_sym L s = (case s of Tm a \<Rightarrow> {[a]} | Nt A \<Rightarrow> L A)"
+
+definition concats :: "'a lang list \<Rightarrow> 'a lang" where
+"concats Ls = foldr (@@) Ls {[]}"
+
+definition inst_syms :: "('n \<Rightarrow> 't lang) \<Rightarrow> ('n, 't) syms \<Rightarrow> 't lang" where
+"inst_syms L w = concats (map (inst_sym L) w)"
+
+definition subst_lang :: "('n,'t)Prods \<Rightarrow> ('n \<Rightarrow> 't lang) \<Rightarrow> ('n \<Rightarrow> 't lang)" where
+"subst_lang P L = (\<lambda>A. \<Union>w \<in> Rhss P A. inst_syms L w)"
+
+definition Lang_lfp :: "('n, 't) Prods \<Rightarrow> 'n \<Rightarrow> 't lang" where
+"Lang_lfp P = lfp (subst_lang P)"
+
+text \<open>Now we show that this \<open>lfp\<close> is a Kleene fixpoint.\<close>
+
+lemma inst_sym_Sup_range:  "inst_sym (Sup(range F)) = (\<lambda>s. UN i. inst_sym (F i) s)"
+  by(auto simp: inst_sym_def fun_eq_iff split: sym.splits)
+
+lemma foldr_map_mono: "F \<le> G \<Longrightarrow> foldr (@@) (map F xs) Ls \<subseteq> foldr (@@) (map G xs) Ls"
+by(induction xs)(auto simp add: le_fun_def subset_eq)
+
+lemma inst_sym_mono: "F \<le> G \<Longrightarrow> inst_sym F s \<subseteq> inst_sym G s"
+by (auto simp add: inst_sym_def le_fun_def subset_iff split: sym.splits)
+
+lemma foldr_conc_map_inst_sym:
+  assumes "omega_chain L"
+  shows "foldr (@@) (map (\<lambda>s. \<Union>i. inst_sym (L i) s) xs) Ls = (\<Union>i. foldr (@@) (map (inst_sym (L i)) xs) Ls)"
+proof(induction xs)
+  case Nil
+  then show ?case by simp 
+next
+  case (Cons a xs)
+  show ?case (is "?l = ?r")
+  proof
+    show "?l \<subseteq> ?r"
+    proof
+      fix w assume "w \<in> ?l"
+      with Cons obtain u v i j
+        where "w = u @ v" "u \<in> inst_sym (L i) a" "v \<in> foldr (@@) (map (inst_sym (L j)) xs) Ls" by(auto)
+      then show "w \<in> ?r"
+        using omega_chain_mono[OF assms, of i "max i j"] omega_chain_mono[OF assms, of j "max i j"]
+          inst_sym_mono foldr_map_mono[of "inst_sym (L j)" "inst_sym (L (max i j))" xs Ls] concI
+        unfolding le_fun_def by(simp) blast
+    qed
+  next
+    show "?r \<subseteq> ?l" using Cons by(fastforce)
+  qed
+qed
+
+lemma omega_cont_Lang_lfp: "omega_cont (subst_lang P)"
+unfolding omega_cont_def subst_lang_def
+proof (safe)
+  fix L :: "nat \<Rightarrow> 'a \<Rightarrow> 'b lang"
+  assume o: "omega_chain L"
+  show "(\<lambda>A. \<Union> (inst_syms (Sup (range L)) ` Rhss P A)) = (SUP i. (\<lambda>A. \<Union> (inst_syms (L i) ` Rhss P A)))"
+    (is "(\<lambda>A. ?l A) = (\<lambda>A. ?r A)")
+  proof
+    fix A :: 'a
+    have "?l A = \<Union>(\<Union>i. (inst_syms (L i) ` Rhss P A))"
+      by(auto simp: inst_syms_def inst_sym_Sup_range concats_def foldr_conc_map_inst_sym[OF o])
+    also have "\<dots> = ?r A"
+      by(auto)
+    finally show "?l A = ?r A" .
+  qed
+qed
+
+theorem Lang_lfp_SUP: "Lang_lfp P = (SUP n. ((subst_lang P)^^n) (\<lambda>A. {}))"
+using Kleene_lfp[OF omega_cont_Lang_lfp] unfolding Lang_lfp_def bot_fun_def by blast
+
+
+subsection \<open>\<open>Lang_lfp = Lang\<close>\<close>
+
+text \<open>We prove that the fixpoint characterization of the language defined by a CFG is equivalent
+to the standard language definition via derivations. Both directions are proved separately\<close>
+
+lemma inst_syms_mono: "(\<And>A. R A \<subseteq> R' A) \<Longrightarrow> w \<in> inst_syms R \<alpha> \<Longrightarrow> w \<in> inst_syms R' \<alpha>"
+unfolding inst_syms_def concats_def
+by (metis (no_types, lifting) foldr_map_mono in_mono inst_sym_mono le_fun_def)
+
+lemma omega_cont_Lang_lfp_iterates: "omega_chain (\<lambda>n. ((subst_lang P)^^n) (\<lambda>A. {}))"
+  using omega_chain_iterates[OF mono_if_omega_cont, OF omega_cont_Lang_lfp]
+  unfolding bot_fun_def by blast
+
+lemma in_subst_langD_inst_syms: "w \<in> subst_lang P L A \<Longrightarrow> \<exists>\<alpha>. (A,\<alpha>)\<in>P \<and> w \<in> inst_syms L \<alpha>"
+unfolding subst_lang_def inst_syms_def Rhss_def by (auto split: prod.splits)
+
+lemma foldr_conc_conc: "foldr (@@) xs {[]} @@ A = foldr (@@) xs A"
+by (induction xs)(auto simp: conc_assoc)
+
+lemma derives_if_inst_syms:
+  "w \<in> inst_syms (\<lambda>A. {w. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w}) \<alpha> \<Longrightarrow> P \<turnstile> \<alpha> \<Rightarrow>* map Tm w"
+proof (induction \<alpha> arbitrary: w)
+  case Nil
+  then show ?case unfolding inst_syms_def concats_def by(auto)
+next
+  case (Cons s \<alpha>)
+  show ?case
+  proof (cases s)
+    case (Nt A)
+    then show ?thesis using Cons
+      unfolding inst_syms_def concats_def inst_sym_def by(fastforce simp: derives_Cons_decomp)
+  next
+    case (Tm a)
+    then show ?thesis using Cons
+      unfolding inst_syms_def concats_def inst_sym_def by(auto simp:derives_Tm_Cons)
+  qed
+qed
+
+lemma derives_if_in_subst_lang: "w \<in> ((subst_lang P)^^n) (\<lambda>A. {}) A \<Longrightarrow> P \<turnstile> [Nt A] \<Rightarrow>* map Tm w"
+proof(induction n arbitrary: w A)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  let ?L = "((subst_lang P)^^n) (\<lambda>A. {})"
+  have *: "?L A \<subseteq> {w. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w}" for A
+    using Suc.IH by blast
+  obtain \<alpha> where \<alpha>: "(A,\<alpha>) \<in> P" "w \<in> inst_syms ?L \<alpha>"
+    using in_subst_langD_inst_syms[OF Suc.prems[simplified]] by blast
+  show ?case using \<alpha>(1) derives_if_inst_syms[OF inst_syms_mono[OF *, of _ "\<lambda>A. A", OF \<alpha>(2)]]
+    by (simp add: derives_Cons_rule)
+qed
+
+lemma derives_if_Lang_lfp: "w \<in> Lang_lfp P A \<Longrightarrow> P \<turnstile> [Nt A] \<Rightarrow>* map Tm w"
+  unfolding Lang_lfp_SUP using derives_if_in_subst_lang
+  by (metis (mono_tags, lifting) SUP_apply UN_E)
+
+lemma Lang_lfp_subset_Lang: "Lang_lfp P A \<subseteq> Lang P A"
+unfolding Lang_def by(blast intro:derives_if_Lang_lfp)
+
+text \<open>The other direction:\<close>
+
+lemma inst_syms_decomp:
+  "\<lbrakk> \<forall>i < length ws. ws ! i \<in> inst_sym L (\<alpha> ! i); length \<alpha> = length ws \<rbrakk>
+  \<Longrightarrow> concat ws \<in> inst_syms L \<alpha>"
+proof (induction ws arbitrary: \<alpha>)
+  case Nil
+  then show ?case unfolding inst_syms_def concats_def by simp
+next
+  case (Cons w ws)
+  then obtain \<alpha>1 \<alpha>r where *: "\<alpha> = \<alpha>1 # \<alpha>r" by (metis Suc_length_conv)
+  with Cons.prems(2) have "length \<alpha>r = length ws" by simp
+  moreover from Cons.prems * have "\<forall>i<length ws. ws ! i \<in> inst_sym L (\<alpha>r ! i)" by auto
+  ultimately have "concat ws \<in> inst_syms L \<alpha>r" using Cons.IH by blast
+  moreover from Cons.prems * have "w \<in> inst_sym L \<alpha>1" by fastforce
+  ultimately show ?case unfolding inst_syms_def concats_def using * by force
+qed
+
+lemma Lang_lfp_if_derives_aux: "P \<turnstile> [Nt A] \<Rightarrow>(n) map Tm w \<Longrightarrow> w \<in> ((subst_lang P)^^n) (\<lambda>A. {}) A"
+proof(induction n arbitrary: w A rule: less_induct)
+  case (less n)
+  show ?case
+  proof (cases n)
+    case 0 then show ?thesis using less.prems by auto
+  next
+    case (Suc m)
+    then obtain \<alpha> where \<alpha>_intro: "(A,\<alpha>) \<in> P" "P \<turnstile> \<alpha> \<Rightarrow>(m) map Tm w"
+      by (metis deriven_start1 less.prems nat.inject)
+    then obtain ws ms where *:
+      "w = concat ws \<and> length \<alpha> = length ws \<and> length \<alpha> = length ms
+        \<and> sum_list ms = m \<and> (\<forall>i < length ws. P \<turnstile> [\<alpha> ! i] \<Rightarrow>(ms ! i) map Tm (ws ! i))"
+      using derive_decomp_Tm by metis
+
+    have "\<forall>i < length ws. ws ! i \<in> inst_sym (\<lambda>A. ((subst_lang P)^^m) (\<lambda>A. {}) A) (\<alpha> ! i)"
+    proof (rule allI | rule impI)+
+      fix i
+      show "i < length ws \<Longrightarrow> ws ! i \<in> inst_sym ((subst_lang P ^^ m) (\<lambda>A. {})) (\<alpha> ! i)"
+      unfolding inst_sym_def
+      proof (induction "\<alpha> ! i")
+        case (Nt B)
+        with * have **: "ms ! i \<le> m"
+          by (metis elem_le_sum_list)
+        with Suc have "ms ! i < n" by force
+        from less.IH[OF this, of B "ws ! i"] Nt *
+          have "ws ! i \<in> (subst_lang P ^^ (ms ! i)) (\<lambda>A. {}) B" by fastforce
+        with omega_chain_mono[OF omega_cont_Lang_lfp_iterates, OF **]
+          have "ws ! i \<in> (subst_lang P ^^ m) (\<lambda>A. {}) B" by (metis le_funD subset_iff)
+        with Nt show ?case by (metis sym.simps(5))
+      next
+        case (Tm a)
+        with * have "P \<turnstile> map Tm [a] \<Rightarrow>(ms ! i) map Tm (ws ! i)" by fastforce
+        then have "ws ! i \<in> {[a]}" using deriven_from_TmsD by fastforce
+        with Tm show ?case by (metis sym.simps(6))
+      qed
+    qed
+
+    from inst_syms_decomp[OF this] * have "w \<in> inst_syms ((subst_lang P ^^ m) (\<lambda>A. {})) \<alpha>" by argo
+    with \<alpha>_intro have "w \<in> (subst_lang P) (\<lambda>A. (subst_lang P ^^ m) (\<lambda>A. {}) A) A"
+      unfolding subst_lang_def Rhss_def by force
+    with Suc show ?thesis by force
+  qed
+qed
+
+
+lemma Lang_lfp_if_derives: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w \<Longrightarrow> w \<in> Lang_lfp P A"
+proof -
+  assume "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w"
+  then obtain n where "P \<turnstile> [Nt A] \<Rightarrow>(n) map Tm w" by (meson rtranclp_power)
+  from Lang_lfp_if_derives_aux[OF this] have "w \<in> ((subst_lang P)^^n) (\<lambda>A. {}) A" by argo
+  with Lang_lfp_SUP show "w \<in> Lang_lfp P A" by (metis (mono_tags, lifting) SUP_apply UNIV_I UN_iff)
+qed
+
+theorem Lang_lfp_eq_Lang: "Lang_lfp P A = Lang P A"
+unfolding Lang_def by(blast intro: Lang_lfp_if_derives derives_if_Lang_lfp)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Disjoint_Union_CFG.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Disjoint_Union_CFG.thy
@@ -0,0 +1,241 @@
+(*
+Authors: Markus Gschoßmann, Tobias Nipkow
+*)
+
+section \<open>Disjoint Union of Sets of Productions\<close>
+
+theory Disjoint_Union_CFG
+imports
+  "Regular-Sets.Regular_Set"
+  Context_Free_Grammar
+begin
+
+text \<open>This theory provides lemmas relevant when combining the productions of two grammars
+with disjoint sets of nonterminals. In particular that the languages of the nonterminals of
+one grammar is unchanged by adding productions involving only disjoint nonterminals.\<close>
+
+lemma derivel_disj_Un_if:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "P \<union> P' \<turnstile> u \<Rightarrow>l v"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<turnstile> u \<Rightarrow>l v \<and> nts_syms v \<inter> Lhss P' = {}"
+proof -
+  from assms(2) obtain A w u' v' where
+        A_w: "(A, w)\<in>(P \<union> P')"
+      and u: "u = map Tm u' @ Nt A # v'"
+      and v: "v = map Tm u' @ w @ v'"
+    unfolding derivel_iff by fast
+  then have "(A,w) \<notin> P'" using assms(3) unfolding nts_syms_def Lhss_def by auto
+  then have "(A,w) \<in> P" using A_w by blast
+  with u v have "(A, w) \<in> P"
+      and u: "u = map Tm u' @ Nt A # v'"
+      and v: "v = map Tm u' @ w @ v'" by auto
+  then have "P \<turnstile> u \<Rightarrow>l v" using derivel.intros by fastforce
+  moreover have "nts_syms v \<inter> Lhss P' = {}"
+    using u v assms \<open>(A, w) \<in> P\<close> unfolding nts_syms_def Nts_def Rhs_Nts_def by auto
+  ultimately show ?thesis by fast
+qed
+
+lemma derive_disj_Un_if:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "P \<union> P' \<turnstile> u \<Rightarrow> v"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<turnstile> u \<Rightarrow> v \<and> nts_syms v \<inter> Lhss P' = {}"
+proof -
+  from assms(2) obtain A w u' v' where
+        A_w: "(A, w) \<in> P \<union> P'"
+      and u: "u = u' @ Nt A # v'"
+      and v: "v = u' @ w @ v'"
+    unfolding derive_iff by fast
+  then have "(A,w) \<notin> P'" using assms(3) unfolding nts_syms_def Lhss_def by auto
+  then have "(A,w) \<in> P" using A_w by blast
+  with u v have "(A, w) \<in> P" and u: "u = u' @ Nt A # v'" and v: "v = u' @ w @ v'" by auto
+  then have "P \<turnstile> u \<Rightarrow> v" using derive.intros by fastforce
+  moreover have "nts_syms v \<inter> Lhss P' = {}"
+    using u v assms \<open>(A, w) \<in> P\<close> unfolding nts_syms_def Nts_def Rhs_Nts_def by auto
+  ultimately show ?thesis by blast
+qed
+
+lemma deriveln_disj_Un_if:
+assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+shows "\<lbrakk> P \<union> P' \<turnstile> u \<Rightarrow>l(n) v;  nts_syms u \<inter> Lhss P' = {} \<rbrakk> \<Longrightarrow>
+  P \<turnstile> u \<Rightarrow>l(n) v \<and> nts_syms v \<inter> Lhss P' = {}"
+proof (induction n arbitrary: v)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n')
+  then obtain v' where split: "P \<union> P' \<turnstile> u \<Rightarrow>l(n') v' \<and> P \<union> P' \<turnstile> v' \<Rightarrow>l v"
+    by (meson relpowp_Suc_E)
+  with Suc have "P \<turnstile> u \<Rightarrow>l(n') v' \<and> nts_syms v' \<inter> Lhss P' = {}"
+    by fast
+  with Suc show ?case using assms derivel_disj_Un_if
+    by (metis split relpowp_Suc_I)
+qed
+
+lemma deriven_disj_Un_if:
+assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+shows "\<lbrakk> P \<union> P' \<turnstile> u \<Rightarrow>(n) v;  nts_syms u \<inter> Lhss P' = {} \<rbrakk> \<Longrightarrow>
+  P \<turnstile> u \<Rightarrow>(n) v \<and> nts_syms v \<inter> Lhss P' = {}"
+proof (induction n arbitrary: v)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n')
+  then obtain v' where split: "P \<union> P' \<turnstile> u \<Rightarrow>(n') v' \<and> P \<union> P' \<turnstile> v' \<Rightarrow> v"
+    by (meson relpowp_Suc_E)
+  with Suc have "P \<turnstile> u \<Rightarrow>(n') v' \<and> nts_syms v' \<inter> Lhss P' = {}"
+    by fast
+  with Suc show ?case using assms derive_disj_Un_if
+    by (metis split relpowp_Suc_I)
+qed
+
+lemma derivel_disj_Un_iff:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<union> P' \<turnstile> u \<Rightarrow>l v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>l v"
+using assms Un_derivel derivel_disj_Un_if by fastforce
+
+lemma derive_disj_Un_iff:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<union> P' \<turnstile> u \<Rightarrow> v \<longleftrightarrow> P \<turnstile> u \<Rightarrow> v"
+using assms Un_derive derive_disj_Un_if by fastforce
+
+lemma deriveln_disj_Un_iff:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<union> P' \<turnstile> u \<Rightarrow>l(n) v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>l(n) v"
+by (metis Un_derivel assms(1,2) deriveln_disj_Un_if relpowp_mono)
+
+lemma deriven_disj_Un_iff:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<union> P' \<turnstile> u \<Rightarrow>(n) v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>(n) v"
+by (metis Un_derive assms(1,2) deriven_disj_Un_if relpowp_mono)
+
+lemma derives_disj_Un_iff:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+      and "nts_syms u \<inter> Lhss P' = {}"
+    shows "P \<union> P' \<turnstile> u \<Rightarrow>* v \<longleftrightarrow> P \<turnstile> u \<Rightarrow>* v"
+by (simp add: deriven_disj_Un_iff[OF assms] rtranclp_power)
+
+lemma Lang_disj_Un1:
+  assumes "Rhs_Nts P \<inter> Lhss P' = {}"
+  and "S \<notin> Lhss P'"
+shows "Lang P S = Lang (P \<union> P') S"
+proof -
+  from assms(2) have "nts_syms [Nt S] \<inter> Lhss P' = {}" unfolding nts_syms_def Lhss_def by simp
+  then show ?thesis unfolding Lang_def
+    by (simp add: derives_disj_Un_iff[OF assms(1)])
+qed
+
+
+subsection \<open>Disjoint Concatenation\<close>
+
+lemma Lang_concat_disj:
+assumes "Nts P1 \<inter> Nts P2 = {}" "S \<notin> Nts P1 \<union> Nts P2 \<union> {S1,S2}" "S1 \<notin> Nts P2" "S2 \<notin> Nts P1"
+shows "Lang ({(S, [Nt S1,Nt S2])} \<union> (P1 \<union> P2)) S = Lang P1 S1 @@ Lang P2 S2"
+proof -
+  let ?P = "{(S, [Nt S1,Nt S2])} \<union> (P1 \<union> P2)"
+  let ?L1 = "Lang P1 S1"
+  let ?L2 = "Lang P2 S2"
+  have "Lang ?P S \<subseteq> ?L1 @@ ?L2" 
+    proof
+      fix w
+      assume "w \<in> Lang ?P S"
+      then have "?P \<turnstile> [Nt S] \<Rightarrow>* map Tm w" using Lang_def by fastforce
+      then obtain \<alpha> where "?P \<turnstile> \<alpha> \<Rightarrow>* map Tm w \<and> (S, \<alpha>) \<in> ?P" using derives_start1 by fast
+      then have dervs: "?P \<turnstile> [Nt S1, Nt S2] \<Rightarrow>* map Tm w" using assms unfolding Nts_def by auto
+      then obtain w1 w2 where "?P \<turnstile> [Nt S1] \<Rightarrow>* map Tm w1" "?P \<turnstile> [Nt S2] \<Rightarrow>* map Tm w2" "w = w1 @ w2"
+        using derives_append_decomp[of ?P "[Nt S1]" "[Nt S2]"] by (auto simp: map_eq_append_conv)
+      then have "P1 \<union> ({(S, [Nt S1,Nt S2])} \<union> P2) \<turnstile> [Nt S1] \<Rightarrow>* map Tm w1"
+        and "P2 \<union> ({(S, [Nt S1,Nt S2])} \<union> P1) \<turnstile> [Nt S2] \<Rightarrow>* map Tm w2" by (simp_all add: Un_commute)
+      from derives_disj_Un_iff[THEN iffD1, OF _ _ this(1)]
+        derives_disj_Un_iff[THEN iffD1, OF _ _ this(2)] assms
+      have "P1 \<turnstile> [Nt S1] \<Rightarrow>* map Tm w1" "P2 \<turnstile> [Nt S2] \<Rightarrow>* map Tm w2"
+        by (auto simp: Nts_Lhss_Rhs_Nts)
+      then have "w1 \<in> ?L1" "w2 \<in> ?L2" unfolding Lang_def by simp_all
+      with \<open>w = _\<close> show "w \<in> ?L1 @@ ?L2" by blast
+    qed
+    moreover
+    have "?L1 @@ ?L2 \<subseteq> Lang ?P S" 
+    proof
+      fix w
+      assume "w \<in> ?L1 @@ ?L2"
+      then obtain w1 w2 where w12_src: "w1 \<in> ?L1 \<and> w2 \<in> ?L2 \<and> w = w1 @ w2" using assms by blast
+      have "P1 \<subseteq> ?P" "P2 \<subseteq> ?P" by auto
+      from w12_src have 1: "P1 \<turnstile> [Nt S1] \<Rightarrow>* map Tm w1" and 2: "P2 \<turnstile> [Nt S2] \<Rightarrow>* map Tm w2"
+        using Lang_def by fast+
+      have "?P \<turnstile> [Nt S] \<Rightarrow> [Nt S1, Nt S2]" 
+        using derive.intros[of "S" "[Nt S1, Nt S2]" ?P "[]"] by auto
+      also have "?P \<turnstile> ... \<Rightarrow>* map Tm w1 @ [Nt S2]" using derives_append derives_mono[OF \<open>P1 \<subseteq> ?P\<close>]
+        using derives_append[OF derives_mono[OF \<open>P1 \<subseteq> ?P\<close> 1], of "[Nt S2]"] by simp
+      also have "?P \<turnstile> \<dots> \<Rightarrow>* map Tm w1 @ map Tm w2"
+        using derives_prepend[OF derives_mono[OF \<open>P2 \<subseteq> ?P\<close> 2]] by simp
+      ultimately have "?P \<turnstile> [Nt S] \<Rightarrow>* map Tm w" using w12_src by simp
+      then show "w \<in> Lang ?P S" unfolding Lang_def by auto
+    qed
+    ultimately show ?thesis by blast
+qed
+
+
+subsection \<open>Disjoint Union including start fork\<close>
+
+lemma derive_from_isolated_fork:
+  "\<lbrakk> A \<notin> Lhss P;  {(A,\<alpha>1),(A,\<alpha>2)} \<union> P \<turnstile> [Nt A] \<Rightarrow> \<beta> \<rbrakk> \<Longrightarrow> \<beta> = \<alpha>1 \<or> \<beta> = \<alpha>2"
+unfolding derive_singleton by(auto simp: Lhss_def)
+
+lemma derives_fork_if_derives1:
+  assumes "P \<turnstile> [Nt B1] \<Rightarrow>* map Tm w"
+  shows "{(A,[Nt B1]), (A,[Nt B2])} \<union> P \<turnstile> [Nt A] \<Rightarrow>* map Tm w" (is "?P \<turnstile> _ \<Rightarrow>* _")
+proof -
+  have "?P \<turnstile> [Nt A] \<Rightarrow> [Nt B1]" using derive_singleton by auto
+  also have "?P \<turnstile> [Nt B1] \<Rightarrow>* map Tm w" using assms
+    by (meson derives_mono sup_ge2)
+  finally show ?thesis .
+qed
+
+lemma derives_disj_if_derives_fork:
+  assumes "A \<notin> Nts P \<union> {B1,B2}"
+  and "{(A,[Nt B1]), (A,[Nt B2])} \<union> P \<turnstile> [Nt A] \<Rightarrow>* map Tm w" (is "?P \<turnstile> _ \<Rightarrow>* _")
+  shows "P \<turnstile> [Nt B1] \<Rightarrow>* map Tm w \<or> P \<turnstile> [Nt B2] \<Rightarrow>* map Tm w"
+proof -
+  obtain \<beta> where steps: "?P \<turnstile> [Nt A] \<Rightarrow> \<beta>" "?P \<turnstile> \<beta> \<Rightarrow>* map Tm w"
+    using converse_rtranclpE[OF assms(2)] by blast
+  have "\<beta> = [Nt B1] \<or> \<beta> = [Nt B2]"
+    using steps(1) derive_from_isolated_fork[of A P] assms(1) by(auto simp: Nts_Lhss_Rhs_Nts)
+  then show ?thesis
+    using steps(2) derives_disj_Un_iff[of P "{(A,[Nt B1]), (A,[Nt B2])}" \<beta>] assms
+    by(auto simp: Nts_Lhss_Rhs_Nts)
+qed
+
+lemma Lang_distrib_eq_Un_Lang2:
+  assumes "A \<notin> Nts P \<union> {B1,B2}"
+  shows "Lang ({(A,[Nt B1]),(A,[Nt B2])} \<union> P) A = (Lang P B1 \<union> Lang P B2)"
+    (is "Lang ?P _ = _" is "?L1 = ?L2")
+proof
+  show "?L2 \<subseteq> ?L1" unfolding Lang_def
+    using derives_fork_if_derives1[of P B1 _ A B2] derives_fork_if_derives1[of P B2 _ A B1]
+    by (auto simp add: insert_commute)
+next
+  show "?L1 \<subseteq> ?L2"
+    unfolding Lang_def using derives_disj_if_derives_fork[OF assms] by auto
+qed
+
+lemma Lang_disj_Un2:
+assumes "Nts P1 \<inter> Nts P2 = {}" "S \<notin> Nts(P1 \<union> P2) \<union> {S1,S2}" "S1 \<notin> Nts P2" "S2 \<notin> Nts P1"
+shows "Lang ({(S,[Nt S1]), (S,[Nt S2])} \<union> (P1 \<union> P2)) S = Lang P1 S1 \<union> Lang P2 S2"
+proof -
+  let ?P = "{(S, [Nt S1]), (S, [Nt S2])} \<union> (P1 \<union> P2)"
+  have "Lang ?P S = Lang (P1 \<union> P2) S1 \<union> Lang (P1 \<union> P2) S2"
+    using Lang_distrib_eq_Un_Lang2[OF assms(2)] by simp
+  moreover have "Lang (P1 \<union> P2) S1 = Lang P1 S1" using assms(1,3)
+    apply(subst Lang_disj_Un1[of P1 P2 S1]) unfolding Nts_Lhss_Rhs_Nts by blast+
+  moreover  have "Lang (P2 \<union> P1) S2 = Lang P2 S2" using assms(1,4)
+    apply(subst Lang_disj_Un1[of P2 P1 S2]) unfolding Nts_Lhss_Rhs_Nts by blast+
+  ultimately show ?thesis
+    by (metis sup_commute)
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Epsilon_Elimination.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Epsilon_Elimination.thy
@@ -0,0 +1,386 @@
+(*
+Author: August Martin Stimpfle
+Based on HOL4 theories by Aditi Barthwal
+*)
+
+section \<open>Elimination of Epsilon Productions\<close>
+
+theory Epsilon_Elimination
+imports Context_Free_Grammar
+begin
+
+inductive nullable :: "('n,'t) prods \<Rightarrow> ('n,'t) sym \<Rightarrow> bool"
+for ps where
+  NullableSym:
+  "\<lbrakk> (A, w) \<in> set ps; \<forall>s \<in> set w. nullable ps s\<rbrakk>
+  \<Longrightarrow> nullable ps (Nt A)"
+
+abbreviation "nullables ps w \<equiv> (\<forall>s \<in> set w. nullable ps s)"
+
+lemma nullables_if: 
+  assumes "set ps \<turnstile> u \<Rightarrow>* v" 
+    and "u=[a]" "nullables ps v"
+  shows "nullables ps u"
+  using assms
+proof(induction arbitrary: a rule: rtranclp.induct)
+  case (rtrancl_refl a)
+  then show ?case by simp
+next
+  case (rtrancl_into_rtrancl u v w)
+  from \<open>set ps \<turnstile> v \<Rightarrow> w\<close> obtain A \<alpha> l r where A\<alpha>: "v = l @ [Nt A] @ r \<and> w = l @ \<alpha> @ r \<and> (A, \<alpha>) \<in> set ps"
+    by (auto simp: derive.simps)
+  from this \<open>nullables ps w\<close> have "nullables ps \<alpha> \<and> nullables ps l \<and> nullables ps r"
+    by simp
+  hence "nullables ps [Nt A]"
+    using A\<alpha> nullable.simps by auto
+  from this \<open>nullables ps \<alpha> \<and> nullables ps l \<and> nullables ps r\<close> have "nullables ps v"
+    using A\<alpha> by auto
+  thus ?case
+    using rtrancl_into_rtrancl.IH rtrancl_into_rtrancl.prems(1) by blast
+qed
+
+lemma nullable_if: "set ps \<turnstile> [a] \<Rightarrow>* [] \<Longrightarrow> nullable ps a"
+  using nullables_if[of ps "[a]" "[]" a] by simp
+
+lemma nullable_aux: "\<forall>s\<in>set gamma. nullable ps s \<and> set ps \<turnstile> [s] \<Rightarrow>* [] \<Longrightarrow> set ps \<turnstile> gamma \<Rightarrow>* []"
+proof (induction gamma)
+  case (Cons a list)
+  hence "set ps \<turnstile> list \<Rightarrow>* []"
+    by simp
+  moreover have "set ps \<turnstile> [a] \<Rightarrow>* []"
+    using Cons by simp
+  ultimately show ?case 
+    using derives_Cons[of \<open>set ps\<close> list \<open>[]\<close> \<open>a\<close>] by simp
+qed simp
+
+lemma if_nullable: "nullable ps a \<Longrightarrow> set ps \<turnstile> [a] \<Rightarrow>* []"
+proof (induction rule: nullable.induct)
+  case (NullableSym x gamma)
+    hence "set ps \<turnstile> [Nt x] \<Rightarrow>* gamma" 
+      using derive_singleton by blast
+    also have "set ps \<turnstile> gamma \<Rightarrow>* []" 
+      using NullableSym nullable_aux by blast
+  finally show ?case .
+qed
+
+corollary nullable_iff: "nullable ps a \<longleftrightarrow> set ps \<turnstile> [a] \<Rightarrow>* []"
+  by (auto simp: nullable_if if_nullable)
+
+fun eps_closure :: "('n, 't) prods \<Rightarrow> ('n, 't) syms \<Rightarrow> ('n, 't) syms list" where
+  "eps_closure ps [] = [[]]" |
+  "eps_closure ps (s#sl) = (
+    if nullable ps s then (map ((#) s) (eps_closure ps sl)) @ eps_closure ps sl 
+    else map ((#) s) (eps_closure ps sl))"
+
+definition eps_elim :: "('n, 't) prods \<Rightarrow> ('n, 't) Prods" where
+"eps_elim ps \<equiv> {(l,r'). \<exists>r. (l,r) \<in> set ps \<and> r' \<in> set (eps_closure ps r) \<and> (r' \<noteq> [])}"
+
+definition eps_elim_rel :: "('n,'t) prods \<Rightarrow> ('n,'t) prods \<Rightarrow> bool" where 
+  "eps_elim_rel ps ps'\<equiv> set ps'= eps_elim ps"
+
+lemma Eps_free_if_eps_elim_rel: "eps_elim_rel ps ps' \<Longrightarrow> Eps_free (set ps')"
+unfolding eps_elim_rel_def eps_elim_def Eps_free_def by blast
+
+(* auxiliary function to prove finiteness *)
+definition eps_elim_fun :: "('n, 't) prods \<Rightarrow> ('n, 't) prod \<Rightarrow> ('n, 't) Prods" where 
+  "eps_elim_fun ps p = {(l',r'). l' = fst p \<and> r' \<in> set (eps_closure ps (snd p)) \<and> (r' \<noteq> [])}"
+
+lemma eps_elim_fun_eq: "eps_elim ps = \<Union>((eps_elim_fun ps) ` set ps)"
+proof 
+  show "eps_elim ps \<subseteq> (\<Union> (eps_elim_fun ps ` set ps))" 
+   unfolding eps_elim_def eps_elim_fun_def by auto
+next
+  show "\<Union>((eps_elim_fun ps) ` set ps) \<subseteq> eps_elim ps"
+  proof
+    fix x
+    assume "x \<in> \<Union>((eps_elim_fun ps) ` set ps)"
+    obtain l r' where "x = (l,r')" by fastforce
+    hence "(l,r') \<in> \<Union>((eps_elim_fun ps) ` set ps)" 
+      using \<open>x \<in> \<Union>((eps_elim_fun ps) ` set ps)\<close> by simp
+    hence 1: "\<exists>r. r' \<in> set (eps_closure ps r) \<and> (r' \<noteq> []) \<and> (l,r) \<in> set ps" 
+      using eps_elim_fun_def by fastforce
+    from this  obtain r where "r' \<in> set (eps_closure ps r) \<and> (l,r) \<in> set ps" 
+      by blast
+    thus "x \<in> eps_elim ps" unfolding eps_elim_fun_def eps_elim_def
+      using 1 \<open>x = (l, r')\<close> by blast 
+  qed
+qed
+
+lemma finite_eps_elim: "finite (eps_elim ps)" 
+proof -
+  have "\<forall>p \<in> set ps. finite (eps_elim_fun ps p)"
+    unfolding eps_elim_fun_def by auto
+  hence "finite (\<Union>((eps_elim_fun ps) ` set ps))"
+    using finite_UN by simp
+  thus ?thesis using eps_elim_fun_eq by metis
+qed
+
+lemma eps_elim_rel_exists: "\<forall>ps. \<exists>ps'. eps_elim_rel ps ps'"
+  unfolding eps_elim_rel_def by (simp add: finite_list finite_eps_elim)
+
+lemma eps_closure_nullable:  "[] \<in> set (eps_closure ps w) \<Longrightarrow> nullables ps w"
+proof (induction w)
+  case Nil
+  then show ?case by simp
+next
+  case (Cons a r)
+  hence "nullable ps a"
+    using image_iff[of \<open>[]\<close> \<open>eps_closure ps\<close> \<open>{a#r}\<close>] by auto
+  then show ?case 
+    using Cons Un_iff by auto
+qed
+
+lemma eps_elim_rel_1: "r' \<in> set (eps_closure ps r) \<Longrightarrow> set ps \<turnstile> r \<Rightarrow>* r'"
+proof (induction r arbitrary: r')
+  case (Cons a r)
+  then show ?case 
+  proof (cases "nullable ps a")
+    case True
+    obtain e where e: "e \<in> set (eps_closure ps r) \<and> (r' = (a#e) \<or> r' = e)"
+      using Cons.prems True by auto
+    hence 1: "set ps \<turnstile> r \<Rightarrow>* e" 
+      using Cons.IH by blast
+    hence 2: "set ps \<turnstile> [a]@r \<Rightarrow>* [a]@e" 
+      using e derives_prepend by blast
+    have "set ps \<turnstile> [a] \<Rightarrow>* []" 
+      using True if_nullable by blast
+    hence "set ps \<turnstile> [a]@r \<Rightarrow>* r" 
+      using derives_append by fastforce
+    thus ?thesis 
+      using 1 2 e by force 
+  next
+    case False
+    obtain e where e: "e \<in> set (eps_closure ps r) \<and> (r' = (a#e))"
+      using Cons.prems False by auto
+    hence "set ps \<turnstile> r \<Rightarrow>* e" 
+      using Cons.IH by simp
+    hence "set ps \<turnstile> [a]@r \<Rightarrow>* [a]@e" 
+      using derives_prepend by blast
+    thus ?thesis
+      using e by simp
+  qed
+qed simp
+
+lemma eps_elim_rel_r2: 
+  assumes "set ps' \<turnstile> u \<Rightarrow> v" and "eps_elim_rel ps ps'" 
+  shows "set ps \<turnstile> u \<Rightarrow>* v"
+  using assms 
+proof -
+  obtain A \<alpha> x y where A: "(A, \<alpha>) \<in> set ps'\<and> u = x @ [Nt A] @ y \<and> v = x @ \<alpha> @ y"
+    using assms derive.cases by meson
+  hence 1: "(A, \<alpha>) \<in> {(l,r'). \<exists>r. (l,r) \<in> set ps \<and> r' \<in> set (eps_closure ps r) \<and> (r' \<noteq> [])}"
+    using assms(2) unfolding eps_elim_rel_def eps_elim_def by simp
+  obtain r where r: "(A, r) \<in> set ps \<and> \<alpha> \<in> set (eps_closure ps r)"
+    using 1 by blast
+  hence "set ps \<turnstile> r \<Rightarrow>* \<alpha>" 
+    using eps_elim_rel_1 by blast
+  hence 2: "set ps \<turnstile> x @ r @ y \<Rightarrow>* x @ \<alpha> @ y"
+    using r derives_prepend derives_append by blast
+  hence "set ps \<turnstile> x @ [Nt A] @ y \<Rightarrow> x @ r @ y" 
+    using r derive.simps by fast
+  thus ?thesis 
+    using 2 by (simp add: A)
+qed
+
+lemma eps_elim_rel_r3:
+  assumes "set ps' \<turnstile> u \<Rightarrow>* v" and "eps_elim_rel ps ps'" 
+  shows "set ps \<turnstile> u \<Rightarrow>* v"
+    using assms by (induction v rule: rtranclp_induct) (auto simp: eps_elim_rel_r2 rtranclp_trans)
+
+lemma eps_elim_rel_r5: "r \<in> set (eps_closure ps r)" 
+  by (induction r) auto
+
+lemma eps_elim_rel_r4:
+  assumes "(l,r) \<in> set ps"
+    and "eps_elim_rel ps ps'"
+    and "(r' \<noteq> [])" 
+    and "r' \<in> set (eps_closure ps r)"
+  shows "(l,r') \<in> set ps'"
+  using assms unfolding eps_elim_rel_def eps_elim_def by blast
+
+lemma eps_elim_rel_r7: 
+  assumes "eps_elim_rel ps ps'"
+    and "set ps \<turnstile> [Nt A] \<Rightarrow> v"
+    and "v' \<in> set (eps_closure ps v) \<and> (v' \<noteq> [])"
+  shows "set ps' \<turnstile> [Nt A] \<Rightarrow> v'"
+proof -
+  have "(A,v) \<in> set ps" 
+    using assms(2) by (simp add: derive_singleton)
+  hence "(A,v') \<in> set ps'" 
+    using assms eps_elim_rel_r4 conjE by fastforce
+  thus ?thesis 
+    using derive_singleton by fast
+qed
+
+lemma eps_elim_rel_r12a: 
+  assumes "x' \<in> set (eps_closure ps x)"
+    and "y' \<in> set (eps_closure ps y)"
+  shows "(x'@y') \<in> set (eps_closure ps (x@y))"
+  using assms by (induction x arbitrary: x' y y' rule: eps_closure.induct) auto
+
+lemma eps_elim_rel_r12b:
+  assumes "x' \<in> set (eps_closure ps x)"
+    and "y' \<in> set (eps_closure ps y)"
+    and "z' \<in> set (eps_closure ps z)"
+  shows "(x'@y'@z') \<in> set (eps_closure ps (x@y@z))"
+  using assms 
+  by (induction x arbitrary: x' y y' z z' rule: eps_closure.induct) (auto simp: eps_elim_rel_r12a)
+
+lemma eps_elim_rel_r14:
+  assumes "r' \<in> set (eps_closure ps (x@y))"
+  shows "\<exists>x' y'. (r'=x'@y') \<and> x' \<in> set (eps_closure ps x) \<and> y' \<in> set (eps_closure ps y)"
+  using assms
+proof (induction x arbitrary: y r' rule: eps_closure.induct)
+  case (2 ps s sl)
+  then show ?case
+  proof -
+    have "\<exists>x' y'. s # x = x' @ y' \<and> (x' \<in> (#) s ` set (eps_closure ps sl) \<or> x' \<in> set (eps_closure ps sl)) \<and> y' \<in> set (eps_closure ps y)"
+      if "\<And>r' y. r' \<in> set (eps_closure ps (sl @ y)) \<Longrightarrow> \<exists>x' y'. r' = x' @ y' \<and> x' \<in> set (eps_closure ps sl) \<and> y' \<in> set (eps_closure ps y)"
+        and "nullable ps s"
+        and "x \<in> set (eps_closure ps (sl @ y))"
+        and "r' = s # x"
+      for x :: "('a, 'b) sym list"
+      using that by (metis append_Cons imageI)
+    moreover have "\<exists>x' y'. r' = x' @ y' \<and> (x' \<in> (#) s ` set (eps_closure ps sl) \<or> x' \<in> set (eps_closure ps sl)) \<and> y' \<in> set (eps_closure ps y)"
+      if "\<And>r' y. r' \<in> set (eps_closure ps (sl @ y)) \<Longrightarrow> \<exists>x' y'. r' = x' @ y' \<and> x' \<in> set (eps_closure ps sl) \<and> y' \<in> set (eps_closure ps y)"
+        and "nullable ps s"
+        and "r' \<in> set (eps_closure ps (sl @ y))"
+      using that by metis
+    moreover have "\<exists>x' y'. s # x = x' @ y' \<and> x' \<in> (#) s ` set (eps_closure ps sl) \<and> y' \<in> set (eps_closure ps y)"
+      if "\<And>r' y. r' \<in> set (eps_closure ps (sl @ y)) \<Longrightarrow> \<exists>x' y'. r' = x' @ y' \<and> x' \<in> set (eps_closure ps sl) \<and> y' \<in> set (eps_closure ps y)"
+        and "\<not> nullable ps s"
+        and "x \<in> set (eps_closure ps (sl @ y))"
+        and "r' = s # x"
+      for x :: "('a, 'b) sym list"
+      using that by (metis append_Cons imageI)
+    ultimately show ?thesis
+      using 2 by auto
+  qed
+qed simp
+
+lemma eps_elim_rel_r15:
+  assumes "set ps \<turnstile> [Nt S] \<Rightarrow>* u"
+    and "eps_elim_rel ps ps'"
+    and "v \<in> set (eps_closure ps u) \<and> (v \<noteq> [])"
+  shows "set ps' \<turnstile> [Nt S] \<Rightarrow>* v"
+  using assms
+proof (induction u arbitrary: v rule: derives_induct)
+  case base
+  then show ?case
+    by (cases "nullable ps (Nt S)") auto
+next
+  case (step x A y w)
+  then obtain x' w' y' where 
+    v: "(v = (x'@w'@y')) \<and> x' \<in> set (eps_closure ps x) \<and> w' \<in> set (eps_closure ps w) \<and> y' \<in> set (eps_closure ps y)"
+    using step eps_elim_rel_r14 by metis
+  then show ?case
+  proof (cases "w' = []")
+    case True
+      hence "v = x'@y'" 
+        using v by simp 
+      have "[] \<in> set (eps_closure ps w)"
+        using True v by simp
+      hence "nullables ps w"
+        using eps_closure_nullable by blast
+      hence "[] \<in> set (eps_closure ps [Nt A])" 
+        using step(2) NullableSym by fastforce
+      hence "(x'@y') \<in> set (eps_closure ps (x@[Nt A]@y))"
+        using eps_elim_rel_r12b[of x' ps x \<open>[]\<close> \<open>[Nt A]\<close> y' y] v by simp
+      then show ?thesis 
+        using \<open>v = x' @ y'\<close> step by blast
+  next
+    case False
+      have "(x'@[Nt A]@y') \<in> set (eps_closure ps (x@[Nt A]@y)) "
+        using eps_elim_rel_r12b[of x' ps x \<open>[Nt A]\<close> \<open>[Nt A]\<close> y' y] eps_elim_rel_r5[of \<open>[Nt A]\<close> ps] v by blast
+      hence 1: "set ps' \<turnstile> [Nt S] \<Rightarrow>* (x'@[Nt A]@y')" 
+        using step by blast
+      have "set ps \<turnstile> [Nt A] \<Rightarrow> w" 
+        using step(2) derive_singleton by blast
+      hence "set ps' \<turnstile> [Nt A] \<Rightarrow> w'"
+        using eps_elim_rel_r7[of ps ps' A w w'] False step v by blast
+      hence "set ps' \<turnstile> (x'@[Nt A]@y') \<Rightarrow> (x'@w'@y')" 
+        using derive_append derive_prepend by blast
+      thus ?thesis using 1
+      by (simp add: v step.prems(2))
+  qed
+qed
+
+theorem eps_elim_rel_eq_if_noe:
+  assumes "eps_elim_rel ps ps'"
+    and "[] \<notin> lang ps S"
+  shows "lang ps S = lang ps' S"
+proof 
+  show "lang ps S \<subseteq> lang ps' S"
+  proof 
+    fix x
+    assume "x \<in> lang ps S"
+    have "\<forall>x. set ps \<turnstile> [Nt S] \<Rightarrow>* x \<longrightarrow> x \<noteq> []"
+      using assms Lang_def by fastforce
+    hence "(map Tm x) \<in> set (eps_closure ps (map Tm x))" 
+      using eps_elim_rel_r5 by auto
+    hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* (map Tm x)"
+      using assms \<open>x \<in> lang ps S\<close> Lang_def eps_elim_rel_r15[of ps S \<open>map Tm x\<close>] by fast
+    thus "x \<in> lang ps' S"
+      using Lang_def \<open>x \<in> lang ps S\<close> by fast 
+  qed
+next
+  show "lang ps' S \<subseteq> lang ps S"
+  proof
+    fix x'
+    assume "x' \<in> lang ps' S"
+    show "x' \<in> lang ps S" 
+      using assms Lang_def \<open>x' \<in> lang ps' S\<close> eps_elim_rel_r3[of ps' \<open>[Nt S]\<close> \<open>map Tm x'\<close> ps] by fast
+  qed
+qed
+
+(* correctness *)
+lemma noe_lang_eps_elim_rel_aux: 
+  assumes "ps \<turnstile> [Nt S] \<Rightarrow>* w" "w = []"  
+  shows "\<exists>A. ps \<turnstile> [Nt S] \<Rightarrow>* [Nt A] \<and> (A, w) \<in> ps"
+  using assms by (induction w rule: rtranclp_induct) (auto simp: derive.simps)
+
+lemma noe_lang_eps_elim_rel: "eps_elim_rel ps ps' \<Longrightarrow> [] \<notin> lang ps' S"
+proof (rule ccontr)
+  assume "eps_elim_rel ps ps'" "\<not>([] \<notin> lang ps' S)"
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* map Tm []"
+    using Lang_def by fast
+  hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* []"
+    by simp
+  hence "\<exists>A. set ps' \<turnstile> [Nt S] \<Rightarrow>* [Nt A] \<and> (A, []) \<in> set ps'"
+    using noe_lang_eps_elim_rel_aux[of \<open>set ps'\<close>] by blast
+  thus False 
+    using \<open>eps_elim_rel ps ps'\<close> unfolding eps_elim_rel_def eps_elim_def by blast
+qed
+
+theorem eps_elim_rel_lang_eq: "eps_elim_rel ps ps'\<Longrightarrow> lang ps' S = lang ps S - {[]}"
+proof 
+  assume "eps_elim_rel ps ps'"
+  show "lang ps' S \<subseteq> lang ps S - {[]}"
+  proof 
+    fix w
+    assume "w \<in> lang ps' S"
+    hence "w \<in> lang ps' S - {[]}"
+      using noe_lang_eps_elim_rel[of ps] \<open>eps_elim_rel ps ps'\<close> by simp
+    thus "w \<in> lang ps S - {[]}"
+      using \<open>eps_elim_rel ps ps'\<close> by (auto simp: Lang_def eps_elim_rel_r3)
+  qed
+next
+  assume "eps_elim_rel ps ps'"
+  show "lang ps S - {[]} \<subseteq> lang ps' S"
+  proof 
+    fix w
+    assume "w \<in> lang ps S - {[]}"
+    hence 1: "(map Tm w) \<noteq> []" 
+      by simp
+    have 2: "set ps \<turnstile> [Nt S] \<Rightarrow>* (map Tm w)"
+      using \<open>w \<in> lang ps S - {[]}\<close> Lang_def by fast
+    have "(map Tm w) \<in> set (eps_closure ps (map Tm w)) "
+      using \<open>w \<in> lang ps S - {[]}\<close> eps_elim_rel_r5 by blast
+    hence "set ps' \<turnstile> [Nt S] \<Rightarrow>* (map Tm w)"
+      using 1 2 eps_elim_rel_r15[of ps] \<open>eps_elim_rel ps ps'\<close> by simp
+    thus "w \<in> lang ps' S"
+      by (simp add: Lang_def)
+  qed
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Inlining1Prod.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Inlining1Prod.thy
@@ -0,0 +1,135 @@
+(*
+Authors: Kaan Taskin, Tobias Nipkow
+*)
+
+section \<open>Inlining a Single Production\<close>
+
+theory Inlining1Prod
+imports Context_Free_Grammar
+begin
+
+text 
+\<open>A single production of \<open>(A,\<alpha>)\<close> can be removed from \<open>ps\<close> by inlining
+(= replacing \<open>Nt A\<close> by \<open>\<alpha>\<close> everywhere in \<open>ps\<close>) without changing the language
+if \<open>Nt A \<notin> set \<alpha>\<close> and \<open>A \<notin> lhss ps\<close>.\<close>
+
+text
+\<open>\<open>substP ps s u\<close> replaces every occurrence of the symbol \<open>s\<close> with the list of symbols \<open>u\<close> on the right-hand sides of the production list \<open>ps\<close>\<close>
+
+definition substP :: "('n, 't) sym \<Rightarrow>  ('n, 't) syms \<Rightarrow> ('n, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "substP s u ps = map (\<lambda>(A,v). (A, substs s u v)) ps"
+
+lemma substP_split: "substP s u (ps @ ps') = substP s u ps @ substP s u ps'"
+  by (simp add: substP_def)
+
+lemma substP_skip1: "s \<notin> set v \<Longrightarrow> substP s u ((A,v) # ps) = (A,v) # substP s u ps"
+  by (auto simp add: substs_skip substP_def)
+
+lemma substP_skip2: "s \<notin> syms ps \<Longrightarrow> substP s u ps = ps"
+  by (induction ps) (auto simp add: substP_def substs_skip)
+
+lemma substP_rev: "Nt B \<notin> syms ps \<Longrightarrow> substP (Nt B) [s] (substP s [Nt B] ps) = ps"
+proof (induction ps)
+  case Nil
+  then show ?case
+    by (simp add: substP_def)
+next
+  case (Cons a ps)
+  let ?A = "fst a" let ?u = "snd a" let ?substs = "substs s [Nt B]"
+  have "substP (Nt B) [s] (substP s [Nt B] (a # ps)) = substP (Nt B) [s] (map (\<lambda>(A,v). (A, ?substs v)) (a#ps))"
+    by (simp add: substP_def)
+  also have "... = substP (Nt B) [s] ((?A, ?substs ?u) # map (\<lambda>(A,v). (A, ?substs v)) ps)"
+    by (simp add: case_prod_unfold)
+  also have "... = map ((\<lambda>(A,v). (A, substsNt B [s] v))) ((?A, ?substs ?u) # map (\<lambda>(A,v). (A, ?substs v)) ps)"
+    by (simp add: substP_def)
+  also have "... = (?A, substsNt B [s] (?substs ?u)) # substP (Nt B) [s] (substP s [Nt B] ps)"
+    by (simp add: substP_def)
+  also from Cons have "... = (?A, substsNt B [s] (?substs ?u)) # ps"
+    using set_subset_Cons unfolding Syms_def by auto
+  also from Cons.prems have "... = (?A, ?u) # ps"
+    using substs_rev unfolding Syms_def by force
+  also have "... = a # ps" by simp
+  finally show ?case .
+qed
+
+lemma substP_der:
+  "(A,u) \<in> set (substP (Nt B) v ps) \<Longrightarrow> set ((B,v) # ps) \<turnstile> [Nt A] \<Rightarrow>* u"
+proof -
+  assume "(A,u) \<in> set (substP (Nt B) v ps)"
+  then have "\<exists>u'. (A,u') \<in> set ps \<and> u = substsNt B v u'" unfolding substP_def by auto
+  then obtain u' where u'_def: "(A,u') \<in> set ps \<and> u = substsNt B v u'" by blast
+  then have path1: "set ((B,v) # ps) \<turnstile> [Nt A] \<Rightarrow>* u'"
+    by (simp add: derive_singleton r_into_rtranclp)
+  have "set ((B,v) # ps) \<turnstile> u' \<Rightarrow>* substsNt B v u'"
+    using substs_der by (metis list.set_intros(1))
+  with u'_def have path2: "set ((B,v) # ps) \<turnstile> u' \<Rightarrow>* u" by simp
+  from path1 path2 show ?thesis by simp
+qed
+
+text\<open>A list of symbols \<open>u\<close> can be derived before inlining if \<open>u\<close> can be derived after inlining.\<close>
+
+lemma if_part:
+  "set (substP (Nt B) v ps) \<turnstile> [Nt A] \<Rightarrow>* u \<Longrightarrow> set ((B,v) # ps) \<turnstile> [Nt A] \<Rightarrow>* u"
+proof (induction rule: derives_induct)
+  case (step u A v w)
+  then show ?case
+    by (meson derives_append derives_prepend rtranclp_trans substP_der)
+qed simp
+
+text
+\<open>For the other implication we need to take care that \<open>B\<close> can be derived in the original \<open>ps\<close>.
+Thus, after inlining, \<open>B\<close> must also be substituted in the derived sentence \<open>u\<close>:\<close>
+
+lemma only_if_lemma:
+assumes "A \<noteq> B"
+    and "B \<notin> lhss ps"
+    and "Nt B \<notin> set v"
+shows "set ((B,v)#ps) \<turnstile> [Nt A] \<Rightarrow>* u \<Longrightarrow> set (substP (Nt B) v ps) \<turnstile> [Nt A] \<Rightarrow>* substsNt B v u"
+proof (induction rule: derives_induct)
+  case base
+  then show ?case using assms(1) by simp
+next
+  case (step s B' w v')
+  then show ?case
+  proof (cases "B = B'")
+    case True
+    then have "v = v'" 
+      using \<open>B \<notin> lhss ps\<close> step.hyps unfolding Lhss_def by auto
+    then have "substsNt B v (s @ [Nt B'] @ w) = substsNt B v (s @ v' @ w)" 
+      using step.prems \<open>Nt B \<notin> set v\<close> True by (simp add: substs_skip)
+    then show ?thesis
+      using step.IH by argo
+  next
+    case False
+    then have "(B',v') \<in> set ps"
+      using step by auto
+    then have "(B', substsNt B v v') \<in> set (substP (Nt B) v ps)"
+      by (metis (no_types, lifting) list.set_map pair_imageI substP_def)
+    from derives_rule[OF this _ rtranclp.rtrancl_refl] step.IH False show ?thesis
+      by simp
+  qed
+qed
+
+text
+\<open>With the assumption that the non-terminal \<open>B\<close> is not in the list of symbols \<open>u\<close>, \<open>substs u (Nt B) v\<close> reduces to \<open>u\<close>\<close>
+
+corollary only_if_part: 
+assumes "A \<noteq> B"
+    and "B \<notin> lhss ps"
+    and "Nt B \<notin> set v"
+    and "Nt B \<notin> set u"
+shows "set ((B,v)#ps) \<turnstile> [Nt A] \<Rightarrow>* u \<Longrightarrow> set (substP (Nt B) v ps) \<turnstile> [Nt A] \<Rightarrow>* u"
+by (metis assms substs_skip only_if_lemma)
+
+text
+\<open>Combining the two implications gives us the desired property of language preservation\<close>
+
+lemma derives_inlining: 
+assumes "B \<notin> lhss ps" and
+        "Nt B \<notin> set v" and
+        "Nt B \<notin> set u" and
+        "A \<noteq> B"
+shows "set (substP (Nt B) v ps) \<turnstile> [Nt A] \<Rightarrow>* u \<longleftrightarrow> set ((B,v) # ps) \<turnstile> [Nt A] \<Rightarrow>* u"
+using assms if_part only_if_part by metis
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/NDA_rlin2.thy
--- /dev/null
+++ thys/Context_Free_Grammar/NDA_rlin2.thy
@@ -0,0 +1,201 @@
+(*
+Authors: Kaan Taskin, Tobias Nipkow
+*)
+
+section \<open>Strongly Right-Linear Grammars as a Nondeterministic Automaton\<close>
+
+theory NDA_rlin2
+imports Right_Linear
+begin
+
+text
+\<open>We define what is essentially the extended transition function of a nondeterministic automaton
+but is driven by a set of strongly right-linear productions \<open>P\<close>, which are of course
+just another representation of the transitions of a nondeterministic automaton.
+Function \<open>nxts_rlin2_set P M w\<close> traverses the terminals list \<open>w\<close>
+starting from the set of non-terminals \<open>M\<close> according to the productions of \<open>P\<close>.
+At the end it returns the reachable non-terminals.\<close>
+
+definition nxt_rlin2 :: "('n,'t)Prods \<Rightarrow> 'n \<Rightarrow> 't \<Rightarrow> 'n set" where
+"nxt_rlin2 P A a = {B. (A, [Tm a, Nt B]) \<in> P}"
+
+definition nxt_rlin2_set :: "('n,'t)Prods \<Rightarrow> 'n set \<Rightarrow> 't \<Rightarrow> 'n set" where
+"nxt_rlin2_set P M a = (\<Union>A\<in>M. nxt_rlin2 P A a)"
+
+definition nxts_rlin2_set :: "('n,'t)Prods \<Rightarrow> 'n set \<Rightarrow> 't list \<Rightarrow> 'n set" where
+"nxts_rlin2_set P = foldl (nxt_rlin2_set P)"
+
+lemma nxt_rlin2_nts:
+  assumes "B\<in>nxt_rlin2 P A a"
+  shows "B\<in>Nts P"
+  using assms nxt_rlin2_def Nts_def nts_syms_def by fastforce
+
+lemma nxts_rlin2_set_app: 
+  "nxts_rlin2_set P M (x @ y) = nxts_rlin2_set P (nxts_rlin2_set P M x) y"
+  unfolding nxts_rlin2_set_def by simp
+
+lemma nxt_rlin2_set_pick:
+  assumes "B \<in> nxt_rlin2_set P M a"
+  shows   "\<exists>C\<in>M. B \<in> nxt_rlin2_set P {C} a"
+  using assms by (simp add:nxt_rlin2_def nxt_rlin2_set_def)
+
+lemma nxts_rlin2_set_pick:
+  assumes "B \<in> nxts_rlin2_set P M w"
+  shows "\<exists>C\<in>M. B \<in> nxts_rlin2_set P {C} w"
+using assms proof (induction w arbitrary: B rule: rev_induct)
+  case Nil
+  then show ?case
+    by (simp add: nxts_rlin2_set_def)
+next
+  case (snoc x xs)
+  from snoc(2) have B_in: "B \<in> nxts_rlin2_set P (nxts_rlin2_set P M xs) [x]"
+    using nxts_rlin2_set_app[of P M xs "[x]"] by simp
+  hence "B \<in> nxt_rlin2_set P (nxts_rlin2_set P M xs) x"
+    by (simp add: nxts_rlin2_set_def)
+  hence "\<exists>C\<in>(nxts_rlin2_set P M xs). B \<in> nxt_rlin2_set P {C} x"
+    using nxt_rlin2_set_pick[of B P "nxts_rlin2_set P M xs" x] by simp
+  then obtain C where C_def: "C \<in> nxts_rlin2_set P M xs" and C_path: "B \<in> nxt_rlin2_set P {C} x"
+    by blast
+  have "\<exists>Ca\<in>M. C \<in> nxts_rlin2_set P {Ca} xs"
+    using snoc.IH[of C, OF C_def] .
+  then obtain D where *: "D \<in> M" and D_path: "C \<in> nxts_rlin2_set P {D} xs"
+    by blast
+  from C_path D_path have **: "B \<in> nxts_rlin2_set P {D} (xs @ [x])"
+    unfolding nxts_rlin2_set_def nxt_rlin2_set_def by auto
+  from * ** show ?case by blast
+qed
+
+lemma nxts_rlin2_set_first_step:
+  assumes "B \<in> nxts_rlin2_set P {A} (a # w)"
+  shows "\<exists>C \<in> nxt_rlin2 P A a. B \<in> nxts_rlin2_set P {C} w"
+proof -
+  from assms have "B \<in> nxts_rlin2_set P {A} ([a]@w)" by simp
+  hence "B \<in> nxts_rlin2_set P (nxts_rlin2_set P {A} [a]) w" 
+    using nxts_rlin2_set_app[of P "{A}" "[a]" w] by simp
+  hence "B \<in> nxts_rlin2_set P (nxt_rlin2 P A a) w"
+    by (simp add: nxt_rlin2_set_def nxts_rlin2_set_def)
+  thus "\<exists>C \<in> nxt_rlin2 P A a. B \<in> nxts_rlin2_set P {C} w"
+    using nxts_rlin2_set_pick[of B P "nxt_rlin2 P A a" w] by simp
+qed
+
+lemma nxts_trans0:
+  assumes "B \<in> nxts_rlin2_set P (nxts_rlin2_set P {A} x) z"
+  shows "B \<in> nxts_rlin2_set P {A} (x@z)"
+  by (metis assms foldl_append nxts_rlin2_set_def)
+
+lemma nxt_mono:
+  assumes "A \<subseteq> B"
+  shows "nxt_rlin2_set P A a \<subseteq> nxt_rlin2_set P B a"
+  unfolding nxt_rlin2_set_def using assms by blast
+
+lemma nxts_mono:
+  assumes "A \<subseteq> B"
+  shows "nxts_rlin2_set P A w \<subseteq> nxts_rlin2_set P B w"
+  unfolding nxts_rlin2_set_def proof (induction w rule:rev_induct)
+  case Nil
+  thus ?case by (simp add: assms)
+next
+  case (snoc x xs)
+  thus  ?case 
+    using nxt_mono[of "foldl (nxt_rlin2_set P) A xs" "foldl (nxt_rlin2_set P) B xs" P x] by simp
+qed
+
+lemma nxts_trans1:
+  assumes "M \<subseteq> nxts_rlin2_set P {A} x"
+      and "B \<in> nxts_rlin2_set P M z"
+  shows "B \<in> nxts_rlin2_set P {A} (x@z)"
+  using assms nxts_trans0[of B P A x z] nxts_mono[of M "nxts_rlin2_set P {A} x" P z, OF assms(1)] by auto
+
+lemma nxts_trans2:
+  assumes "C \<in> nxts_rlin2_set P {A} x"
+      and "B \<in> nxts_rlin2_set P {C} z"
+    shows "B \<in> nxts_rlin2_set P {A} (x@z)"
+  using assms nxts_trans1[of "{C}" P A x B z] by auto
+
+lemma nxts_to_mult_derive:
+  "B \<in> nxts_rlin2_set P M w \<Longrightarrow> (\<exists>A\<in>M. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt B])"
+unfolding nxts_rlin2_set_def proof (induction w arbitrary: B rule: rev_induct)
+  case Nil
+  hence 1: "B \<in> M" by simp
+  have 2: "P \<turnstile> [Nt B] \<Rightarrow>* map Tm [] @ [Nt B]" by simp
+  from 1 2 show ?case by blast
+next
+  case (snoc x xs)
+  from snoc.prems have "\<exists>C. C \<in> foldl (nxt_rlin2_set P) M xs \<and> (C, [Tm x, Nt B]) \<in> P"
+    unfolding nxt_rlin2_set_def nxt_rlin2_def by auto
+  then obtain C where C_xs: "C \<in> foldl (nxt_rlin2_set P) M xs" and C_prod: "(C, [Tm x, Nt B]) \<in> P" by blast
+  from C_xs obtain A where A_der: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ [Nt C]" and A_in: "A \<in> M" 
+    using snoc.IH[of C] by auto
+  from C_prod have "P \<turnstile> [Nt C] \<Rightarrow> [Tm x, Nt B]"
+    using derive_singleton[of P "Nt C" "[Tm x, Nt B]"] by blast
+  hence "P \<turnstile> map Tm xs @ [Nt C] \<Rightarrow> map Tm xs @ [Tm x, Nt B]"
+    using derive_prepend[of P "[Nt C]" "[Tm x, Nt B]" "map Tm xs"] by simp
+  hence C_der: "P \<turnstile> map Tm xs @ [Nt C] \<Rightarrow> map Tm (xs @ [x]) @ [Nt B]" by simp
+  from A_der C_der have "P \<turnstile> [Nt A] \<Rightarrow>* map Tm (xs @ [x]) @ [Nt B]" by simp
+  with A_in show ?case by blast
+qed
+
+lemma mult_derive_to_nxts:
+  assumes "rlin2 P"
+  shows "A\<in>M \<Longrightarrow> P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt B] \<Longrightarrow> B \<in> nxts_rlin2_set P M w"
+unfolding nxts_rlin2_set_def proof (induction w arbitrary: B rule: rev_induct)
+  case Nil
+  with assms have "A = B"
+    using rlin2_nts_derive_eq[of P A B] by simp
+  with Nil.prems(1) show ?case by simp
+next
+  case (snoc x xs)
+  from snoc.prems(2) have "P \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ [Tm x,Nt B]" by simp
+  with assms obtain C where C_der: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ [Nt C]"
+                        and C_prods: "(C,[Tm x, Nt B]) \<in> P" using rlin2_introduce_tm[of P A xs x B] by fast
+  from \<open>A \<in> M\<close> C_der have "C \<in> foldl (nxt_rlin2_set P) M xs"
+    using snoc.IH[of C] by auto
+  with C_prods show ?case
+    unfolding nxt_rlin2_set_def nxt_rlin2_def by auto
+qed
+
+text
+\<open>Acceptance of a word \<open>w\<close> w.r.t. \<open>P\<close> (starting from \<open>A\<close>), \<open>accepted P A w\<close>, means that we can reach
+an ``accepting'' nonterminal \<open>Z\<close>, i.e. one with a production \<open>(Z,[])\<close>.
+On the automaton level \<open>Z\<close> reachable final state.
+We show that \<open>accepted P A w\<close> iff \<open>w\<close> is in the language of \<open>A\<close> w.r.t. \<open>P\<close>.\<close>
+
+definition "accepted P A w = (\<exists>Z \<in> nxts_rlin2_set P {A} w. (Z,[]) \<in> P)"
+
+theorem accepted_if_Lang:
+  assumes "rlin2 P"
+      and "w \<in> Lang P A"
+    shows "accepted P A w"
+proof -
+  from assms obtain B where A_der: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt B]" and B_in: "(B,[]) \<in> P" 
+    unfolding Lang_def using rlin2_tms_eps[of P A w] by auto
+  from A_der have "B \<in> nxts_rlin2_set P {A} w" 
+    using mult_derive_to_nxts[OF assms(1)] by auto
+  with B_in show ?thesis 
+    unfolding accepted_def by blast
+qed
+
+theorem Lang_if_accepted: 
+  assumes "accepted P A w"
+    shows "w \<in> Lang P A"
+proof -
+  from assms obtain Z where Z_nxts: "Z \<in> nxts_rlin2_set P {A} w" and Z_eps: "(Z,[]) \<in> P"
+    unfolding accepted_def by blast
+  from Z_nxts obtain B where B_der: "P \<turnstile> [Nt B] \<Rightarrow>* map Tm w @ [Nt Z]" and B_in: "B \<in> {A}"
+    using nxts_to_mult_derive by fast
+  from B_in have A_eq_B: "A = B" by simp
+  from Z_eps have "P \<turnstile> [Nt Z] \<Rightarrow> []" 
+    using derive_singleton[of P "Nt Z" "[]"] by simp
+  hence "P \<turnstile> map Tm w @ [Nt Z] \<Rightarrow> map Tm w"
+    using derive_prepend[of P "[Nt Z]" "[]" "map Tm w"] by simp
+  with B_der A_eq_B have "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w" by simp
+  thus ?thesis 
+    unfolding Lang_def by blast
+qed
+
+theorem Lang_iff_accepted_if_rlin2:
+assumes "rlin2 P"
+shows "w \<in> Lang P A \<longleftrightarrow> accepted P A w"
+  using accepted_if_Lang[OF assms] Lang_if_accepted by fast
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Parse_Tree.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Parse_Tree.thy
@@ -0,0 +1,106 @@
+(* Author: Tobias Nipkow *)
+
+section \<open>Parse Trees\<close>
+
+theory Parse_Tree
+imports Context_Free_Grammar
+begin
+
+datatype ('n,'t) tree = Sym "('n,'t) sym"  | Rule 'n "('n,'t) tree list"
+datatype_compat tree
+
+fun root :: "('n,'t) tree \<Rightarrow> ('n,'t) sym" where
+"root(Sym s) = s" |
+"root(Rule A _) = Nt A"
+
+fun fringe :: "('n,'t) tree \<Rightarrow> ('n,'t) syms" where
+"fringe(Sym s) = [s]" |
+"fringe(Rule _ ts) = concat(map fringe ts)"
+
+abbreviation "fringes ts \<equiv> concat(map fringe ts)"
+
+fun parse_tree :: "('n,'t)Prods \<Rightarrow> ('n,'t) tree \<Rightarrow> bool" where
+"parse_tree P (Sym s) = True" |
+"parse_tree P (Rule A ts) = ((\<forall>t \<in> set ts. parse_tree P t) \<and> (A,map root ts) \<in> P)"
+
+lemma fringe_steps_if_parse_tree: "parse_tree P t \<Longrightarrow> P \<turnstile> [root t] \<Rightarrow>* fringe t"
+proof(induction t)
+  case (Sym s)
+  then show ?case by (auto)
+next
+  case (Rule A ts)
+  have "P \<turnstile> [Nt A] \<Rightarrow> map root ts"
+    using Rule.prems by (simp add: derive_singleton)
+  with Rule show ?case apply(simp)
+    using derives_concat1 by (metis converse_rtranclp_into_rtranclp)
+qed
+
+fun subst_pt and subst_pts where
+"subst_pt t' 0 (Sym _) = t'" |
+"subst_pt t' m (Rule A ts) = Rule A (subst_pts t' m ts)" |
+"subst_pts t' m (t#ts) =
+  (let n = length(fringe t) in if m < n then subst_pt t' m t # ts
+   else t # subst_pts t' (m-n) ts)"
+
+lemma fringe_subst_pt: "i < length(fringe t) \<Longrightarrow>
+  fringe(subst_pt t' i t) = take i (fringe t) @ fringe t' @ drop (Suc i) (fringe t)" and
+fringe_subst_pts: "i < length(fringes ts) \<Longrightarrow>
+  fringes (subst_pts t' i ts) =
+    take i (fringes ts) @ fringe t' @ drop (Suc i) (fringes ts)"
+proof(induction t' i t and t' i ts rule: subst_pt_subst_pts.induct)
+  case (3 t' m t ts)
+  let ?n = "length (fringe t)"
+  show ?case
+  proof (cases "m < ?n")
+    case True
+    thus ?thesis using "3.IH"(1)[OF refl] by (simp add: Let_def)
+  next
+    case False
+    thus ?thesis using "3.IH"(2)[OF refl] "3.prems" by (simp add: Suc_diff_le)
+  qed
+qed auto
+
+lemma root_subst_pt: "\<lbrakk> i < length(fringe t); fringe t ! i = Nt A; root t' = Nt A \<rbrakk>
+ \<Longrightarrow> root (subst_pt t' i t) = root t" and
+map_root_subst_pts: "\<lbrakk> i < length(fringes ts); fringes ts ! i = Nt A; root t' = Nt A \<rbrakk>
+ \<Longrightarrow> map root (subst_pts t' i ts) = map root ts"
+proof(induction t' i t and t' i ts rule: subst_pt_subst_pts.induct)
+  case (3 t' m t ts)
+  then show ?case by (auto simp add: nth_append Let_def)
+qed auto
+
+lemma parse_tree_subst_pt:
+  "\<lbrakk> parse_tree P t; i < length(fringe t); fringe t ! i = Nt A; parse_tree P t'; root t'= Nt A \<rbrakk>
+  \<Longrightarrow> parse_tree P (subst_pt t' i t)"
+and parse_tree_subst_pts:
+  "\<lbrakk> \<forall>t \<in> set ts. parse_tree P t; i < length(fringes ts); fringes ts ! i = Nt A; parse_tree P t'; root t'= Nt A \<rbrakk>
+  \<Longrightarrow> \<forall>t' \<in> set(subst_pts t' i ts). parse_tree P t'"
+proof(induction t' i t and t' i ts rule: subst_pt_subst_pts.induct)
+  case (2 m A ts)
+  then show ?case
+    using map_root_subst_pts by fastforce
+next
+  case (3 m t ts)
+  then show ?case by(auto simp add: nth_append Let_def)
+qed auto
+
+lemma parse_tree_if_derives: "P \<turnstile> [Nt A] \<Rightarrow>* w \<Longrightarrow> \<exists>t. parse_tree P t \<and> fringe t = w \<and> root t = Nt A"
+proof(induction rule: derives_induct)
+  case base
+  then show ?case
+    using fringe.simps(1) parse_tree.simps(1) root.simps(1) by blast
+next
+  case (step u A' v w)
+  then obtain t where 1: "parse_tree P t" and 2: "fringe t = u @ [Nt A'] @ v" and 3: \<open>root t = Nt A\<close>
+    by blast
+  let ?t' = "Rule A' (map Sym w)"
+  let ?t = "subst_pt ?t' (length u) t"
+  have "fringe ?t = u @ w @ v"
+    using 2 fringe_subst_pt[of "length u" t ?t'] by(simp add: o_def)
+  moreover have "parse_tree P ?t"
+    using parse_tree_subst_pt[OF 1, of "length u"] step.hyps(2) 2 by(simp add: o_def)
+  moreover have \<open>root ?t = Nt A\<close> by (simp add: 2 3 root_subst_pt)
+  ultimately show ?case by blast
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Pumping_Lemma_CFG.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Pumping_Lemma_CFG.thy
@@ -0,0 +1,382 @@
+(* Author: August Martin Stimpfle, Tobias Nipkow
+   Original theory by Thomas Ammer
+*)
+
+section "Pumping Lemma for Context Free Grammars"
+
+theory Pumping_Lemma_CFG
+imports
+  "List_Power.List_Power"
+  Chomsky_Normal_Form
+begin
+
+text \<open>Paths in the (implicit) parse tree of the derivation of some terminal word;
+specialized for productions in CNF.\<close>
+
+inductive path :: "('n, 't) Prods \<Rightarrow> 'n \<Rightarrow> 'n list \<Rightarrow> 't list \<Rightarrow> bool"
+   ("(2_ \<turnstile>/ (_/ \<Rightarrow>\<langle>_\<rangle>/ _))" [50, 0, 0, 50] 50) where
+terminal:  "(A, [Tm a]) \<in> P \<Longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>[A]\<rangle> [a]" |
+left:  "\<lbrakk>(A, [Nt B, Nt C]) \<in> P \<and> (P \<turnstile> B \<Rightarrow>\<langle>p\<rangle> w) \<and> (P \<turnstile> C \<Rightarrow>\<langle>q\<rangle> v)\<rbrakk> \<Longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p\<rangle> (w@v)" |
+right:  "\<lbrakk>(A, [Nt B, Nt C]) \<in> P \<and> (P \<turnstile> B \<Rightarrow>\<langle>p\<rangle> w) \<and> (P \<turnstile> C \<Rightarrow>\<langle>q\<rangle> v)\<rbrakk> \<Longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#q\<rangle> (w@v)"
+
+(* rules for derivations of words if the grammar is in cnf, resembles binary syntax trees *)
+inductive cnf_derives :: "('n, 't) Prods \<Rightarrow> 'n \<Rightarrow> 't list \<Rightarrow> bool"
+  ("(2_ \<turnstile>/ (_/ \<Rightarrow>cnf/ _))" [50, 0, 50] 50) where
+step:  "(A, [Tm a]) \<in> P \<Longrightarrow> P \<turnstile> A \<Rightarrow>cnf [a]"|
+trans:  "\<lbrakk>(A, [Nt B, Nt C]) \<in> P \<and> P \<turnstile> B \<Rightarrow>cnf w \<and> P \<turnstile> C \<Rightarrow>cnf v\<rbrakk> \<Longrightarrow> P \<turnstile> A \<Rightarrow>cnf (w@v)"
+
+lemma path_if_cnf_derives: "P \<turnstile> S \<Rightarrow>cnf w \<Longrightarrow> \<exists>p. P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w"
+  by (induction rule: cnf_derives.induct) (auto intro: path.intros)
+
+lemma cnf_derives_if_path: "P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w \<Longrightarrow> P \<turnstile> S \<Rightarrow>cnf w"
+  by (induction rule: path.induct) (auto intro: cnf_derives.intros)
+
+corollary cnf_path: "P \<turnstile> S \<Rightarrow>cnf w \<longleftrightarrow> (\<exists>p. P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w)"
+  using path_if_cnf_derives[of P] cnf_derives_if_path[of P] by blast
+
+lemma cnf_der: 
+  assumes "P \<turnstile> [Nt S] \<Rightarrow>* map Tm w" "CNF P" 
+  shows "P \<turnstile> S \<Rightarrow>cnf w"
+using assms proof (induction w arbitrary: S rule: length_induct)
+  case (1 w)
+  have "Eps_free P"
+    using assms(2) CNF_eq by blast
+  hence "\<not>(P \<turnstile> [Nt S] \<Rightarrow>* [])"
+    by (simp add: Eps_free_derives_Nil)
+  hence 2: "length w \<noteq> 0" 
+    using 1 by auto
+  thus ?case
+  proof (cases "length w \<le> 1")
+    case True
+    hence "length w = 1"
+      using 2 by linarith
+    then obtain t where "w = [t]"
+      using length_Suc_conv[of w 0] by auto 
+    hence "(S, [Tm t]) \<in> P"
+      using 1 assms(2) cnf_single_derive[of P S t] by simp
+    thus ?thesis
+      by (simp add: \<open>w = _\<close> cnf_derives.intros(1))
+  next
+    case False 
+    obtain A B u v where ABuv: "(S, [Nt A, Nt B]) \<in> P \<and> P \<turnstile> [Nt A] \<Rightarrow>* map Tm u \<and> P \<turnstile> [Nt B] \<Rightarrow>* map Tm v \<and> u@v = w \<and> u \<noteq> [] \<and> v \<noteq> []"
+      using False assms(2) 1 cnf_word[of P S w] by auto
+    have "length u < length w \<and> length v < length w"
+      using ABuv by auto
+    hence "cnf_derives P A u \<and> cnf_derives P B v"
+      using 1 ABuv by blast
+    thus ?thesis
+      using ABuv cnf_derives.intros(2)[of S A B P u v] by blast
+  qed
+qed
+
+lemma der_cnf:
+  assumes "P \<turnstile> S \<Rightarrow>cnf w" "CNF P"
+  shows "P \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+using assms proof (induction rule: cnf_derives.induct)
+  case (step A a P)
+  then show ?case 
+    by (simp add: derive_singleton r_into_rtranclp)
+next
+  case (trans A B C P w v)
+  hence "P \<turnstile> [Nt A] \<Rightarrow> [Nt B, Nt C]"
+    by (simp add: derive_singleton)
+  moreover have "P \<turnstile> [Nt B] \<Rightarrow>* map Tm w \<and> P \<turnstile> [Nt C] \<Rightarrow>* map Tm v"
+    using trans by blast
+  ultimately show ?case 
+    using derives_append_decomp[of P \<open>[Nt B]\<close> \<open>[Nt C]\<close> \<open>map Tm (w @ v)\<close>] by auto
+qed
+
+corollary cnf_der_eq: "CNF P \<Longrightarrow> (P \<turnstile> [Nt S] \<Rightarrow>* map Tm w \<longleftrightarrow> P \<turnstile> S \<Rightarrow>cnf w)"
+  using cnf_der[of P S w] der_cnf[of P S w] by blast
+
+lemma path_if_derives: 
+  assumes "P \<turnstile> [Nt S] \<Rightarrow>* map Tm w" "CNF P" 
+  shows "\<exists>p. P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w"
+  using assms cnf_der[of P S w] path_if_cnf_derives[of P S w] by blast
+
+lemma derives_if_path:
+  assumes "P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w" "CNF P"
+  shows "P \<turnstile> [Nt S] \<Rightarrow>* map Tm w"
+  using assms der_cnf[of P S w] cnf_derives_if_path[of P S p w] by blast
+
+text \<open>\<open>lpath\<close> = longest path, similar to \<open>path\<close>; \<open>lpath\<close> always chooses the longest path in a syntax tree\<close>
+inductive lpath :: "('n, 't) Prods \<Rightarrow> 'n \<Rightarrow> 'n list \<Rightarrow> 't list \<Rightarrow> bool" 
+   ("(2_ \<turnstile>/ (_/ \<Rightarrow>\<llangle>_\<rrangle>/ _))" [50, 0, 0, 50] 50) where
+terminal:  "(A, [Tm a]) \<in> P \<Longrightarrow> P \<turnstile> A \<Rightarrow>\<llangle>[A]\<rrangle> [a]" |
+nonTerminals:  "\<lbrakk>(A, [Nt B, Nt C]) \<in> P \<and> (P \<turnstile> B \<Rightarrow>\<llangle>p\<rrangle> w) \<and> (P \<turnstile> C \<Rightarrow>\<llangle>q\<rrangle> v)\<rbrakk> \<Longrightarrow> 
+                P \<turnstile> A \<Rightarrow>\<llangle>A#(if length p \<ge> length q then p else q)\<rrangle> (w@v)" 
+
+lemma path_lpath: "P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w \<Longrightarrow> \<exists>p'. (P \<turnstile> S \<Rightarrow>\<llangle>p'\<rrangle> w) \<and> length p' \<ge> length p"
+  by (induction rule: path.induct) (auto intro: lpath.intros)
+
+lemma lpath_path: "(P \<turnstile> S \<Rightarrow>\<llangle>p\<rrangle> w) \<Longrightarrow> P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w"
+  by (induction rule: lpath.induct) (auto intro: path.intros)
+
+(*Property of the a binary tree: Since the tree is a binary tree, word lengths are always \<le> 2^(longest path-1).
+However, this version is easier to use/proof and suffices *)
+lemma lpath_length: "(P \<turnstile> S \<Rightarrow>\<llangle>p\<rrangle> w) \<Longrightarrow> length w \<le> 2^length p"
+proof (induction rule: lpath.induct)
+  case (terminal A a P)
+  then show ?case by simp
+next
+  case (nonTerminals A B C P p w q v)
+  then show ?case 
+  proof (cases "length p \<ge> length q")
+    case True
+    hence "length v \<le> 2^length p"
+      using nonTerminals order_subst1[of \<open>length v\<close> \<open>\<lambda>x. 2^x\<close> \<open>length q\<close> \<open>length p\<close>] by simp
+    hence "length w +length v \<le> 2*2^length p"
+      by (simp add: nonTerminals.IH(1) add_le_mono mult_2)
+    then show ?thesis
+      by (simp add: True)
+  next
+    case False
+    hence "length w \<le> 2^length q"
+      using nonTerminals order_subst1[of \<open>length w\<close> \<open>\<lambda>x. 2^x\<close> \<open>length p\<close> \<open>length q\<close>] by simp 
+    hence "length w +length v \<le> 2*2^length q"
+      by (simp add: nonTerminals.IH add_le_mono mult_2)
+    then show ?thesis
+      by (simp add: False)
+  qed
+qed
+
+lemma step_decomp:
+  assumes "P \<turnstile> A \<Rightarrow>\<langle>[A]@p\<rangle> w" " length p \<ge> 1" 
+  shows "\<exists>B C p' a b. (A, [Nt B, Nt C]) \<in> P \<and> w =a@b \<and>
+        ((P \<turnstile> B \<Rightarrow>\<langle>p\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>p'\<rangle> b) \<or> (P \<turnstile> B \<Rightarrow>\<langle>p'\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>p\<rangle> b))"
+  using assms by (cases) fastforce+
+
+lemma steplp_decomp:
+  assumes "P \<turnstile> A \<Rightarrow>\<llangle>[A]@p\<rrangle> w" " length p \<ge> 1" 
+  shows "\<exists>B C p' a b. (A, [Nt B, Nt C]) \<in> P \<and> w =a@b \<and>
+      ((P \<turnstile> B \<Rightarrow>\<llangle>p\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p'\<rrangle> b \<and> length p \<ge> length p') \<or>
+       (P \<turnstile> B \<Rightarrow>\<llangle>p'\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p\<rrangle> b \<and> length p \<ge> length p'))"
+  using assms by (cases) fastforce+
+
+lemma path_first_step: "P \<turnstile> A \<Rightarrow>\<langle>p\<rangle> w \<Longrightarrow> \<exists>q. p = [A]@q "
+  by (induction rule: path.induct) simp_all
+
+lemma no_empty: "P \<turnstile> A \<Rightarrow>\<langle>p\<rangle> w \<Longrightarrow> length w > 0"
+  by (induction rule: path.induct) simp_all
+
+lemma substitution: 
+  assumes "P \<turnstile> A \<Rightarrow>\<langle>p1@[X]@p2\<rangle> z"
+  shows "\<exists>v w x. P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> w \<and> z = v@w@x \<and>
+        (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>p1@[X]@p'\<rangle> v@w'@x) \<and>
+        (length p1 > 0 \<longrightarrow> length (v@x) > 0)"
+using assms proof (induction p1 arbitrary: P A z)
+  case Nil
+  hence "\<forall>w' p'. ((P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> z) \<and> z = []@z@[] \<and>
+        (P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>[]@[X]@p'\<rangle> ([]@w'@[])) \<and>
+        (length [] > 0 \<longrightarrow> length ([]@[]) > 0))"
+    using path_first_step[of P A] by auto
+  then show ?case 
+    by blast
+next
+  case (Cons A p1 P Y)
+  hence 0: "A = Y"
+    using path_first_step[of P Y] by fastforce 
+  have "length (p1@[X]@p2) \<ge> 1"
+    by simp
+  hence "\<exists>B C a b q. (A, [Nt B, Nt C]) \<in> P \<and> a@b = z \<and> 
+        ((P \<turnstile> B \<Rightarrow>\<langle>q\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>p1@[X]@p2\<rangle> b) \<or> (P \<turnstile> B \<Rightarrow>\<langle>p1@[X]@p2\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>q\<rangle> b))"
+    using Cons.prems path_first_step step_decomp by fastforce
+  then obtain B C a b q where BC: "(A, [Nt B, Nt C]) \<in> P \<and> a@b = z \<and> 
+        ((P \<turnstile> B \<Rightarrow>\<langle>q\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>p1@[X]@p2\<rangle> b) \<or> (P \<turnstile> B \<Rightarrow>\<langle>p1@[X]@p2\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>q\<rangle> b))"
+    by blast
+  then show ?case
+  proof (cases "P \<turnstile> B \<Rightarrow>\<langle>q\<rangle> a \<and> P \<turnstile> C \<Rightarrow>\<langle>p1@[X]@p2\<rangle> b")
+    case True
+    then obtain v w x where vwx: "P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> w \<and> b = v@w@x \<and> 
+          (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> C \<Rightarrow>\<langle>p1@[X]@p'\<rangle> (v@w'@x))"
+      using Cons.IH by blast
+    hence 1: "\<forall>w' p'. (P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w') \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> (a@v@w'@x)"
+      using BC by (auto intro: path.intros(3))
+    obtain v' where "v' = a@v"
+      by simp
+    hence "length (v'@x) > 0"
+      using True no_empty by fast
+    hence "P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> w \<and> z = v'@w@x \<and> (\<forall>w' p'.
+           P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> (v'@w'@x)) \<and>
+          (length (A#p1) > 0 \<longrightarrow> length (v'@x) >0)"
+      using vwx 1 BC \<open>v' = _\<close> by simp
+    thus ?thesis
+      using 0 by auto
+  next
+    case False
+    then obtain v w x where vwx: "(P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> w) \<and> a = v@w@x \<and> 
+          (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> B \<Rightarrow>\<langle>p1@[X]@p'\<rangle> (v@w'@x))"
+      using Cons.IH BC by blast
+    hence 1: "\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> (v@w'@x@b)"
+      using BC left[of A B C P] by fastforce
+    hence "P \<turnstile> X \<Rightarrow>\<langle>[X]@p2\<rangle> w \<and> z = v@w@x@b \<and> (\<forall>w' p'.
+           P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> (v@w'@x@b)) \<and>
+          (length (A#p1) > 0 \<longrightarrow> length (a@v@x) >0)"
+      using vwx BC no_empty by fastforce
+    moreover have "length (v@x@b) > 0"
+      using no_empty BC by fast
+    ultimately show ?thesis
+    using 0 by auto
+  qed
+qed
+
+lemma substitution_lp: 
+  assumes "P \<turnstile> A \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> z"
+  shows "\<exists>v w x. P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> w \<and> z = v@w@x \<and>
+        (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>p1@[X]@p'\<rangle> v@w'@x)"
+using assms proof (induction p1 arbitrary: P A z)
+  case Nil
+  hence "\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>[]@[X]@p'\<rangle> ([]@w'@[])"
+    using path_first_step lpath_path by fastforce
+  moreover have "P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> z \<and> z = []@z@[]"
+    using Nil lpath.cases[of P A \<open>[X] @ p2\<close> z] by auto
+  ultimately show ?case 
+    by blast
+next
+  case (Cons A p1 P Y)
+  hence 0: "A = Y"
+    using path_first_step[of P Y] lpath_path by fastforce 
+  have "length (p1@[X]@p2) \<ge> 1"
+    by simp
+  hence "\<exists>B C p' a b. (A, [Nt B, Nt C]) \<in> P \<and> z = a@b \<and> 
+        ((P \<turnstile> B \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p'\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p') \<or> 
+         (P \<turnstile> B \<Rightarrow>\<llangle>p'\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p'))"
+    using steplp_decomp[of P A \<open>p1@[X]@p2\<close> z] 0 Cons by simp
+  then obtain B C p' a b where BC: "(A, [Nt B, Nt C]) \<in> P \<and> z = a@b \<and> 
+        ((P \<turnstile> B \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p'\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p') \<or> 
+         (P \<turnstile> B \<Rightarrow>\<llangle>p'\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p'))"
+    by blast
+  then show ?case
+  proof (cases "P \<turnstile> B \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p'\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p'")
+    case True
+    then obtain v w x where vwx: "P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> w \<and> a = v@w@x \<and> 
+          (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> B \<Rightarrow>\<langle>p1@[X]@p'\<rangle> v@w'@x)"
+      using Cons.IH by blast
+    hence "P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> w \<and> z = v@w@x@b \<and>
+       (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> v@w'@x@b)"
+      using BC lpath_path[of P] path.intros(2)[of A B C P] by fastforce
+    then show ?thesis
+      using 0 by auto
+  next
+    case False
+    hence "(P \<turnstile> B \<Rightarrow>\<llangle>p'\<rrangle> a \<and> P \<turnstile> C \<Rightarrow>\<llangle>p1@[X]@p2\<rrangle> b \<and> length (p1@[X]@p2) \<ge> length p')"
+      using BC by blast
+    then obtain v w x where vwx: "P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> w \<and> b = v@w@x \<and> 
+          (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> C \<Rightarrow>\<langle>p1@[X]@p'\<rangle> v@w'@x)"
+      using Cons.IH by blast
+    then obtain v' where "v' = a@v"
+      by simp
+    hence "P \<turnstile> X \<Rightarrow>\<llangle>[X]@p2\<rrangle> w \<and> z = v'@w@x \<and>
+       (\<forall>w' p'. P \<turnstile> X \<Rightarrow>\<langle>[X]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>A#p1@[X]@p'\<rangle> v'@w'@x)"
+      using BC lpath_path[of P] path.intros(3)[of A B C P] vwx by fastforce
+    then show ?thesis
+      using 0 by auto
+  qed
+qed
+
+lemma path_nts: "P \<turnstile> S \<Rightarrow>\<langle>p\<rangle> w \<Longrightarrow> set p \<subseteq> Nts P"
+  unfolding Nts_def by (induction rule: path.induct) auto
+
+lemma inner_aux: 
+  assumes "\<forall>w' p'. P \<turnstile> A \<Rightarrow>\<langle>[A]@p3\<rangle> w \<and> (P \<turnstile> A \<Rightarrow>\<langle>[A]@p'\<rangle> w' \<longrightarrow>
+          P \<turnstile> A \<Rightarrow>\<langle>[A]@p2@[A]@p'\<rangle> v@w'@x)"
+  shows "P \<turnstile> A \<Rightarrow>\<langle>(([A]@p2)^^(Suc i)) @ [A]@p3\<rangle> v^^(Suc i) @ w @ x^^(Suc i)"
+  using assms proof (induction i)
+  case 0
+  then show ?case by simp
+next
+  case (Suc i)
+  hence 1: "P \<turnstile> A \<Rightarrow>\<langle>[A]@p2@ ([A] @ p2)^^i @ [A]@p3\<rangle> v^^(Suc i) @ w @ x^^(Suc i)"
+    by simp
+  hence "P \<turnstile> A \<Rightarrow>\<langle>[A] @ p2 @ [A] @ p2@ ([A] @ p2)^^i @ [A]@p3\<rangle> v @ v^^(Suc i) @ w @ x^^(Suc i) @ x"
+    using assms by fastforce
+  thus ?case 
+    using pow_list_Suc2[of \<open>Suc i\<close> x] by simp
+qed
+
+lemma inner_pumping: 
+  assumes "CNF P" "finite P"
+    and "m = card (Nts P)"
+    and "z \<in> Lang P S"
+    and "length z \<ge> 2^(m+1)"
+  shows "\<exists>u v w x y . z = u@v@w@x@y \<and> length (v@w@x) \<le> 2^(m+1) \<and> length (v@x) \<ge> 1 \<and> (\<forall>i. u@(v^^i)@w@(x^^i)@y \<in> Lang P S)"
+proof -
+  obtain p' where p': "P \<turnstile> S \<Rightarrow>\<langle>p'\<rangle> z"
+    using assms Lang_def[of P S] path_if_derives[of P S] by blast
+  then obtain lp where lp: "P \<turnstile> S \<Rightarrow>\<llangle>lp\<rrangle> z"
+    using path_lpath[of P] by blast
+  hence 1: "set lp \<subseteq> Nts P"
+    using lpath_path[of P] path_nts[of P] by blast
+  have "length lp > m"
+  proof -
+    have "(2^(m+1)::nat) \<le> 2^length lp"
+      using lp lpath_length[of P S lp z] assms(5) le_trans by blast
+    hence "m+1 \<le> length lp" 
+      using power_le_imp_le_exp[of 2 \<open>m+1\<close> \<open>length lp\<close>] by auto
+    thus ?thesis
+      by simp
+  qed
+  then obtain l p where p: "lp = l@p \<and> length p = m+1"
+    using less_Suc_eq by (induction lp) fastforce+
+  hence "set l \<subseteq> Nts P \<and> set p \<subseteq> Nts P \<and> finite (Nts P)"
+    using \<open>finite P\<close> 1 finite_Nts[of P] assms(1) by auto
+  hence "card (set p) < length p"
+    using p assms(3) card_mono[of \<open>Nts P\<close> \<open>set p\<close>] by simp
+  then obtain A p1 p2 p3 where "p = p1@[A]@p2@[A]@p3"
+    by (metis distinct_card nat_neq_iff not_distinct_decomp)
+  then obtain u vwx y where uy: "(P \<turnstile> A \<Rightarrow>\<llangle>[A]@p2@[A]@p3\<rrangle> vwx \<and> z = u@vwx@y \<and>
+        (\<forall>w' p'. (P \<turnstile> A \<Rightarrow>\<langle>[A]@p'\<rangle> w' \<longrightarrow> P \<turnstile> S \<Rightarrow>\<langle>l@p1@[A]@p'\<rangle> u@w'@y)))"
+    using substitution_lp[of P S \<open>l@p1\<close> A \<open>p2@[A]@p3\<close> z] lp p by auto
+  hence "length vwx \<le> 2^(m+1)"
+    using \<open>p = _\<close> p lpath_length[of P A \<open>[A] @ p2 @ [A] @ p3\<close> vwx] order_subst1 by fastforce
+  then obtain v w x where vwx: "P \<turnstile> A \<Rightarrow>\<langle>[A]@p3\<rangle> w \<and> vwx = v@w@x \<and>
+        (\<forall>w' p'. (P \<turnstile> A \<Rightarrow>\<langle>[A]@p'\<rangle> w' \<longrightarrow> P \<turnstile> A \<Rightarrow>\<langle>[A]@p2@[A]@p'\<rangle> v@w'@x)) \<and>
+        (length ([A]@p2) > 0 \<longrightarrow> length (v@x) > 0)"
+    using substitution[of P A \<open>[A]@p2\<close> A p3 vwx] uy lpath_path[of P A] by auto
+  have "P \<turnstile> S \<Rightarrow>\<langle>l@p1@ (([A]@p2)^^i) @[A]@p3\<rangle> u@(v^^i)@w@(x^^i)@y" for i
+  proof -
+    have "\<forall>i. P \<turnstile> A \<Rightarrow>\<langle>([A]@p2)^^(Suc i) @ [A]@p3\<rangle> v^^(Suc i) @ w @ x^^(Suc i)"
+      using vwx inner_aux[of P A] by blast
+    hence "\<forall>i. P \<turnstile> S \<Rightarrow>\<langle>l@p1@(([A]@p2)^^(Suc i)) @[A]@p3\<rangle> u@ (v^^(Suc i)) @ w @ (x^^(Suc i)) @y"
+      using uy by fastforce
+    moreover have "P \<turnstile> S \<Rightarrow>\<langle>l@p1@(([A]@p2)^^0) @[A]@p3\<rangle> u@ (v^^0) @ w @ (x^^0) @y"
+      using vwx uy by auto
+    ultimately show "P \<turnstile> S \<Rightarrow>\<langle>l@p1@ (([A]@p2)^^i) @[A]@p3\<rangle> u@(v^^i)@w@(x^^i)@y"
+      by (induction i) simp_all
+  qed
+  hence "\<forall>i. u@(v^^i)@w@(x^^i)@y \<in> Lang P S"
+    unfolding Lang_def using assms(1) assms(2) derives_if_path[of P S] by blast
+  hence "z = u@v@w@x@y \<and> length (v@w@x) \<le> 2^(m+1) \<and> 1 \<le> length (v@x) \<and> (\<forall> i. u@(v^^i)@w@(x^^i)@ y \<in> Lang P S)"
+    using vwx uy \<open>length vwx \<le> 2 ^ (m + 1)\<close> by (simp add: Suc_leI)
+  then show ?thesis
+    by blast
+qed
+
+abbreviation "pumping_property L n \<equiv> \<forall>z \<in> L. length z \<ge> n \<longrightarrow>
+  (\<exists>u v w x y. z = u @ v @ w @ x @ y \<and> length (v@w@x) \<le> n \<and> length (v@x) \<ge> 1
+        \<and> (\<forall>i. u @ v^^i @ w @ x^^i @ y \<in> L))"
+
+theorem Pumping_Lemma_CNF:
+  assumes "CNF P" "finite P"
+  shows "\<exists>n. pumping_property (Lang P S) n"
+using inner_pumping[OF assms, of \<open>card (Nts P)\<close>] by blast
+
+theorem Pumping_Lemma:
+  assumes "finite (P :: ('n::infinite,'t)Prods)"
+  shows "\<exists>n. pumping_property (Lang P S) n"
+proof -
+  obtain ps where "set ps = P" using finite_list[OF assms] by blast
+  obtain ps' :: "('n,'t)prods" where ps': "CNF(set ps')" "lang ps' S= lang ps S - {[]}"
+    using cnf_exists[of S ps] by auto
+  let ?P' = "set ps'"
+  have P': "CNF ?P'" "finite ?P'" using ps'(1) by auto
+  from Pumping_Lemma_CNF[OF P', of S] obtain n where
+    pump: "pumping_property (Lang ?P' S) n" by blast
+  then have "pumping_property (Lang ?P' S) (Suc n)"
+    by (metis Suc_leD nle_le)
+  then have "pumping_property (Lang P S) (Suc n)"
+    using ps'(2) \<open>set ps = P\<close> by (metis Diff_iff list.size(3) not_less_eq_eq singletonD zero_le)
+  then show ?thesis by blast
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Pumping_Lemma_Regular.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Pumping_Lemma_Regular.thy
@@ -0,0 +1,359 @@
+(*
+Authors: Kaan Taskin
+*)
+
+section \<open>Pumping Lemma for Strongly Right-Linear Grammars\<close>
+
+theory Pumping_Lemma_Regular
+imports NDA_rlin2 "List_Power.List_Power"
+begin
+
+text \<open>The proof is on the level of strongly right-linear grammars.
+Currently there is no proof on the automaton level but now it would be easy to obtain one.\<close>
+
+lemma not_distinct:
+  assumes "m = card P"
+      and "m \<ge> 1"
+      and "\<forall>i < length w. w ! i \<in> P"
+      and "length w \<ge> Suc m"
+    shows "\<exists>xs ys zs y. w = xs @ [y] @ ys @ [y] @ zs \<and> length (xs @ [y] @ ys @ [y]) \<le> Suc m"
+using assms proof (induction w arbitrary: P m rule: length_induct)
+  case (1 aw)
+  from "1.prems"(4) obtain a w where aw_cons[simp]: "aw = a#w" and w_len: "m \<le> length w"
+    using Suc_le_length_iff[of m aw] by blast
+  show ?case proof (cases "a \<in> set w")
+    case True
+    hence "\<not> distinct aw" by simp
+    then obtain xs ys zs y where aw_split: "aw = xs @ [y] @ ys @ [y] @ zs"
+      using not_distinct_decomp by blast
+    show ?thesis proof (cases "length (xs @ [y] @ ys @ [y]) \<le> Suc m")
+      case True
+      with aw_split show ?thesis by blast
+    next
+      case False
+      let ?xsyys = "xs @ [y] @ ys"
+      from False have a4: "length ?xsyys \<ge> Suc m" by simp
+      from aw_split have a5: "length ?xsyys < length aw" by simp
+      with "1.prems"(3) have "\<forall>i<length ?xsyys. aw ! i \<in> P" by simp
+      with aw_split have a3: "\<forall>i < length ?xsyys. ?xsyys ! i \<in> P"
+        by (metis append_assoc nth_append)
+      from "1.prems"(2) "1.prems"(1) a3 a4 a5 have "\<exists>xs' ys' zs' y'. ?xsyys = xs' @ [y'] @ ys' @ [y'] @ zs' \<and> length (xs' @ [y'] @ ys' @ [y']) \<le> Suc m"
+        using "1.IH" by simp
+      then obtain xs' ys' zs' y' where xsyys_split: "?xsyys = xs' @ [y'] @ ys' @ [y'] @ zs'" and xsyys'_len: "length (xs' @ [y'] @ ys' @ [y']) \<le> Suc m" by blast
+      let ?xs = xs' let ?y = y' let ?ys = ys' let ?zs = "zs' @ [y] @ zs" 
+      from xsyys_split aw_split have *: "aw = ?xs @ [?y] @ ?ys @ [?y] @ ?zs" by simp
+      from xsyys'_len have **: "length (?xs @ [?y] @ ?ys @ [?y]) \<le> Suc m" by simp
+      from * ** show ?thesis by blast
+    qed
+  next
+    case False
+    let ?P' = "P - {a}"
+    from "1.prems"(3) have a_in: "a \<in> P" by auto
+    with "1.prems"(1) have a1: "m-1 = card ?P'" by simp
+    from "1.prems"(2) w_len have "w \<noteq> []" by auto
+    with "1.prems"(3) False have b_in: "\<exists>b\<noteq>a. b \<in> P" by force
+    from a_in b_in "1.prems"(2) "1.prems"(1) have "m \<ge> 2"
+      by (metis Suc_1 card_1_singletonE not_less_eq_eq singletonD verit_la_disequality)
+    hence a2: "m-1 \<ge> 1" by simp
+    from False "1.prems"(3) have a3: "\<forall>i<length w. w ! i \<in> ?P'"
+      using DiffD2 by auto
+    from "1.prems"(2) w_len have a4: "Suc (m-1) \<le> length w" by simp
+    from a1 a2 a3 a4 have "\<exists>xs ys zs y. w = xs @ [y] @ ys @ [y] @ zs \<and> length (xs @ [y] @ ys @ [y]) \<le> Suc (m - 1)"
+      using "1.IH" by simp
+    then obtain xs ys zs y where w_split: "w = xs @ [y] @ ys @ [y] @ zs" and xsys_len: "length (xs @ [y] @ ys @ [y]) \<le> m" by auto
+    from w_split have *: "a#w = (a#xs) @ [y] @ ys @ [y] @ zs" by simp
+    from xsys_len have **: "length ((a#xs) @ [y] @ ys @ [y]) \<le> Suc m" by simp
+    from * ** aw_cons show ?thesis by blast
+  qed
+qed
+
+text
+\<open>We define the function \<open>nxts_nts P a w\<close> that collects all paths traversing the word \<open>w\<close> starting from the non-terminal \<open>A\<close> in the 
+ production set \<open>P\<close>. \<open>nxts_nts0\<close> appends the non-terminal \<open>A\<close> in front of every list produced by \<open>nxts_nts\<close>\<close>
+
+fun nxts_nts :: "('n,'t)Prods \<Rightarrow> 'n \<Rightarrow> 't list \<Rightarrow> 'n list set" where
+  "nxts_nts P A [] = {[]}"
+| "nxts_nts P A (a#w) = (\<Union>B\<in>nxt_rlin2 P A a. (Cons B)`nxts_nts P B w)"
+
+definition nxts_nts0 where
+"nxts_nts0 P A w \<equiv> ((#) A) ` nxts_nts P A w"
+
+subsection \<open>Properties of \<open>nxts_nts\<close> and \<open>nxts_nts0\<close>\<close>
+
+lemma nxts_nts0_i0:
+  "\<forall>e \<in> nxts_nts0 P A w. e!0 = A"
+  unfolding nxts_nts0_def by auto
+
+lemma nxts_nts0_shift:
+  assumes "i < length w"
+  shows "\<forall>e \<in> nxts_nts0 P A w. \<exists>e' \<in> nxts_nts P A w. e ! (Suc i) = e' ! i"
+  unfolding nxts_nts0_def by auto
+
+lemma nxts_nts_pick_nt:
+  assumes "e \<in> nxts_nts P A (a#w)"
+  shows "\<exists>C\<in>nxt_rlin2 P A a. \<exists>e' \<in> nxts_nts P C w. e = C#e'"
+  using assms by auto
+
+lemma nxts_nts0_len:
+  "\<forall>e \<in> nxts_nts0 P A w. length e = Suc (length w)"
+  unfolding nxts_nts0_def
+  by (induction P A w rule: nxts_nts.induct) auto
+
+lemma nxts_nts0_nxt: 
+  assumes "i < length w"
+  shows "\<forall>e \<in> nxts_nts0 P A w. e!(Suc i) \<in> nxt_rlin2 P (e!i) (w!i)"
+  unfolding nxts_nts0_def using assms proof (induction P A w arbitrary: i rule: nxts_nts.induct)
+  case (1 P A)
+  thus ?case by simp
+next
+  case (2 P A a w)
+  thus ?case
+    using less_Suc_eq_0_disj by auto
+qed
+
+lemma nxts_nts0_path:
+  assumes "i1 \<le> length w"
+      and "i2 \<le> length w"
+      and "i1 \<le> i2"
+    shows "\<forall>e \<in> nxts_nts0 P A w. e!i2 \<in> nxts_rlin2_set P {e!i1} (drop i1 (take i2 w))"
+proof
+  fix e
+  assume "e \<in> nxts_nts0 P A w"
+  with assms show "e ! i2 \<in> nxts_rlin2_set P {e ! i1} (drop i1 (take i2 w))" proof (induction "i2-i1" arbitrary: i2)
+    case 0
+    thus ?case
+      by (simp add: nxts_rlin2_set_def)
+  next
+    case (Suc x)
+      let ?i2' = "i2 - 1"
+      from Suc.hyps(2) have x_def: "x = ?i2' - i1" by simp
+      from Suc.prems(2) have i2'_len: "?i2' \<le> length w" by simp
+      from Suc.prems(3) Suc.hyps(2) have i1_i2': "i1 \<le> ?i2'" by simp
+      have IH: "e ! ?i2' \<in> nxts_rlin2_set P {e ! i1} (drop i1 (take ?i2' w))"
+        using Suc.hyps(1)[of ?i2', OF x_def Suc.prems(1) i2'_len i1_i2' Suc.prems(4)] .
+      from Suc.hyps(2) Suc.prems(2) Suc.prems(4) have "e ! i2 \<in> nxt_rlin2 P (e!(i2-1)) (w!(i2-1))"
+        using nxts_nts0_nxt[of ?i2' w P A] by simp
+      hence e_i2: "e ! i2 \<in> nxts_rlin2_set P {e!(i2-1)} [w!(i2-1)]"
+        unfolding nxts_rlin2_set_def nxt_rlin2_set_def by simp
+      have "drop i1 (take (i2 - 1) w) @ [w ! (i2 - 1)] = drop i1 (take i2 w)"
+        by (smt (verit) Cons_nth_drop_Suc Suc.hyps(2) Suc.prems(2) Suc.prems(3) add_Suc drop_drop drop_eq_Nil drop_take i1_i2' i2'_len le_add_diff_inverse2 le_less_Suc_eq nle_le nth_via_drop order.strict_iff_not take_Suc_conv_app_nth x_def)
+      thus ?case 
+        using nxts_trans2[of "e ! (i2 - 1)" P "e ! i1" "drop i1 (take (i2 - 1) w)" "e ! i2" "[w!(i2-1)]", OF IH e_i2] by argo
+  qed
+qed
+
+lemma nxts_nts0_path_start:
+  assumes "i \<le> length w"
+  shows "\<forall>e \<in> nxts_nts0 P A w. e ! i \<in> nxts_rlin2_set P {A} (take i w)"
+  using assms nxts_nts0_path[of 0 w i P A] by (simp add: nxts_nts0_def)
+
+lemma nxts_nts_elem:
+  assumes "i < length w"
+  shows "\<forall>e \<in> nxts_nts P A w. e ! i \<in> Nts P"
+proof
+  fix e
+  assume "e \<in> nxts_nts P A w"
+  with assms show "e ! i \<in> Nts P" proof (induction P A w arbitrary: i e rule: nxts_nts.induct)
+    case (1 P A)
+    thus ?case by simp
+  next
+    case (2 P A a w)
+    from 2(3) obtain C e' where C_def: "C \<in> nxt_rlin2 P A a" and e'_def: "e' \<in> nxts_nts P C w" and e_app: "e = C#e'"
+      using nxts_nts_pick_nt[of e P A a w] by blast
+    show ?case proof (cases "i = 0")
+      case True
+      with e_app C_def show ?thesis
+        using nxt_rlin2_nts by simp
+    next
+      case False
+      from False 2(2) have i_len: "i - 1 < length w" by simp
+      have "e' ! (i - 1) \<in> Nts P" 
+        using "2.IH"[of C "i-1" e', OF C_def i_len e'_def] .
+      with e_app False have "e ! i \<in> Nts P" by simp
+      thus ?thesis .
+    qed
+  qed
+qed
+
+lemma nxts_nts0_elem:
+  assumes "A \<in> Nts P"
+      and "i \<le> length w"
+  shows "\<forall>e \<in> nxts_nts0 P A w. e ! i \<in> Nts P"
+proof (cases "i = 0")
+  case True
+  thus ?thesis
+    by (simp add: assms(1) nxts_nts0_i0)
+next
+  case False
+  show ?thesis proof
+    fix e
+    assume e_def: "e \<in> nxts_nts0 P A w"
+    from False e_def assms(2) have "\<exists>e' \<in> nxts_nts P A w. e ! i = e' ! (i-1)"
+      using nxts_nts0_shift[of "i-1" w P A] by simp
+    then obtain e' where e'_def: "e' \<in> nxts_nts P A w" and e_ind: "e ! i = e' ! (i-1)"
+      by blast
+    from False e'_def assms(2) have "e' ! (i-1) \<in> Nts P"
+      using nxts_nts_elem[of "i-1" w P A] by simp
+    with e_ind show "e ! i \<in> Nts P" by simp
+  qed
+qed
+
+lemma nxts_nts0_pick:
+  assumes "B \<in> nxts_rlin2_set P {A} w"
+  shows "\<exists>e \<in> nxts_nts0 P A w. last e = B"
+unfolding nxts_nts0_def using assms proof (induction P A w arbitrary: B rule: nxts_nts.induct)
+  case (1 P A)
+  thus ?case
+    by (simp add: nxts_rlin2_set_def)
+next
+  case (2 P A a w)
+  from 2(2) obtain C where C_def: "C \<in> nxt_rlin2 P A a" and C_path: "B \<in> nxts_rlin2_set P {C} w"
+    using nxts_rlin2_set_first_step[of B P A a w] by blast
+  have "\<exists>e \<in> nxts_nts0 P C w. last e = B"
+    using "2.IH"[of C B, OF C_def C_path] by (simp add: nxts_nts0_def)
+  then obtain e where e_def: "e \<in> nxts_nts0 P C w" and e_last: "last e = B" 
+    by blast
+  from e_def C_def have *: "A#e \<in> nxts_nts0 P A (a#w)"
+    unfolding nxts_nts0_def by auto
+  from e_last e_def have **: "last (A#e) = B"
+    using nxts_nts0_len[of P C w] by auto
+  from * ** show ?case
+    unfolding nxts_nts0_def by blast
+qed
+
+subsection \<open>Pumping Lemma\<close>
+
+text
+\<open>The following lemma states that in the automata level there exists a cycle occurring in the first \<open>m\<close> symbols where \<open>m\<close> is the cardinality 
+ of the non-terminals set, under the following assumptions\<close>
+
+lemma nxts_split_cycle:
+  assumes "finite P"
+      and "A \<in> Nts P"
+      and "m = card (Nts P)"
+      and "B \<in> nxts_rlin2_set P {A} w"
+      and "length w \<ge> m"
+    shows "\<exists>x y z C. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> m \<and>
+              C \<in> nxts_rlin2_set P {A} x \<and> C \<in> nxts_rlin2_set P {C} y \<and> B \<in> nxts_rlin2_set P {C} z"
+proof -
+  let ?nts = "nxts_nts0 P A w"
+  obtain e where e_def: "e \<in> ?nts" and e_last: "last e = B"
+    using nxts_nts0_pick[of B P A w, OF assms(4)] by auto
+  from e_def have e_len: "length e = Suc (length w)"
+    using nxts_nts0_len[of P A w] by simp
+  from e_len e_def have e_elem: "\<forall>i < length e. e!i \<in> Nts P"
+    using nxts_nts0_elem[OF assms(2)] by (auto simp: less_Suc_eq_le)
+  have "finite (Nts P)"
+    using finite_Nts[of P, OF assms(1)] .
+  with assms(2) assms(3) have m_geq_1: "m \<ge> 1"
+    using less_eq_Suc_le by fastforce
+  from assms(5) e_len have "\<exists>xs ys zs y. e = xs @ [y] @ ys @ [y] @ zs \<and> length (xs @ [y] @ ys @ [y]) \<le> Suc m"
+    using not_distinct[OF assms(3) m_geq_1 e_elem] by simp
+  then obtain xs ys zs C where e_split: "e = xs @ [C] @ ys @ [C] @ zs" and xy_len: "length (xs @ [C] @ ys @ [C]) \<le> Suc m"
+    by blast
+  let ?e1 = "xs @ [C]" let ?e2 = "ys @ [C]" let ?e3 = zs
+  let ?x = "take (length ?e1 - 1) w" let ?y = "drop (length ?e1 - 1) (take (length ?e1+length ?e2 - 1) w)"
+  let ?z = "drop (length ?e1+length ?e2 - 1) w"
+  have *: "w = ?x@?y@?z"
+    by (metis Nat.add_diff_assoc2 append_assoc append_take_drop_id diff_add_inverse drop_take le_add1 length_append_singleton plus_1_eq_Suc take_add)
+  from e_len e_split have **: "length ?y \<ge> 1" by simp
+  from xy_len have ***: "length (?x@?y) \<le> m" by simp
+  have x_fac: "?x = take (length xs) w" by simp
+  from ** have x_fac2: "length xs \<le> length w"  by simp
+  from e_split have x_fac3: "e ! length xs = C" by simp
+  from e_def x_fac x_fac3 have ****: "C \<in> nxts_rlin2_set P {A} ?x"
+    using nxts_nts0_path_start[of "length xs" w P A, OF x_fac2] by auto
+  have y_fac: "?y = drop (length xs) (take (length xs + length ys + 1) w)" by simp
+  from e_len e_split have y_fac2: "length xs + length ys + 1 \<le> length w" by simp
+  have y_fac3: "length xs \<le> length xs + length ys + 1" by simp
+  have y_fac4: "e ! (length xs + length ys + 1) = C"
+    by (metis add.right_neutral add_Suc_right append.assoc append_Cons e_split length_Cons length_append list.size(3) nth_append_length plus_1_eq_Suc)
+  from e_def y_fac x_fac3 y_fac4 have *****: "C \<in> nxts_rlin2_set P {C} ?y"
+    using nxts_nts0_path[of "length xs" w "length xs + length ys + 1" P A, OF x_fac2 y_fac2 y_fac3] by auto
+  have z_fac: "?z = drop (length xs + length ys + 1) (take (length w) w)" by simp
+  from e_last e_len have z_fac2: "e ! (length w) = B"
+    by (metis Zero_not_Suc diff_Suc_1 last_conv_nth list.size(3))
+  from e_def z_fac y_fac2 y_fac4 z_fac2 have ******: "B \<in> nxts_rlin2_set P {C} ?z"
+    using nxts_nts0_path[of "length xs + length ys + 1" w "length w" P A] by auto
+  from * ** *** **** ***** ****** show ?thesis by blast
+qed
+
+text
+\<open>We also show that a cycle can be pumped in the automata level\<close>
+
+lemma pump_cycle:
+  assumes "B \<in> nxts_rlin2_set P {A} x"
+      and "B \<in> nxts_rlin2_set P {B} y"
+    shows "B \<in> nxts_rlin2_set P {A} (x@(y^^i))"
+using assms proof (induction i)
+  case 0
+  thus ?case by (simp add: assms(1))
+next
+  case (Suc i)
+  have "B \<in> nxts_rlin2_set P {A} (x@(y^^i))"
+    using Suc.IH[OF assms] .
+  with assms(2) have "B \<in> nxts_rlin2_set P {A} (x@(y^^i)@y)"
+    using nxts_trans2[of B P A "x@(y^^i)" B y] by simp
+  thus ?case
+    by (simp add: pow_list_comm)
+qed
+
+text
+\<open>Combining the previous lemmas we can prove the pumping lemma where the starting non-terminal is in the production set. We simply extend the
+ lemma for non-terminals that are not part of the production set, as these non-terminals will produce the empty language\<close>
+
+lemma pumping_re_aux:
+  assumes "finite P"
+      and "A \<in> Nts P"
+      and "m = card (Nts P)"
+      and "accepted P A w"
+      and "length w \<ge> m"
+    shows "\<exists>x y z. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> m \<and> (\<forall>i. accepted P A (x@(y^^i)@z))"
+proof -
+  from assms(4) obtain Z where Z_in:"Z \<in> nxts_rlin2_set P {A} w" and Z_eps:"(Z,[])\<in>P"
+    by (auto simp: accepted_def)
+  obtain x y z C where *: "w = x@y@z" and **: "length y \<ge> 1" and ***: "length (x@y) \<le> m" and
+              1: "C \<in> nxts_rlin2_set P {A} x" and 2: "C \<in> nxts_rlin2_set P {C} y" and 3: "Z \<in> nxts_rlin2_set P {C} z"
+    using nxts_split_cycle[OF assms(1) assms(2) assms(3) Z_in assms(5)] by auto
+  have "\<forall>i. C \<in> nxts_rlin2_set P {A} (x@(y^^i))"
+    using pump_cycle[OF 1 2] by simp
+  with 3 have "\<forall>i. Z \<in> nxts_rlin2_set P {A} (x@(y^^i)@z)"
+    using nxts_trans2[of C P A] by fastforce
+  with Z_eps have ****: "(\<forall>i. accepted P A (x@(y^^i)@z))"
+    by (auto simp: accepted_def)
+  from * ** *** **** show ?thesis by auto
+qed
+
+theorem pumping_lemma_re_nts:
+  assumes "rlin2 P"
+      and "finite P"
+      and "A \<in> Nts P"
+  shows "\<exists>n. \<forall>w \<in> Lang P A. length w \<ge> n \<longrightarrow>
+          (\<exists>x y z. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> n \<and> (\<forall>i. x@(y^^i)@z \<in> Lang P A))" 
+  using assms pumping_re_aux[of P A "card (Nts P)"] Lang_iff_accepted_if_rlin2[OF assms(1)] by metis
+
+theorem pumping_lemma_regular:
+  assumes "rlin2 P" and "finite P"
+  shows "\<exists>n. \<forall>w \<in> Lang P A. length w \<ge> n \<longrightarrow>
+          (\<exists>x y z. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> n \<and> (\<forall>i. x@(y^^i)@z \<in> Lang P A))" 
+proof (cases "A \<in> Nts P")
+  case True
+  thus ?thesis
+    using pumping_lemma_re_nts[OF assms True] by simp
+next
+  case False
+  hence "Lang P A = {}"
+    by (auto intro!: Lang_empty_if_notin_Lhss simp add: Lhss_def Nts_def)
+  thus ?thesis by simp
+qed
+
+text \<open>Most of the time pumping lemma is used in the contrapositive form
+to prove that no right-linear set of productions exists.\<close>
+
+corollary pumping_lemma_regular_contr:
+  assumes "finite P"
+      and "\<forall>n. \<exists>w \<in> Lang P A. length w \<ge> n \<and> (\<forall>x y z. w = x@y@z \<and> length y \<ge> 1 \<and> length (x@y) \<le> n \<longrightarrow> (\<exists>i. x@(y^^i)@z \<notin> Lang P A))" 
+    shows "\<not>rlin2 P"
+using assms pumping_lemma_regular[of P A] by metis
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/ROOT
--- /dev/null
+++ thys/Context_Free_Grammar/ROOT
@@ -0,0 +1,22 @@
+chapter AFP
+
+session Context_Free_Grammar = Finite_Automata_HF +
+  options [timeout=300]
+  sessions
+    "HOL-Library"
+    List_Power
+  theories
+    Context_Free_Grammar
+    Parse_Tree
+    Context_Free_Language
+    Chomsky_Normal_Form
+    Pumping_Lemma_CFG
+    AnBnCn_not_CFL
+    CFL_Not_Intersection_Closed
+    Right_Linear
+    Right_Linear_Automata
+    Pumping_Lemma_Regular
+    AnBn_Not_Regular
+  document_files
+    "root.tex"
+    "root.bib"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Renaming_CFG.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Renaming_CFG.thy
@@ -0,0 +1,168 @@
+(*
+Authors: Markus Gschoßmann, Tobias Nipkow
+*)
+
+section \<open>Renaming Nonterminals\<close>
+
+theory Renaming_CFG
+imports Context_Free_Grammar
+begin
+
+text \<open>This theory provides lemmas that relate derivations w.r.t. some set of productions \<open>P\<close>
+to derivations w.r.t. a renaming of the nonterminals in \<open>P\<close>.\<close>
+
+fun rename_sym :: "('old \<Rightarrow> 'new) \<Rightarrow> ('old,'t) sym \<Rightarrow> ('new,'t) sym" where
+"rename_sym f (Nt n) = Nt (f n)" |
+"rename_sym f (Tm t) = Tm t"
+
+abbreviation "rename_syms f \<equiv> map (rename_sym f)"
+
+fun rename_prod :: "('old \<Rightarrow> 'new) \<Rightarrow> ('old,'t) prod \<Rightarrow> ('new,'t) prod" where
+"rename_prod f (A,w) = (f A, rename_syms f w)"
+
+abbreviation "rename_Prods f P \<equiv> rename_prod f ` P"
+
+lemma rename_sym_o_Tm[simp]: "rename_sym f \<circ> Tm = Tm"
+by(rule ext) simp
+
+lemma Nt_notin_rename_syms_if_notin_range:
+  "x \<notin> range f \<Longrightarrow> Nt x \<notin> set (rename_syms f w)"
+by(auto elim!: rename_sym.elims[OF sym])
+
+lemma in_Nts_rename_Prods: "B \<in> Nts (rename_Prods f P) = (\<exists>A \<in> Nts P. f A = B)"
+unfolding Nts_def nts_syms_def by(force split: prod.splits elim!: rename_sym.elims[OF sym])
+
+lemma rename_preserves_deriven:
+   "P \<turnstile> \<alpha> \<Rightarrow>(n) \<beta> \<Longrightarrow> rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>(n) rename_syms f \<beta>"
+proof (induction rule: deriven_induct)
+  case 0
+  then show ?case by simp
+next
+  let ?F = "rename_syms f"
+  case (Suc n u A v w)
+  then have "(f A, rename_syms f w) \<in> rename_Prods f P"
+    by (metis image_eqI rename_prod.simps) 
+  from derive.intros[OF this] have "rename_Prods f P \<turnstile> ?F u @ ?F [Nt A] @ ?F v \<Rightarrow> ?F u @ ?F w @ ?F v"
+    by auto
+  with Suc show ?case
+    by (simp add: relpowp_Suc_I)
+qed
+
+lemma rename_preserves_derives:
+   "P \<turnstile> \<alpha> \<Rightarrow>* \<beta> \<Longrightarrow> rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>* rename_syms f \<beta>"
+by (meson rename_preserves_deriven rtranclp_power)
+
+lemma rename_preserves_derivel:
+  assumes "P \<turnstile> \<alpha> \<Rightarrow>l \<beta>"
+  shows "rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>l rename_syms f \<beta>"
+proof -
+  from assms obtain A w u1 u2
+    where A_w_u1_u2: "(A,w) \<in> P \<and> \<alpha> = map Tm u1 @ Nt A # u2 \<and> \<beta> = map Tm u1 @ w @ u2"
+    unfolding derivel_iff by fast
+  then have "(f A, rename_syms f w) \<in> (rename_Prods f P) \<and> 
+             rename_syms f \<alpha> = map Tm u1 @ Nt (f A) # rename_syms f u2 \<and>
+             rename_syms f \<beta> = map Tm u1 @ rename_syms f w @ rename_syms f u2"
+    by force
+  then have "\<exists> (A,w) \<in> rename_Prods f P.
+        \<exists>u1 u2. rename_syms f \<alpha> = map Tm u1 @ Nt A # u2 \<and> rename_syms f \<beta> = map Tm u1 @ w @ u2"
+    by blast
+  then show ?thesis by (simp only: derivel_iff)
+qed
+
+lemma rename_preserves_deriveln:
+  "P \<turnstile> \<alpha> \<Rightarrow>l(n) \<beta> \<Longrightarrow> rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>l(n) rename_syms f \<beta>"
+proof (induction n arbitrary: \<beta>)
+  case 0
+  then show ?case by simp
+next
+  case Suc then show ?case
+    by (metis relpowp_Suc_E relpowp_Suc_I rename_preserves_derivel)
+qed
+
+lemma rename_preserves_derivels:
+  "P \<turnstile> \<alpha> \<Rightarrow>l* \<beta> \<Longrightarrow> rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>l* rename_syms f \<beta>"
+by (meson rename_preserves_deriveln rtranclp_power)
+
+lemma rename_deriven_iff_inj:
+fixes P :: "('a,'t)Prods"
+assumes "inj_on f (Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>)"
+shows "rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>(n) rename_syms f \<beta> \<longleftrightarrow> P \<turnstile> \<alpha> \<Rightarrow>(n) \<beta>" (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?r \<Longrightarrow> ?l" by (rule rename_preserves_deriven)
+next
+  (* since f is injective, the second direction follows from the first by using the inverse *)
+  let ?M = "Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>"
+  obtain "g" where "g = the_inv_into ?M f" and inv: "(\<And>x. x \<in> ?M \<Longrightarrow> (g (f x) = x))"
+    using assms by (simp add: the_inv_into_f_f inj_on_Un)
+  then have "s \<in> Syms P \<union> set \<alpha> \<union> set \<beta> \<Longrightarrow> rename_sym g (rename_sym f s) = s" for s::"('a,'t) sym"
+    by (cases s) (auto simp: Nts_def Syms_def nts_syms_def)
+  then have inv_rename_syms: "\<And>(ss::('a,'t) syms). set ss \<subseteq> Syms P \<union> set \<alpha> \<union> set \<beta> \<Longrightarrow> rename_syms g (rename_syms f ss) = ss"
+    by (simp add: list.map_ident_strong subset_iff)
+  with inv have "p \<in> P \<Longrightarrow> rename_prod g (rename_prod f p) = p" for p::"('a,'t) prod"
+    by(cases p)(auto simp: Nts_def Syms_def)
+  then have inv_rename_prods: "rename_Prods g (rename_Prods f P) = P"
+    using image_iff by fastforce
+  then show "?l \<Longrightarrow> ?r"
+    using rename_preserves_deriven inv_rename_syms by (metis Un_upper2 le_supI1)
+qed
+
+lemma rename_derives_iff_inj:
+  assumes "inj_on f (Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>)"
+  shows "rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>* rename_syms f \<beta> \<longleftrightarrow> P \<turnstile> \<alpha> \<Rightarrow>* \<beta>"
+by (meson assms relpowp_imp_rtranclp rename_deriven_iff_inj rtranclp_imp_relpowp)
+
+lemma rename_deriveln_iff_inj:
+fixes P :: "('a,'t)Prods"
+assumes "inj_on f (Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>)"
+shows "rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>l(n) rename_syms f \<beta> \<longleftrightarrow> P \<turnstile> \<alpha> \<Rightarrow>l(n) \<beta>" (is "?l \<longleftrightarrow> ?r")
+proof
+  show "?r \<Longrightarrow> ?l" by (rule rename_preserves_deriveln)
+next
+  let ?M = "Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>"
+  obtain "g" where "g = the_inv_into ?M f" and inv: "(\<And>x. x \<in> ?M \<Longrightarrow> (g (f x) = x))"
+    using assms by (simp add: the_inv_into_f_f inj_on_Un)
+  then have "s \<in> Syms P \<union> set \<alpha> \<union> set \<beta> \<Longrightarrow> rename_sym g (rename_sym f s) = s" for s::"('a,'t) sym"
+    by (cases s) (auto simp: Nts_def Syms_def nts_syms_def)
+  then have inv_rename_syms: "\<And>(ss::('a,'t) syms). set ss \<subseteq> Syms P \<union> set \<alpha> \<union> set \<beta> \<Longrightarrow> rename_syms g (rename_syms f ss) = ss"
+    by (simp add: list.map_ident_strong subset_iff)
+  with inv have "p \<in> P \<Longrightarrow> rename_prod g (rename_prod f p) = p" for p::"('a,'t) prod"
+    by(cases p)(auto simp: Nts_def Syms_def)
+  then have inv_rename_prods: "rename_Prods g (rename_Prods f P) = P"
+    using image_iff by fastforce
+  then show "?l \<Longrightarrow> ?r"
+    using rename_preserves_deriveln inv_rename_syms by (metis Un_upper2 le_supI1)
+qed
+
+lemma rename_derivels_iff_inj:
+  assumes "inj_on f (Nts P \<union> nts_syms \<alpha> \<union> nts_syms \<beta>)"
+  shows "rename_Prods f P \<turnstile> rename_syms f \<alpha> \<Rightarrow>l* rename_syms f \<beta> \<longleftrightarrow> P \<turnstile> \<alpha> \<Rightarrow>l* \<beta>"
+by (meson assms relpowp_imp_rtranclp rename_deriveln_iff_inj rtranclp_imp_relpowp)
+
+lemma Lang_rename_Prods:
+  assumes "inj_on f (Nts P \<union> {S})"
+  shows "Lang (rename_Prods f P) (f S) = Lang P S"
+proof -
+  from assms rename_derives_iff_inj[of f P "[Nt S]" "map Tm _"]
+  show ?thesis unfolding Lang_def by (simp)
+qed
+
+lemma derives_preserves_renaming:
+  assumes "rename_Prods f P \<turnstile> rename_syms f u \<Rightarrow>* fv"
+  shows "\<exists>v. fv = rename_syms f v"
+  using assms
+proof(induction rule: derives_induct)
+  case base
+  then show ?case by auto
+next
+  case (step u A v w)
+  then obtain A' where A'_src: "rename_syms f [Nt A'] = [Nt A]" by auto
+  from step obtain drvW where "rename_syms f drvW = u @ [Nt A] @ v" by auto
+  then have uAv_split: "u @ rename_syms f [Nt A'] @ v = rename_syms f drvW" using A'_src by simp
+  from uAv_split obtain u' where u'_src: "rename_syms f u' = u" by (metis map_eq_append_conv)
+  from uAv_split obtain v' where v'_src: "rename_syms f v' = v" by (metis map_eq_append_conv)
+  from step obtain w' where "rename_syms f w' = w" by auto
+  then have "u @ w @ v = rename_syms f (u' @ w' @ v')" using u'_src v'_src by auto
+  then show ?case by fast
+qed
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Right_Linear.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Right_Linear.thy
@@ -0,0 +1,689 @@
+(* Author: Kaan Taskin *)
+
+section \<open>Right-Linear Grammars\<close>
+
+theory Right_Linear
+imports Unit_Elimination Binarize
+begin
+
+text \<open>This theory defines (strongly) right-linear grammars and proves that every
+right-linear grammar can be transformed into a strongly right-linear grammar.\<close>
+
+text\<open>In a \emph{right linear} grammar every production has the form \<open>A \<rightarrow> wB\<close> or \<open>A \<rightarrow> w\<close>
+where \<open>w\<close> is a sequence of terminals:\<close>
+
+definition rlin :: "('n, 't) Prods \<Rightarrow> bool" where
+  "rlin ps = (\<forall>(A,w) \<in> ps. \<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B]))"
+
+definition rlin_noterm :: "('n, 't) Prods \<Rightarrow> bool" where
+  "rlin_noterm ps = (\<forall>(A,w) \<in> ps. w = [] \<or> (\<exists>u B. w = map Tm u @ [Nt B]))"
+
+definition rlin_bin :: "('n, 't) Prods \<Rightarrow> bool" where
+  "rlin_bin ps = (\<forall>(A,w) \<in> ps. w = [] \<or> (\<exists>B. w = [Nt B] \<or> (\<exists>a. w = [Tm a, Nt B])))"
+
+text\<open>In a \emph{strongly right linear} grammar every production has the form \<open>A \<rightarrow> aB\<close> or \<open>A \<rightarrow> \<epsilon>\<close>
+where \<open>a\<close> is a terminal:\<close>
+
+definition rlin2 :: "('a, 't) Prods \<Rightarrow> bool" where
+  "rlin2 ps = (\<forall>(A,w) \<in> ps. w = [] \<or> (\<exists>a B. w = [Tm a, Nt B]))"
+
+text\<open>A straightforward property:\<close>
+
+lemma rlin_if_rlin2: 
+  assumes "rlin2 ps"
+  shows "rlin ps"
+proof -
+  have "\<exists>u. x2 = map Tm u \<or> (\<exists>B. x2 = map Tm u @ [Nt B])"
+    if "\<forall>x\<in>ps. \<forall>x1 x2. x = (x1, x2) \<longrightarrow> x2 = [] \<or> (\<exists>a B. x2 = [Tm a, Nt B])"
+      and "(x1, x2) \<in> ps"
+    for x1 :: 'a and x2 :: "('a, 'b) sym list"
+    using that by (metis append_Cons append_Nil list.simps(8,9))
+  with assms show ?thesis
+    unfolding rlin_def rlin2_def
+    by (auto split: prod.splits)
+qed
+
+lemma rlin_cases:
+  assumes rlin_ps: "rlin ps" 
+      and elem: "(A,w) \<in> ps"
+    shows "(\<exists>B. w = [Nt B]) \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))"
+proof -
+  from rlin_ps have "\<forall>(A,w) \<in> ps. (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u \<le> 0)) 
+                   \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))"
+    using rlin_def by fastforce
+  with elem have "(\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u \<le> 0)) 
+                   \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))" by auto
+  hence "(\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u = 0)) 
+                   \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))" by simp
+  hence "(\<exists>u. w = map Tm u \<or> (\<exists>B. w = [Nt B])) 
+                   \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))" by auto
+  hence "(\<exists>B. w = [Nt B]) \<or> (\<exists>u. w = map Tm u \<or> (\<exists>B. w = map Tm u @ [Nt B] \<and> length u > 0))" by blast
+  thus ?thesis .
+qed
+
+
+subsection \<open>From \<open>rlin\<close> to \<open>rlin2\<close>\<close>
+
+text
+\<open>The \<open>finalize\<close> function is responsible of the transformation of a production list from \<open>rlin\<close> to \<open>rlin_noterm\<close>, while 
+ preserving the language. We make use of fixpoint iteration and define the function \<open>finalize1\<close> that adds a 
+ fresh non-terminal \<open>B\<close> at the end of every production that consists only of terminals and has at least length one. The
+ function also introduces the new production \<open>(B,[])\<close> in the production list. The step function of the fixpoint iteration is
+ then the auxiliary function \<open>finalize'\<close>. We also define the count function as \<open>countfin\<close> which counts the the productions that
+ consists only of terminal and has at least length one\<close>
+
+fun finalize1 :: "('n :: infinite, 't) prods \<Rightarrow> ('n, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "finalize1 ps' [] = []"
+| "finalize1 ps' ((A,[])#ps) = (A,[]) # finalize1 ps' ps"
+| "finalize1 ps' ((A,w)#ps) = 
+    (if \<exists>u. w = map Tm u then let B = fresh(nts ps') in (A,w @ [Nt B])#(B,[])#ps else (A,w) # finalize1 ps' ps)"
+
+definition finalize' :: "('n::infinite, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "finalize' ps = finalize1 ps ps"
+
+fun countfin :: "('n::infinite, 't) prods \<Rightarrow> nat" where
+  "countfin [] = 0"
+| "countfin ((A,[])#ps) = countfin ps"
+| "countfin ((A,w) # ps) = (if \<exists>u. w = map Tm u then 1 + countfin ps else countfin ps)"
+
+definition finalize :: "('n::infinite, 't) prods \<Rightarrow> ('n, 't) prods" where
+  "finalize ps = (finalize' ^^ (countfin ps)) ps"
+
+text
+\<open>Firstly we show that \<open>finalize\<close> indeed does the intended transformation\<close>
+
+lemma countfin_dec1:
+  assumes "finalize1 ps' ps \<noteq> ps" 
+  shows "countfin ps > countfin (finalize1 ps' ps)"
+using assms proof (induction ps' ps rule: finalize1.induct)
+  case (3 ps' A v va ps)
+  thus ?case proof (cases "\<exists>u. v # va = map Tm u")
+    case True
+    let ?B = "fresh(nts ps')"
+    have not_tm: "\<nexists>u. v # va @ [Nt ?B] = map Tm u"
+      by (simp add: ex_map_conv)
+    from True have "countfin (finalize1 ps' ((A, v # va) # ps)) = countfin ((A,v#va @ [Nt ?B])#(?B,[])#ps)"
+      by (metis append_Cons finalize1.simps(3))
+    also from not_tm have "... = countfin ps" by simp
+    also have "... < countfin ps + 1" by simp
+    also from True have "... = countfin ((A, v # va) # ps)" by simp
+    finally show ?thesis .
+  next
+    case False
+    with 3 show ?thesis by simp
+  qed
+qed simp_all
+
+lemma countfin_dec':
+  assumes "finalize' ps \<noteq> ps" 
+  shows "countfin ps > countfin (finalize' ps)"
+  using finalize'_def assms countfin_dec1 by metis
+
+lemma finalize_ffpi:
+  "finalize'((finalize' ^^ countfin x) x) = (finalize' ^^ countfin x) x"
+proof -
+  have "\<And>x. countfin(finalize' x) < countfin x \<or> finalize' x = x"
+    using countfin_dec' by blast
+  thus ?thesis using funpow_fix by metis
+qed
+
+lemma finalize_rlinnoterm1:
+  assumes "rlin (set ps)"
+      and "ps = finalize1 ps' ps"
+    shows "rlin_noterm (set ps)"
+  using assms proof (induction ps' ps rule: finalize1.induct)
+  case (1 ps')
+  thus ?case
+    by (simp add: rlin_noterm_def)
+next
+  case (2 ps' A ps)
+  thus ?case
+    by (simp add: rlin_def rlin_noterm_def)
+next
+  case (3 ps' A v va ps)
+  thus ?case proof (cases "\<exists>u. v#va = map Tm u")
+    case True
+    with 3 show ?thesis 
+      by simp (meson list.inject not_Cons_self)
+  next
+    case False
+    with 3 show ?thesis
+      by (simp add: rlin_def rlin_noterm_def)
+  qed
+qed
+
+lemma finalize_rlin1:
+  "rlin (set ps) \<Longrightarrow> rlin (set (finalize1 ps' ps))"
+proof (induction ps' ps rule: finalize1.induct)
+  case (2 ps' A ps)
+  thus ?case
+    by (simp add: rlin_def)
+next
+  case (3 ps' A v va ps)
+  thus ?case proof (cases "\<exists>u. v#va = map Tm u")
+    case True
+    with 3 show ?thesis
+      by (auto simp: Let_def rlin_def split_beta map_eq_append_conv Cons_eq_append_conv intro: exI[of _ "_ # _"])
+  next
+    case False
+    with 3 show ?thesis
+      by (simp add: rlin_def)
+  qed
+qed simp
+
+lemma finalize_rlin':
+  "rlin (set ps) \<Longrightarrow> rlin (set (finalize' ps))"
+  unfolding finalize'_def using finalize_rlin1 by blast
+
+lemma finalize_rlin'n:
+  "rlin (set ps) \<Longrightarrow> rlin (set ((finalize'^^n) ps))"
+  by (induction n) (auto simp add: finalize_rlin')
+
+lemma finalize_rlinnoterm':
+  assumes "rlin (set ps)"
+      and "ps = finalize' ps"
+  shows "rlin_noterm (set ps)"
+  using finalize'_def assms finalize_rlinnoterm1 by metis
+
+lemma finalize_rlinnoterm: 
+  "rlin (set ps) \<Longrightarrow> rlin_noterm (set (finalize ps))"
+proof -
+  assume asm: "rlin (set ps)"
+  hence 1: "rlin (set ((finalize' ^^ countfin ps) ps))"
+    using finalize_rlin'n by auto
+  have "finalize'((finalize' ^^ countfin ps) ps) = (finalize' ^^ countfin ps) ps"
+    using finalize_ffpi by blast
+  with 1 have "rlin_noterm (set ((finalize' ^^ countfin ps) ps))"
+    using finalize_rlinnoterm' by metis
+  hence "rlin_noterm (set (finalize ps))"
+    by (simp add: finalize_def)
+  thus ?thesis .
+qed
+
+text
+\<open>Now proving the language preservation property of \<open>finalize\<close>, similarly to \<open>binarize\<close>\<close>
+
+lemma finalize1_cases:
+  "finalize1 ps' ps = ps \<or> (\<exists>A w ps'' B. set ps = {(A, w)} \<union> set ps'' \<and> set (finalize1 ps' ps) = {(A,w @ [Nt B]),(B,[])} \<union> set ps'' \<and> Nt B \<notin> syms ps')"
+proof (induction ps' ps rule: finalize1.induct)
+  case (2 ps' C ps)
+  thus ?case proof (cases "finalize1 ps' ps = ps")
+    case False
+    then obtain A w ps'' B where defs: "set ps = {(A, w)} \<union> set ps'' \<and> set (finalize1 ps' ps) = {(A, w @ [Nt B]), (B, [])} \<union> set ps'' \<and> Nt B \<notin> syms ps'"
+    using 2 by blast
+    from defs have wit: "set ((C, []) # ps) = {(A, w)} \<union> set ((C, []) # ps'')" by simp
+    from defs have wit2: "set (finalize1 ps' ((C, []) # ps)) = {(A, w @ [Nt B]), (B, [])} \<union> set ((C, []) # ps'')" by simp
+    from defs have wit3: "Nt B \<notin> syms ps'" by simp
+    from wit wit2 wit3 show ?thesis by blast
+  qed simp
+next
+  case (3 ps' C v va ps)
+  thus ?case proof (cases "\<exists>u. v#va = map Tm u")
+    case True
+    thus ?thesis using fresh_nts in_Nts_iff_in_Syms
+      by (simp add: Let_def) fastforce
+  next
+    case false1: False
+    thus ?thesis proof (cases "finalize1 ps' ps = ps")
+    case False
+    with false1 obtain A w ps'' B where defs: "set ps = {(A, w)} \<union> set ps'' \<and> set (finalize1 ps' ps) = {(A, w @ [Nt B]), (B, [])} \<union> set ps'' \<and> Nt B \<notin> syms ps'"
+    using 3 by blast
+    from defs have wit: "set ((C, v#va) # ps) = {(A, w)} \<union> set ((C, v#va) # ps'')" by simp
+    from defs false1 have wit2: "set (finalize1 ps' ((C, v#va) # ps)) = {(A, w @ [Nt B]), (B, [])} \<union> set ((C, v#va) # ps'')" by simp
+    from defs have wit3: "Nt B \<notin> syms ps'" by simp
+    from wit wit2 wit3 show ?thesis by blast
+  qed simp
+  qed
+qed simp
+
+lemma finalize_der':
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (finalize' ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+  unfolding finalize'_def proof (cases "finalize1 ps ps = ps")
+  case False
+  then obtain C w ps'' B where defs: "set ps = {(C, w)} \<union> set ps'' \<and> set (finalize1 ps ps) = {(C, w @ [Nt B]), (B, [])} \<union> set ps'' \<and> Nt B \<notin> syms ps"
+    by (meson finalize1_cases)
+  from defs have a_not_b: "C \<noteq> B" unfolding Syms_def by fast
+  from defs assms have a1: "A \<noteq> B" unfolding Lhss_def Syms_def by auto
+  from defs have a2: "Nt B \<notin> set (map Tm x)" by auto
+  from defs have a3: "Nt B \<notin> set []" by simp
+  from defs have "set ps = set ((C, w) # ps'')" by simp
+  with defs a_not_b have a4: "B \<notin> lhss ((C, w @ [Nt B]) # ps'')"
+    unfolding Lhss_def Syms_def by auto
+  from defs have notB: "Nt B \<notin> syms ps''" unfolding Syms_def by blast
+  then have 1: "set ps = set (substP (Nt B) [] ((C, w @ [Nt B]) # ps''))" proof -
+    from defs have s1: "Nt B \<notin> syms ps" unfolding Syms_def by meson
+    from defs have s2: "(C,w) \<in> set ps" by blast
+    from s1 s2 have b_notin_w: "Nt B \<notin> set w" unfolding Syms_def by fastforce
+    from defs have "set ps = {(C, w)} \<union> set ps''" by simp
+    also have "... = set ((C, w) # ps'')" by simp
+    also have "... = set ([(C, w)] @ ps'')" by simp
+    also from defs have "... = set ([(C,substsNt B [] (w @ [Nt B]))] @ ps'')" using b_notin_w
+      by (simp add: substs_skip)
+    also have "... = set ((substP (Nt B) [] [(C, w @ [Nt B])]) @ ps'')" by (simp add: substP_def)
+    also have "... = set ((substP (Nt B) [] [(C, w @ [Nt B])]) @ substP (Nt B) [] ps'')" using notB by (simp add: substP_skip2)
+    also have "... = set (substP (Nt B) [] ((C, w @ [Nt B]) # ps''))" by (simp add: substP_def)
+    finally show ?thesis .
+  qed
+  from defs have 2: "set (finalize1 ps ps) = set ((C, w @ [Nt B]) # (B, []) # ps'')" by auto
+  with 1 2 a1 a2 a3 a4 show "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (finalize1 ps ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+    by (simp add: derives_inlining insert_commute)
+qed simp
+
+lemma lhss_finalize1:
+  "lhss ps \<subseteq> lhss (finalize1 ps' ps)"
+proof (induction ps' ps rule: finalize1.induct)
+  case (2 ps' A ps)
+  thus ?case unfolding Lhss_def by auto
+next
+  case (3 ps' A v va ps)
+  thus ?case unfolding Lhss_def by (auto simp: Let_def)
+qed simp
+
+lemma lhss_binarize'n:
+  "lhss ps \<subseteq> lhss ((finalize'^^n) ps)"
+proof (induction n)
+  case (Suc n)
+  thus ?case unfolding finalize'_def using lhss_finalize1 by auto
+qed simp
+
+lemma finalize_der'n: 
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set ((finalize'^^n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+using assms proof (induction n)
+  case (Suc n)
+  hence "A \<in> lhss ((finalize' ^^ n) ps)"
+    using lhss_binarize'n by blast
+  hence "set ((finalize' ^^ n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (finalize' ((finalize' ^^ n) ps)) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+    using finalize_der' by blast
+  hence"set ((finalize' ^^ n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set ((finalize' ^^ Suc n) ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+    by simp
+  with Suc show ?case by blast
+qed simp
+
+lemma finalize_der: 
+  assumes "A \<in> lhss ps"
+  shows "set ps \<turnstile> [Nt A] \<Rightarrow>* map Tm x \<longleftrightarrow> set (finalize ps) \<turnstile> [Nt A] \<Rightarrow>* map Tm x"
+  unfolding finalize_def using finalize_der'n[OF assms] by simp
+
+lemma lang_finalize_lhss:
+  assumes "A \<in> lhss ps"
+  shows "lang ps A = lang (finalize ps) A"
+  using finalize_der[OF assms] Lang_eqI_derives by metis
+
+lemma finalize_syms1:
+  assumes  "Nt A \<in> syms ps"
+    shows  "Nt A \<in> syms (finalize1 ps' ps)"
+using assms proof (induction ps' ps rule: finalize1.induct)
+  case (3 ps' A v va ps)
+  thus ?case proof (cases "\<exists>u. v#va = map Tm u")
+    case True
+    with 3 show ?thesis unfolding Syms_def by (auto simp: Let_def)
+  next
+    case False
+    with 3 show ?thesis unfolding Syms_def by auto
+  qed
+qed auto
+
+lemma finalize_nts'n:
+  assumes "A \<in> nts ps"
+  shows   "A \<in> nts ((finalize' ^^ n) ps)"
+  using assms proof (induction n)
+  case (Suc n)
+  thus ?case
+    unfolding finalize'_def by (simp add: finalize_syms1 in_Nts_iff_in_Syms)
+qed simp
+
+lemma finalize_nts:
+  assumes "A \<in> nts ps"
+  shows   "A \<in> nts (finalize ps)"
+  unfolding finalize_def using finalize_nts'n[OF assms] by simp
+
+lemma finalize_lhss_nts1:
+  assumes "A \<notin> lhss ps"
+      and "A \<in> nts ps'"
+    shows "A \<notin> lhss (finalize1 ps' ps)"
+using assms proof (induction ps' ps rule: finalize1.induct)
+  case (3 ps' A v va ps)
+  thus ?case proof (cases "\<exists>u. v#va = map Tm u")
+    case True
+    with 3 show ?thesis unfolding Lhss_def by (auto simp: Let_def fresh_nts)
+  next
+    case False
+    with 3 show ?thesis unfolding Lhss_def by (auto simp: Let_def)
+  qed
+qed simp_all
+
+lemma finalize_lhss_nts'n:
+  assumes "A \<notin> lhss ps"
+      and "A \<in> nts ps"
+    shows "A \<notin> lhss ((finalize'^^n) ps) \<and> A \<in> nts ((finalize'^^n) ps)"
+  using assms proof (induction n)
+  case (Suc n)
+  thus ?case
+    unfolding finalize'_def by (simp add: finalize_lhss_nts1 finalize_syms1 in_Nts_iff_in_Syms)
+qed simp
+
+lemma finalize_lhss_nts:
+   assumes "A \<notin> lhss ps"
+      and  "A \<in> nts ps"
+    shows "A \<notin> lhss (finalize ps) \<and> A \<in> nts (finalize ps)"
+  unfolding finalize_def using finalize_lhss_nts'n[OF assms] by simp
+
+lemma lang_finalize: 
+  assumes "A \<in> nts ps"
+  shows "lang (finalize ps) A = lang ps A"
+proof (cases "A \<in> lhss ps")
+  case True
+  thus ?thesis
+    using lang_finalize_lhss by blast
+next
+  case False
+  thus ?thesis
+    using assms finalize_lhss_nts Lang_empty_if_notin_Lhss by fast
+qed
+
+text
+\<open>Next step is to define the transformation from \<open>rlin_noterm\<close> to \<open>rlin_bin\<close>. For this we use the function \<open>binarize\<close>.
+ The language preservation property of \<open>binarize\<close> is already proven\<close>
+
+lemma binarize_rlinbin1: 
+  assumes "rlin_noterm (set ps)"
+      and "ps = binarize1 ps' ps"
+  shows "rlin_bin (set ps)"
+  using assms proof (induction ps' ps rule: binarize1.induct)
+  case (1 ps')
+  thus ?case
+    by (simp add: rlin_bin_def)
+next
+  case (2 ps' A ps)
+  thus ?case
+    by (simp add: rlin_noterm_def rlin_bin_def)
+next
+  case (3 ps' A s0 u ps)
+  from "3.prems"(2) have a1: "length u \<le> 1" by simp (meson list.inject not_Cons_self)
+  with "3.prems"(2) have a2: "ps = binarize1 ps' ps" by simp
+  from "3.prems"(1) have a3: "rlin_noterm (set ps)"
+    by (simp add: rlin_noterm_def)
+  from a1 a2 a3 have 1: "rlin_bin (set ps)"
+    using "3.IH" by blast
+  from "3.prems"(1) have ex: "\<exists>v B. s0 # u = map Tm v @ [Nt B]"
+    by (simp add: rlin_noterm_def)
+  with a1 have 2: "\<exists>B. s0 # u = [Nt B] \<or> (\<exists>a. s0 # u = [Tm a, Nt B])" proof (cases "length u = 0")
+    case True
+    with ex show ?thesis by simp
+  next
+    case False
+    with a1 have "length u = 1" by linarith
+    show ?thesis
+    proof -
+      have "\<exists>B. s0 = Nt B \<and> u = [] \<or> (\<exists>a. s0 = Tm a) \<and> u = [Nt B]"
+        if "length u = Suc 0" and "s0 # u = map Tm v @ [Nt B]"
+        for v :: "'b list" and B :: 'a
+        using that by (metis append_Cons append_Nil append_butlast_last_id butlast_snoc diff_Suc_1 hd_map last_snoc length_0_conv length_butlast list.sel(1) list.simps(8))
+      with ex \<open>length u = 1\<close> show ?thesis
+        by auto
+    qed
+  qed
+  from 1 2 show ?case
+    by (simp add: rlin_bin_def)
+qed
+
+lemma binarize_noterm1:
+  "rlin_noterm (set ps) \<Longrightarrow> rlin_noterm (set (binarize1 ps' ps))"
+proof (induction ps' ps rule: binarize1.induct)
+  case (2 ps' A ps)
+  thus ?case
+    by (simp add: rlin_noterm_def)
+next
+  case (3 ps' A s0 u ps)
+  thus ?case proof (cases "length u \<le> 1")
+    case True
+    with 3 show ?thesis
+      by (simp add: rlin_noterm_def)
+  next
+    case False
+    let ?B = "fresh(nts ps')"
+    from "3.prems" have a1: "rlin_noterm (set ps)"
+      by (simp add: rlin_noterm_def)
+    from "3.prems" have ex: "\<exists>v B. s0 # u = map Tm v @ [Nt B]"
+      by (simp add: rlin_noterm_def)
+    with False have a2: "\<exists>v B. [s0, Nt ?B] = map Tm v @ [Nt B]"
+      by (auto simp: Cons_eq_append_conv neq_Nil_conv intro: exI[of _ "[_]"])
+    from ex False have a3: "\<exists>v B. u = map Tm v @ [Nt B]"
+      by (auto simp: Cons_eq_append_conv)
+    from False a1 a2 a3 show ?thesis
+      by (auto simp: Let_def rlin_noterm_def)
+  qed
+qed simp
+
+lemma binarize_noterm':
+  "rlin_noterm (set ps) \<Longrightarrow> rlin_noterm (set (binarize' ps))"
+  unfolding binarize'_def using binarize_noterm1 by blast
+
+lemma binarize_noterm'n:
+  "rlin_noterm (set ps) \<Longrightarrow> rlin_noterm (set ((binarize'^^n) ps))"
+  by (induction n) (auto simp add: binarize_noterm')
+
+lemma binarize_rlinbin':
+  assumes "rlin_noterm (set ps)"
+      and "ps = binarize' ps"
+  shows "rlin_bin (set ps)"
+  using binarize'_def assms binarize_rlinbin1 by metis
+
+lemma binarize_rlinbin: 
+  "rlin_noterm (set ps) \<Longrightarrow> rlin_bin (set (binarize ps))"
+proof -
+  assume asm: "rlin_noterm (set ps)"
+  hence 1: "rlin_noterm (set ((binarize' ^^ count ps) ps))"
+    using binarize_noterm'n by auto
+  have "binarize'((binarize' ^^ count ps) ps) = (binarize' ^^ count ps) ps"
+    using binarize_ffpi by blast
+  with 1 have "rlin_bin (set ((binarize' ^^ count ps) ps))"
+    using binarize_rlinbin' by fastforce
+  hence "rlin_bin (set (binarize ps))"
+    by (simp add: binarize_def)
+  thus ?thesis .
+qed
+
+text
+\<open>The last transformation takes a production set from \<open>rlin_bin\<close> and converts it to \<open>rlin2\<close>. That is, we need to remove unit
+ productions of the from \<open>(A, [Nt B])\<close>. In \<open>uProds.thy\<close> is the predicate \<open>\<U> ps' ps\<close> defined that is satisfied if \<open>ps\<close> is the
+ same production set as \<open>ps'\<close> without the unit productions. The language preservation property is already given\<close>
+
+lemma uppr_rlin2:
+  assumes rlinbin: "rlin_bin (set ps')"
+    and uppr_ps': "unit_elim_rel ps' ps"
+  shows "rlin2 (set ps)"
+proof - 
+  from rlinbin have "rlin2 (set ps' - {(A,w) \<in> set ps'. \<exists>B. w = [Nt B]})"
+    using rlin2_def rlin_bin_def by fastforce
+  hence "rlin2 (set ps' - (unit_prods ps'))"
+    by (simp add: unit_prods_def)
+  hence 1: "rlin2 (unit_rm ps')"
+    by (simp add: unit_rm_def)
+  hence 2: "rlin2 (new_prods ps')"
+    unfolding new_prods_def rlin2_def by fastforce
+  from 1 2 have "rlin2 (unit_rm ps' \<union> new_prods ps')"
+    unfolding rlin2_def by auto
+  with uppr_ps' have "rlin2 (set ps)"
+    by (simp add: unit_elim_rel_def)
+  thus ?thesis .
+qed
+
+text
+\<open>The transformation \<open>rlin2_of_rlin\<close> applies the presented functions in the right order. At the end, we show that \<open>rlin2_of_rlin\<close>
+ converts a production set from \<open>rlin\<close> to a production set from \<open>rlin2\<close>, without changing the language\<close>
+
+definition rlin2_of_rlin :: "('n::infinite,'t) prods \<Rightarrow> ('n,'t)prods" where
+  "rlin2_of_rlin ps = unit_elim (binarize (finalize ps))"
+
+theorem rlin_to_rlin2: 
+  assumes "rlin (set ps)" 
+  shows "rlin2 (set (rlin2_of_rlin ps))"
+using assms proof -
+  assume "rlin (set ps)"
+  hence "rlin_noterm (set (finalize ps))"
+    using finalize_rlinnoterm by blast
+  hence "rlin_bin (set (binarize (finalize ps)))"
+    by (simp add: binarize_rlinbin)
+  hence "rlin2 (set (unit_elim (binarize (finalize ps))))"
+    by (simp add: unit_elim_rel_unit_elim uppr_rlin2)
+  thus "rlin2 (set (rlin2_of_rlin ps))"
+    by (simp add: rlin2_of_rlin_def)
+qed
+
+lemma lang_rlin2_of_rlin:
+  "A \<in> Nts (set ps) \<Longrightarrow> lang (rlin2_of_rlin ps) A = lang ps A"
+by(simp add: rlin2_of_rlin_def lang_unit_elim finalize_nts lang_binarize lang_finalize)
+
+
+subsection \<open>Properties of \<open>rlin2\<close> derivations\<close>
+
+text
+\<open>In the following we present some properties for list of symbols that are derived from a production set satisfying \<open>rlin2\<close>\<close>
+
+lemma map_Tm_single_nt:
+  assumes "map Tm w @ [Tm a, Nt A] = u1 @ [Nt B] @ u2"
+  shows "u1 = map Tm (w @ [a]) \<and> u2 = []"
+proof -
+  from assms have *: "map Tm (w @ [a]) @ [Nt A] = u1 @ [Nt B] @ u2" by simp
+  have 1: "Nt B \<notin> set (map Tm (w @ [a]))" by auto
+  have 2: "Nt B \<in> set (u1 @ [Nt B] @ u2)" by simp
+  from * 1 2 have "Nt B \<in> set ([Nt A])"
+    by (metis list.set_intros(1) rotate1.simps(2) set_ConsD set_rotate1 sym.inject(1))
+  hence "[Nt B] = [Nt A]" by simp
+  with 1 * show ?thesis
+    by (metis append_Cons append_Cons_eq_iff append_self_conv emptyE empty_set)
+qed
+
+text
+\<open>A non-terminal can only occur as the rightmost symbol\<close>
+
+lemma rlin2_derive:
+  assumes "P \<turnstile> v1 \<Rightarrow>* v2" 
+      and "v1 = [Nt A]"
+      and "v2 = u1 @ [Nt B] @ u2" 
+      and "rlin2 P"
+    shows "\<exists>w. u1 = map Tm w \<and> u2 = []"
+using assms proof (induction arbitrary: u1 B u2 rule: derives_induct)
+  case base
+  then show ?case
+    by (simp add: append_eq_Cons_conv)
+next
+  case (step u C v w)
+  from step.prems(1) step.prems(3) have "\<exists>w. u = map Tm w \<and> v = []" 
+    using step.IH[of u C v] by simp
+  then obtain wh where u_def: "u = map Tm wh" by blast
+  have v_eps: "v = []"
+    using \<open>\<exists>w. u = map Tm w \<and> v = []\<close> by simp
+  from step.hyps(2) step.prems(3) have w_cases: "w = [] \<or> (\<exists>d D. w = [Tm d, Nt D])"
+    unfolding rlin2_def by auto
+  then show ?case proof cases
+    assume "w=[]"
+    with v_eps step.prems(2) have "u = u1 @ [Nt B] @ u2" by simp
+    with u_def show ?thesis by (auto simp: append_eq_map_conv)
+  next
+    assume "\<not>w=[]"
+    then obtain d D where "w = [Tm d, Nt D]" 
+      using w_cases by blast
+    with u_def v_eps step.prems(2) have "u1 = map Tm (wh @ [d]) \<and> u2 = []"
+      using map_Tm_single_nt[of wh d D u1 B u2] by simp
+    thus ?thesis by blast
+  qed
+qed
+
+text
+\<open>A new terminal is introduced by a production of the form \<open>(C, [Tm x, Nt B])\<close>\<close>
+
+lemma rlin2_introduce_tm:
+  assumes "rlin2 P"
+      and "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Tm x, Nt B]"
+    shows "\<exists>C. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt C] \<and> (C,[Tm x, Nt B]) \<in> P"
+proof -
+  from assms(2) have "\<exists>v. P \<turnstile> [Nt A] \<Rightarrow>* v \<and> P \<turnstile> v \<Rightarrow> map Tm w @ [Tm x, Nt B]"
+    using rtranclp.cases by fastforce
+  then obtain v where v_star: "P \<turnstile> [Nt A] \<Rightarrow>* v" and v_step: "P \<turnstile> v \<Rightarrow> map Tm w @ [Tm x, Nt B]" by blast
+  from v_step have "\<exists>u1 u2 C \<alpha>. v = u1 @ [Nt C] @ u2 \<and> map Tm w @ [Tm x, Nt B] = u1 @ \<alpha> @ u2 \<and> (C,\<alpha>) \<in> P"
+    using derive.cases by fastforce
+  then obtain u1 u2 C \<alpha> where v_def: "v = u1 @ [Nt C] @ u2" and w_def: "map Tm w @ [Tm x, Nt B] = u1 @ \<alpha> @ u2" 
+                          and C_prod: "(C,\<alpha>) \<in> P" by blast
+  from assms(1) v_star v_def have u2_eps: "u2 = []"
+    using rlin2_derive[of P "[Nt A]"] by simp
+  from assms(1) v_star v_def obtain wa where u1_def: "u1 = map Tm wa"
+    using rlin2_derive[of P "[Nt A]" "u1 @ [Nt C] @ u2" A u1] by auto
+  from w_def u2_eps u1_def have "map Tm w @ [Tm x, Nt B] = map Tm wa @ \<alpha>" by simp
+  then have "map Tm (w @ [x]) @ [Nt B] = map Tm wa @ \<alpha>" by simp
+  then have "\<alpha> \<noteq> []" 
+    by (metis append.assoc append.right_neutral list.distinct(1) map_Tm_single_nt)
+  with assms(1) C_prod obtain d D where "\<alpha> = [Tm d, Nt D]"
+    using rlin2_def by fastforce
+  from w_def u2_eps have x_d: "x = d" 
+    using \<open>\<alpha> = [Tm d, Nt D]\<close> by simp
+  from w_def u2_eps have B_D: "B = D"
+    using \<open>\<alpha> = [Tm d, Nt D]\<close> by simp
+  from x_d B_D have alpha_def: "\<alpha> = [Tm x, Nt B]"
+    using \<open>\<alpha> = [Tm d, Nt D]\<close> by simp
+  from w_def u2_eps alpha_def have "map Tm w = u1" by simp
+  with u1_def have w_eq_wa: "w = wa"
+    by (metis list.inj_map_strong sym.inject(2))
+  from v_def u1_def w_eq_wa u2_eps have "v = map Tm w @ [Nt C]" by simp
+  with v_star have 1: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt C]" by simp
+  from C_prod alpha_def have 2: "(C,[Tm x, Nt B]) \<in> P" by simp
+  from 1 2 show ?thesis by auto
+qed
+
+lemma rlin2_nts_derive_eq: 
+  assumes "rlin2 P"
+      and "P \<turnstile> [Nt A] \<Rightarrow>* [Nt B]"
+    shows "A = B"
+proof -
+  from assms(2) have star_cases: "[Nt A] = [Nt B] \<or> (\<exists>w. P \<turnstile> [Nt A] \<Rightarrow> w \<and> P \<turnstile> w \<Rightarrow>* [Nt B])"
+    using converse_rtranclpE by force
+  show ?thesis proof cases
+    assume "\<not>[Nt A] = [Nt B]"
+    then obtain w where w_step: "P \<turnstile> [Nt A] \<Rightarrow> w" and w_star: "P \<turnstile> w \<Rightarrow>* [Nt B]"
+      using star_cases by auto
+    from assms(1) w_step have w_cases: "w = [] \<or> (\<exists>a C. w = [Tm a, Nt C])"
+      unfolding rlin2_def using derive_singleton[of P "Nt A" w] by auto
+    show ?thesis proof cases
+      assume "w = []"
+      with w_star show ?thesis by simp
+    next
+      assume "\<not>w = []"
+      with w_cases obtain a C where "w = [Tm a, Nt C]" by blast
+      with w_star show ?thesis
+        using derives_Tm_Cons[of P a "[Nt C]" "[Nt B]"] by simp
+    qed
+  qed simp
+qed
+
+text
+\<open>If the list of symbols consists only of terminals, the last production used is of the form \<open>B,[]\<close>\<close>
+
+lemma rlin2_tms_eps:
+  assumes "rlin2 P"
+      and "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w"
+    shows "\<exists>B. P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt B] \<and> (B,[]) \<in> P"
+proof -
+  from assms(2) have "\<exists>v. P \<turnstile> [Nt A] \<Rightarrow>* v \<and> P \<turnstile> v \<Rightarrow> map Tm w"
+    using rtranclp.cases by force
+  then obtain v where v_star: "P \<turnstile> [Nt A] \<Rightarrow>* v" and v_step: "P \<turnstile> v \<Rightarrow> map Tm w" by blast
+  from v_step have "\<exists>u1 u2 C \<alpha>. v = u1 @ [Nt C] @ u2 \<and> map Tm w = u1 @ \<alpha> @ u2 \<and> (C,\<alpha>) \<in> P"
+    using derive.cases by fastforce
+  then obtain u1 u2 C \<alpha> where v_def: "v = u1 @ [Nt C] @ u2" and w_def: "map Tm w = u1 @ \<alpha> @ u2" and C_prod: "(C,\<alpha>) \<in> P" by blast
+  have "\<nexists>A. Nt A \<in> set (map Tm w)" by auto
+  with w_def have "\<nexists>A. Nt A \<in> set \<alpha>" 
+    by (metis Un_iff set_append)
+  then have "\<nexists>a A. \<alpha> = [Tm a, Nt A]" by auto
+  with assms(1) C_prod have alpha_eps: "\<alpha> = []"
+    using rlin2_def by force
+  from v_star assms(1) v_def have u2_eps: "u2 = []"
+    using rlin2_derive[of P "[Nt A]"] by simp
+  from w_def alpha_eps u2_eps have u1_def: "u1 = map Tm w" by simp
+  from v_star v_def u1_def u2_eps have 1: "P \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt C]" by simp
+  from alpha_eps C_prod have 2: "(C,[]) \<in> P"  by simp
+  from 1 2 show ?thesis by auto
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Right_Linear_Automata.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Right_Linear_Automata.thy
@@ -0,0 +1,199 @@
+(*
+Authors: Kaan Taskin, Tobias Nipkow
+*)
+
+section \<open>Relating Strongly Right-Linear Grammars and Automata\<close>
+
+theory Right_Linear_Automata
+imports
+  NDA_rlin2
+  "Finite_Automata_HF.Finite_Automata_HF"
+  HereditarilyFinite.Finitary
+begin
+
+subsection \<open>From Strongly Right-Linear Grammar to NFA\<close>
+
+definition nfa_rlin2 :: "('n,'t)Prods \<Rightarrow> 'n \<Rightarrow> ('t,'n) nfa" where
+"nfa_rlin2 P S =
+  \<lparr> states = {S} \<union> Nts P,
+    init = {S},
+    final = {A \<in> Nts P. (A,[]) \<in> P},
+    nxt = \<lambda>q a. nxt_rlin2 P q a,
+    eps = Id \<rparr>"               
+
+context
+  fixes P :: "('n,'t)Prods"
+  assumes "finite P"
+begin
+
+interpretation NFA_rlin2: nfa "nfa_rlin2 P S"
+unfolding nfa_rlin2_def proof (standard, goal_cases)
+  case 1
+  then show ?case by(simp)
+next
+  case 2
+  then show ?case by auto
+next
+  case (3 q x)
+  then show ?case by(auto simp add: nxt_rlin2_nts)
+next
+  case 4
+  then show ?case using \<open>finite P\<close> by (simp add: Nts_def finite_nts_syms split_def)
+qed
+print_theorems
+
+lemma nfa_init_nfa_rlin2: "nfa.init (nfa_rlin2 P S) = {S}"
+by (simp add: nfa_rlin2_def)
+
+lemma nfa_final_nfa_rlin2: "nfa.final (nfa_rlin2 P S) = {A \<in> Nts P. (A,[]) \<in> P}"
+by (simp add: nfa_rlin2_def)
+
+lemma nfa_nxt_nfa_rlin2: "nfa.nxt (nfa_rlin2 P S) A a = nxt_rlin2 P A a"
+by (simp add: nfa_rlin2_def)
+
+lemma nfa_epsclo_nfa_rlin2: "M \<subseteq> {S} \<union> Nts P \<Longrightarrow> nfa.epsclo (nfa_rlin2 P S) M = M"
+unfolding NFA_rlin2.epsclo_def unfolding nfa_rlin2_def by(auto)
+
+lemma nfa_nextl_nfa_rlin2: "M \<subseteq> {S} \<union> Nts P
+  \<Longrightarrow> nfa.nextl (nfa_rlin2 P S) M xs = nxts_rlin2_set P M xs"
+proof(induction xs arbitrary: M)
+  case Nil
+  then show ?case
+    by (simp add: nxts_rlin2_set_def)(fastforce intro!: nfa_epsclo_nfa_rlin2)
+next
+  case (Cons a xs)
+  let ?epsclo = "nfa.epsclo (nfa_rlin2 P S)"
+  let ?nxt = "nfa.nxt (nfa_rlin2 P S)"
+  let ?nxts = "nfa.nextl (nfa_rlin2 P S)"
+  have "?nxts M (a # xs) = ?nxts (\<Union>x\<in>?epsclo M. ?nxt x a) xs"
+    by simp
+  also have "\<dots> = ?nxts (\<Union>x\<in>M. ?nxt x a) xs"
+    using Cons.prems by(subst nfa_epsclo_nfa_rlin2) auto
+  also have "\<dots> = ?nxts (\<Union>m\<in>M. nxt_rlin2 P m a) xs"
+    by (simp add: nfa_nxt_nfa_rlin2)
+  also have "\<dots> = nxts_rlin2_set P (\<Union>m\<in>M. nxt_rlin2 P m a) xs"
+    using Cons.prems by(subst Cons.IH)(auto simp add: nxt_rlin2_nts)
+  also have "\<dots> = nxts_rlin2_set P M (a # xs)"
+    by (simp add: nxt_rlin2_set_def nxts_rlin2_set_def)
+  finally show ?case .
+qed
+
+lemma lang_pres_nfa_rlin2: assumes "rlin2 P"
+shows "nfa.language (nfa_rlin2 P S) = Lang P S"
+proof -
+  have 1: "\<And>A xs. \<lbrakk> A \<in> nxts_rlin2_set P {S} xs; A \<in> Nts P; (A, []) \<in> P\<rbrakk> \<Longrightarrow>
+    P \<turnstile> [Nt S] \<Rightarrow>* map Tm xs"
+    using nxts_to_mult_derive by (metis (no_types, opaque_lifting) append.right_neutral derive.intros
+      r_into_rtranclp rtranclp_trans singletonD)
+  have "\<And>A B. Nt B \<notin> Syms P \<Longrightarrow> (A, []) \<in> P \<Longrightarrow> A \<noteq> B" by(auto simp: Syms_def)
+  hence 2: "\<And>xs. rlin2 P \<Longrightarrow> P \<turnstile> [Nt S] \<Rightarrow>* map Tm xs \<Longrightarrow>
+          nxts_rlin2_set P {S} xs \<inter> {A \<in> Nts P. (A, []) \<in> P} \<noteq> {}"
+   using in_Nts_iff_in_Syms mult_derive_to_nxts rlin2_tms_eps
+   by (metis (no_types, lifting) Int_Collect empty_iff singletonI)
+  show ?thesis
+    unfolding NFA_rlin2.language_def Lang_def nfa_init_nfa_rlin2 nfa_final_nfa_rlin2
+      nfa_nextl_nfa_rlin2[OF Un_upper1]
+    using 2[OF assms] by (auto simp: 1)
+qed
+
+lemma regular_if_rlin2: assumes "rlin2 P"
+  shows "regular (Lang P S)"
+using lang_pres_nfa_rlin2[OF assms] NFA_rlin2.imp_regular[of S]
+by metis
+
+end
+
+
+subsection \<open>From DFA to Strongly Right-Linear Grammar\<close>
+
+context dfa
+begin
+
+text
+\<open>We define \<open>Prods_dfa\<close> that collects the production set from the deterministic finite automata \<open>M\<close>\<close>
+
+definition Prods_dfa :: "('s, 'a) Prods" where
+"Prods_dfa =
+  (\<Union>q \<in> dfa.states M. \<Union>x. {(q,[Tm x, Nt(dfa.nxt M q x)])}) \<union> (\<Union>q \<in> dfa.final M. {(q,[])})"
+
+lemma rlin2_prods_dfa: "rlin2 (Prods_dfa)"
+  unfolding rlin2_def Prods_dfa_def by blast
+
+text
+\<open>We show that a word can be derived from the production set \<open>Prods_dfa\<close> if and only if traversing the word in the deterministic finite 
+ automata \<open>M\<close> ends in a final state. The proofs are very similar to those in \<open>DFA_rlin2.thy\<close>\<close>
+
+lemma mult_derive_to_nxtl:
+  "Prods_dfa \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt B] \<Longrightarrow> nextl A w = B"
+proof (induction w arbitrary: B rule: rev_induct)
+  case Nil
+  thus ?case
+    using rlin2_nts_derive_eq[OF rlin2_prods_dfa, of A B] by simp
+next
+  case (snoc x xs)
+  from snoc.prems have "Prods_dfa \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ [Tm x,Nt B]" by simp
+  then obtain C where C_der: "Prods_dfa \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ [Nt C]"
+               and C_prods: "(C,[Tm x, Nt B]) \<in> Prods_dfa" using rlin2_introduce_tm[OF rlin2_prods_dfa, of A xs x B] by auto 
+  have 1: "nextl A xs = C"
+    using snoc.IH[OF C_der] .
+  from C_prods have 2: "B = dfa.nxt M C x"
+    unfolding Prods_dfa_def by blast
+  from 1 2 show ?case by simp
+qed
+
+lemma nxtl_to_mult_derive:
+  assumes "A \<in> dfa.states M"
+    shows "Prods_dfa \<turnstile> [Nt A] \<Rightarrow>* map Tm w @ [Nt (nextl A w)]"
+proof (induction w rule: rev_induct)
+  case (snoc x xs)
+  let ?B = "nextl A xs"
+  have "?B \<in> dfa.states M" 
+    using nextl_state[OF assms, of xs] .
+  hence "(?B, [Tm x, Nt (dfa.nxt M ?B x)]) \<in> Prods_dfa"
+    unfolding Prods_dfa_def by blast
+  hence "Prods_dfa \<turnstile> [Nt ?B] \<Rightarrow> [Tm x] @ [Nt (dfa.nxt M ?B x)]"
+    by (simp add: derive_singleton)
+  hence "Prods_dfa \<turnstile> [Nt A] \<Rightarrow>* map Tm xs @ ([Tm x] @ [Nt (dfa.nxt M ?B x)])"
+    using snoc.IH by (meson derive_prepend rtranclp.simps)
+  thus ?case by auto
+qed simp
+
+theorem Prods_dfa_iff_dfa:
+  "q \<in> dfa.states M \<Longrightarrow> Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w \<longleftrightarrow> nextl q w \<in> dfa.final M"
+proof
+  show "Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w \<Longrightarrow> nextl q w \<in> dfa.final M"
+  proof -
+    assume asm: "Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w"
+    obtain B where q_der: "Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w @ [Nt B]" and B_in: "(B,[]) \<in> Prods_dfa" 
+      unfolding Lang_def using rlin2_tms_eps[OF rlin2_prods_dfa asm] by auto
+    have 1: "nextl q w = B"
+      using mult_derive_to_nxtl[OF q_der] .
+    from B_in have 2: "B \<in> dfa.final M"
+      unfolding Prods_dfa_def by blast
+    from 1 2 show ?thesis by simp
+  qed
+next
+  assume asm1: "q \<in> dfa.states M"
+  show "nextl q w \<in> dfa.final M \<Longrightarrow> Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w"
+  proof -
+    assume asm2: "nextl q w \<in> dfa.final M"
+    let ?Z = "nextl q w"
+    from asm2 have Z_eps: "(?Z,[]) \<in> Prods_dfa"
+      unfolding Prods_dfa_def by blast
+    have "Prods_dfa \<turnstile> [Nt q] \<Rightarrow>* map Tm w @ [Nt ?Z]"
+      using nxtl_to_mult_derive[OF asm1, of w] .
+    with Z_eps show ?thesis
+      by (metis derives_rule rtranclp.rtrancl_refl self_append_conv)
+  qed
+qed
+
+corollary dfa_language_eq_Lang: "dfa.language M = Lang Prods_dfa (dfa.init M)"
+unfolding language_def Lang_def by (simp add: Prods_dfa_iff_dfa)
+
+end
+
+corollary rlin2_if_regular:
+  "regular L \<Longrightarrow> \<exists>P S::hf. rlin2 P \<and> L = Lang P S"
+by (metis dfa.dfa_language_eq_Lang dfa.rlin2_prods_dfa regular_def)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/Unit_Elimination.thy
--- /dev/null
+++ thys/Context_Free_Grammar/Unit_Elimination.thy
@@ -0,0 +1,306 @@
+(*
+Author: August Martin Stimpfle
+Based on HOL4 theories by Aditi Barthwal
+*)
+
+section \<open>Elimination of Unit Productions\<close>
+
+theory Unit_Elimination
+imports Context_Free_Grammar
+begin
+
+(* Rules of the form A\<rightarrow>B, where A and B are in nonterminals ps *)
+definition unit_prods :: "('n,'t) prods \<Rightarrow> ('n,'t) Prods" where
+"unit_prods ps = {(l,r) \<in> set ps. \<exists>A. r = [Nt A]}"
+
+(* A \<Rightarrow>* B where A and B are in nonTerminals g *)
+definition unit_rtc :: "('n, 't) Prods \<Rightarrow> ('n \<times> 'n) set" where
+"unit_rtc Ps = {(A,B). Ps \<turnstile> [Nt A] \<Rightarrow>* [Nt B] \<and> {A,B} \<subseteq> Nts Ps}"
+
+definition unit_rm :: "('n, 't) prods \<Rightarrow> ('n, 't) Prods" where
+"unit_rm ps = (set ps - unit_prods ps)"
+
+definition new_prods :: "('n, 't) prods \<Rightarrow> ('n, 't) Prods" where 
+"new_prods ps = {(A,r). \<exists>B. (B,r) \<in> (unit_rm ps) \<and> (A, B) \<in> unit_rtc (unit_prods ps)}"
+
+definition unit_elim_rel :: "('n, 't) prods \<Rightarrow> ('n, 't) prods \<Rightarrow> bool" where
+"unit_elim_rel ps ps' \<equiv> set ps' = (unit_rm ps \<union> new_prods ps)"
+
+definition Unit_free :: "('n, 't) Prods \<Rightarrow> bool" where
+"Unit_free P = (\<nexists>A B. (A,[Nt B]) \<in> P)"
+
+lemma Unit_free_if_unit_elim_rel: "unit_elim_rel ps ps' \<Longrightarrow> Unit_free (set ps')" 
+unfolding unit_elim_rel_def unit_rm_def new_prods_def unit_prods_def Unit_free_def by simp
+
+lemma unit_elim_rel_Eps_free:
+  assumes "Eps_free (set ps)" and "unit_elim_rel ps ps'"
+  shows "Eps_free (set ps')"
+  using assms 
+  unfolding unit_elim_rel_def Eps_free_def unit_rm_def unit_prods_def new_prods_def by auto
+
+(* Finiteness & Existence *)
+
+(* finiteness unit_prods which also implies finiteness for unit_rm *)
+fun uprods :: "('n,'t) prods \<Rightarrow> ('n,'t) prods" where
+"uprods [] = []" |
+"uprods (p#ps) = (if \<exists>A. (snd p) = [Nt A] then p#uprods ps else uprods ps)"
+
+lemma unit_prods_uprods: "set (uprods ps) = unit_prods ps"
+unfolding unit_prods_def by (induction ps) auto
+
+lemma finiteunit_prods: "finite (unit_prods ps)"
+using unit_prods_uprods by (metis List.finite_set)
+
+(* finiteness for unit_rtc *)
+definition NtsCross :: "('n, 't) Prods  \<Rightarrow> ('n \<times> 'n) set" where
+"NtsCross Ps = {(A, B). A \<in> Nts Ps \<and> B \<in> Nts Ps }"
+
+lemma finite_unit_rtc: 
+  assumes "finite ps" 
+  shows  "finite (unit_rtc ps)"
+proof -
+  have "finite (Nts ps)"
+    unfolding Nts_def using assms finite_nts_syms by auto
+  hence "finite (NtsCross ps)"
+    unfolding NtsCross_def by auto
+  moreover have "unit_rtc ps \<subseteq> NtsCross ps"
+    unfolding unit_rtc_def NtsCross_def by blast
+  ultimately show ?thesis
+    using assms infinite_super by fastforce 
+qed
+
+(* finiteness for new_prods *)
+definition nPSlambda :: "('n, 't) Prods \<Rightarrow> ('n \<times> 'n) \<Rightarrow> ('n, 't) Prods" where
+"nPSlambda Ps d = {fst d} \<times> {r. (snd d, r) \<in> Ps}"
+
+lemma npsImage: "\<Union>((nPSlambda (unit_rm ps)) ` (unit_rtc (unit_prods ps))) = new_prods ps"
+  unfolding new_prods_def nPSlambda_def by fastforce
+
+lemma finite_nPSlambda:
+  assumes "finite Ps" 
+  shows "finite (nPSlambda Ps d)"
+proof -
+  have "{(B, r). (B, r) \<in> Ps \<and> B = snd d} \<subseteq> Ps" 
+    by blast
+  hence "finite {(B, r). (B, r) \<in> Ps \<and> B = snd d}"
+    using assms finite_subset by blast
+  hence "finite (snd ` {(B, r). (B, r) \<in> Ps \<and> B = snd d})" 
+    by simp
+  moreover have "{r. (snd d, r) \<in> Ps} = (snd ` {(B, r). (B, r) \<in> Ps \<and> B = snd d})"
+    by force
+  ultimately show ?thesis
+    using assms unfolding nPSlambda_def by simp
+qed
+
+lemma finite_new_prods: "finite (new_prods ps)"
+proof -
+  have "finite (unit_rm ps)"
+    unfolding unit_rm_def using finiteunit_prods by blast
+  moreover have "finite (unit_rtc (unit_prods ps))"
+    using finiteunit_prods finite_unit_rtc by blast
+  ultimately show ?thesis
+    using npsImage finite_nPSlambda finite_UN by metis
+qed
+
+lemma finiteunit_elim_relRules: "finite (unit_rm ps \<union> new_prods ps)"
+proof -
+  have "finite (unit_rm ps)"
+    unfolding unit_rm_def using finiteunit_prods by blast
+  moreover have "finite (new_prods ps)"
+    using finite_new_prods by blast
+  ultimately show ?thesis by blast
+qed
+
+lemma unit_elim_rel_exists: "\<forall>ps. \<exists>ps'. unit_elim_rel ps ps'"
+unfolding unit_elim_rel_def using finite_list[OF finiteunit_elim_relRules] by blast
+
+definition unit_elim where
+"unit_elim ps = (SOME ps'. unit_elim_rel ps ps')"
+
+lemma unit_elim_rel_unit_elim: "unit_elim_rel ps (unit_elim ps)"
+by (simp add: someI_ex unit_elim_def unit_elim_rel_exists)
+
+(* towards theorem 4.4 *)
+
+lemma inNonUnitProds:
+  "p \<in> unit_rm ps \<Longrightarrow> p \<in> set ps"
+  unfolding unit_rm_def by blast
+
+lemma psubDeriv:
+  assumes "ps \<turnstile> u \<Rightarrow> v"
+    and "\<forall>p \<in> ps. p \<in> ps'"
+  shows "ps' \<turnstile> u \<Rightarrow> v"
+  using assms by (meson derive_iff)
+
+lemma psubRtcDeriv:
+  assumes "ps \<turnstile> u \<Rightarrow>* v"
+    and "\<forall>p \<in> ps. p \<in> ps'"
+  shows "ps' \<turnstile> u \<Rightarrow>* v"
+  using assms by (induction rule: rtranclp.induct) (auto simp: psubDeriv rtranclp.rtrancl_into_rtrancl)
+
+lemma unit_prods_deriv: 
+  assumes "unit_prods ps \<turnstile> u \<Rightarrow>* v"
+  shows "set ps \<turnstile> u \<Rightarrow>* v"
+proof -
+  have "\<forall>p \<in> unit_prods ps. p \<in> set ps"
+    unfolding unit_prods_def by blast
+  thus ?thesis 
+    using assms psubRtcDeriv by blast
+qed
+
+lemma unit_elim_rel_r3:
+  assumes "unit_elim_rel ps ps'" and "set ps' \<turnstile> u \<Rightarrow> v"
+  shows "set ps \<turnstile> u \<Rightarrow>* v"
+proof -
+  obtain A \<alpha> r1 r2 where A: "(A, \<alpha>) \<in> set ps' \<and> u = r1 @ [Nt A] @ r2 \<and> v = r1 @ \<alpha> @ r2"
+    using assms derive.cases by meson
+  hence "(A, \<alpha>) \<in> unit_rm ps \<or> (A, \<alpha>) \<in> new_prods ps"
+    using assms(1) unfolding unit_elim_rel_def by simp
+  thus ?thesis
+  proof
+    assume "(A, \<alpha>) \<in> unit_rm ps"
+    hence "(A, \<alpha>) \<in> set ps"
+      using inNonUnitProds by blast
+    hence "set ps \<turnstile> r1 @ [Nt A] @ r2 \<Rightarrow> r1 @ \<alpha> @ r2"
+      by (auto simp: derive.simps)
+    thus ?thesis using A by simp
+  next 
+    assume "(A, \<alpha>) \<in> new_prods ps"
+    from this obtain B where B: "(B, \<alpha>) \<in> unit_rm ps \<and> (A, B) \<in> unit_rtc (unit_prods ps)"
+      unfolding new_prods_def by blast
+    hence "unit_prods ps \<turnstile> [Nt A] \<Rightarrow>* [Nt B]"
+      unfolding unit_rtc_def by simp
+    hence "set ps \<turnstile> [Nt A] \<Rightarrow>* [Nt B]"
+      using unit_prods_deriv by blast
+    hence 1: "set ps \<turnstile> r1 @ [Nt A] @ r2 \<Rightarrow>* r1 @ [Nt B] @ r2"
+      using derives_append derives_prepend by blast
+    have "(B, \<alpha>) \<in> set ps"
+      using B inNonUnitProds by blast
+    hence "set ps \<turnstile> r1 @ [Nt B] @ r2 \<Rightarrow> r1 @ \<alpha> @ r2"
+      by (auto simp: derive.simps)
+    thus ?thesis 
+      using 1 A by simp
+  qed
+qed
+
+lemma unit_elim_rel_r4: 
+  assumes "set ps' \<turnstile> u \<Rightarrow>* v"
+    and "unit_elim_rel ps ps'"
+  shows "set ps \<turnstile> u \<Rightarrow>* v"
+  using assms by (induction rule: rtranclp.induct) (auto simp: unit_elim_rel_r3 rtranclp_trans)
+
+lemma deriv_unit_rtc:
+  assumes "set ps \<turnstile> [Nt A] \<Rightarrow> [Nt B]"
+  shows "(A, B) \<in> unit_rtc (unit_prods ps)"
+proof -
+  have "(A, [Nt B]) \<in> set ps"
+    using assms by (simp add: derive_singleton)
+  hence "(A, [Nt B]) \<in> unit_prods ps"
+    unfolding unit_prods_def by blast
+  hence "unit_prods ps \<turnstile> [Nt A] \<Rightarrow> [Nt B]"
+    by (simp add: derive_singleton)
+  moreover have "B \<in> Nts (unit_prods ps) \<and> A \<in> Nts (unit_prods ps)"
+    using \<open>(A, [Nt B]) \<in> unit_prods ps\<close> Nts_def nts_syms_def by fastforce
+  ultimately show ?thesis
+    unfolding unit_rtc_def by blast
+qed
+
+lemma unit_elim_rel_r12: 
+  assumes "unit_elim_rel ps ps'" "(A, \<alpha>) \<in> set ps'"
+  shows "(A, \<alpha>) \<notin> unit_prods ps"
+  using assms unfolding unit_elim_rel_def unit_rm_def unit_prods_def new_prods_def by blast
+
+lemma unit_elim_rel_r14: 
+  assumes "unit_elim_rel ps ps'" 
+    and "set ps \<turnstile> [Nt A] \<Rightarrow> [Nt B]" "set ps' \<turnstile> [Nt B] \<Rightarrow> v"
+  shows "set ps' \<turnstile> [Nt A] \<Rightarrow> v"
+proof -
+  have 1: "(A, B) \<in> unit_rtc (unit_prods ps)"
+    using deriv_unit_rtc assms(2) by fast
+  have 2: "(B, v) \<in> set ps'"
+    using assms(3) by (simp add: derive_singleton)
+  thus ?thesis
+  proof (cases "(B, v) \<in> set ps")
+    case True
+    hence "(B, v) \<in> unit_rm ps"
+      unfolding unit_rm_def using assms(1) assms(3) unit_elim_rel_r12[of ps ps' B v] by (simp add: derive_singleton)
+    then show ?thesis
+    using 1 assms(1) unfolding unit_elim_rel_def new_prods_def derive_singleton by blast
+  next
+    case False
+    hence "(B, v) \<in> new_prods ps"
+      using assms(1) 2 unfolding unit_rm_def unit_elim_rel_def  by simp
+    from this obtain C where C: "(C, v) \<in> unit_rm ps \<and> (B, C) \<in> unit_rtc (unit_prods ps)"
+      unfolding new_prods_def by blast
+    hence "unit_prods ps \<turnstile> [Nt A] \<Rightarrow>* [Nt C]"
+      using 1 unfolding unit_rtc_def by auto
+    hence "(A, C) \<in> unit_rtc (unit_prods ps)"
+      unfolding unit_rtc_def using 1 C unit_rtc_def by fastforce
+    hence "(A, v) \<in> new_prods ps"
+      unfolding new_prods_def using C by blast
+    hence "(A, v) \<in> set ps'"
+      using assms(1) unfolding unit_elim_rel_def  by blast
+    thus ?thesis by (simp add: derive_singleton)
+  qed
+qed
+
+lemma unit_elim_rel_r20_aux:
+  assumes "set ps \<turnstile> l @ [Nt A] @ r \<Rightarrow>* map Tm v" 
+  shows "\<exists>\<alpha>. set ps \<turnstile> l @ [Nt A] @ r \<Rightarrow> l @ \<alpha> @ r \<and> set ps \<turnstile> l @ \<alpha> @ r \<Rightarrow>* map Tm v \<and> (A, \<alpha>) \<in> set ps"
+proof -
+  obtain l' w r' where w: "set ps \<turnstile> l \<Rightarrow>* l'  \<and> set ps \<turnstile> [Nt A] \<Rightarrow>* w \<and>  set ps \<turnstile> r \<Rightarrow>* r' \<and> map Tm v = l' @ w @ r'"
+    using assms(1) by (metis derives_append_decomp)
+  have "Nt A \<notin> set (map Tm v)" 
+    using assms(1) by auto
+  hence "[Nt A] \<noteq> w" 
+    using w by auto
+  from this obtain \<alpha> where \<alpha>: "set ps \<turnstile> [Nt A] \<Rightarrow> \<alpha> \<and> set ps \<turnstile> \<alpha> \<Rightarrow>* w"
+    by (metis w converse_rtranclpE)
+  hence "(A, \<alpha>) \<in> set ps" 
+    by (simp add: derive_singleton)
+  thus ?thesis by (metis \<alpha> w derive.intros derives_append_decomp) 
+qed
+
+lemma unit_elim_rel_r20: 
+  assumes "set ps \<turnstile> u \<Rightarrow>* map Tm v" "unit_elim_rel ps ps'"
+  shows "set ps' \<turnstile> u \<Rightarrow>* map Tm v"
+  using assms proof (induction rule: converse_derives_induct)
+  case base
+  then show ?case by blast
+next
+  case (step l A r w)
+  then show ?case 
+  proof (cases "(A, w) \<in> unit_prods ps")
+    case True
+    from this obtain B where "w = [Nt B]"
+      unfolding unit_prods_def by blast
+    have "set ps' \<turnstile> l @ w @ r \<Rightarrow>* map Tm v \<and> Nt B \<notin> set (map Tm v)"
+      using step.IH assms(2) by auto
+    obtain \<alpha> where \<alpha>: "set ps' \<turnstile> l @ [Nt B] @ r \<Rightarrow> l @ \<alpha> @ r \<and> set ps' \<turnstile> l @ \<alpha> @ r \<Rightarrow>* map Tm v \<and> (B, \<alpha>) \<in> set ps'"
+      using assms(2) step.IH \<open>w=_\<close>  unit_elim_rel_r20_aux[of ps' l B r v] by blast
+    hence "(A, \<alpha>) \<in> set ps'"
+      using assms(2) step.hyps(2) \<open>w=_\<close> unit_elim_rel_r14[of ps ps' A B \<alpha>] by (simp add: derive_singleton)
+    hence "set ps' \<turnstile> l @ [Nt A] @ r \<Rightarrow>* l @ \<alpha> @ r"
+      using derive.simps by fastforce
+    then show ?thesis 
+      using \<alpha> by auto
+  next
+    case False
+    hence "(A, w) \<in> unit_rm ps"
+      unfolding unit_rm_def using step.hyps(2) by blast
+    hence "(A, w) \<in> set ps'"
+      using assms(2) unfolding unit_elim_rel_def  by simp
+    hence "set ps' \<turnstile> l @ [Nt A] @ r \<Rightarrow> l @ w @ r"
+      by (auto simp: derive.simps)
+    then show ?thesis
+      using step by simp
+  qed
+qed
+
+theorem unit_elim_rel_lang_eq: "unit_elim_rel ps ps' \<Longrightarrow> lang ps' S = lang ps S"
+  unfolding Lang_def using unit_elim_rel_r4 unit_elim_rel_r20 by blast
+
+corollary lang_unit_elim: "lang (unit_elim ps) A = lang ps A"
+by (rule unit_elim_rel_lang_eq[OF unit_elim_rel_unit_elim])
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/document/root.bib
--- /dev/null
+++ thys/Context_Free_Grammar/document/root.bib
@@ -0,0 +1,34 @@
+@book{HopcroftMU,
+  title = {Introduction to Automata Theory, Languages, and Computation},
+  author = {J. E. Hopcroft and R. Motwani and J. D. Ullman},
+  year = {2006},
+  edition = {3rd},
+  publisher = {Addison Wesley}
+}
+
+@inproceedings{RamosAMQ,
+  author       = {Marcus V. M. Ramos and
+                  Jos{\'{e}} Carlos Bacelar Almeida and
+                  Nelma Moreira and
+                  Ruy J. G. B. de Queiroz},
+  editor       = {Beniamino Accattoli and
+                  Carlos Olarte},
+  title        = {Some Applications of the Formalization of the Pumping Lemma for Context-Free
+                  Languages},
+  booktitle    = {Proceedings of the 13th Workshop on Logical and Semantic Frameworks
+                  with Applications, {LSFA} 2018},
+  series       = {Electronic Notes in Theoretical Computer Science},
+  volume       = {344},
+  pages        = {151--167},
+  publisher    = {Elsevier},
+  year         = {2018},
+  url          = {https://doi.org/10.1016/j.entcs.2019.07.010},
+}
+
+@misc{RamosGithub,
+    title = {GitHub repository},
+    url = {https://github.com/mvmramos/intersection},
+    author = {M.V.M. Ramos},
+    year = {2018},
+    note = {Accessed on 8/5/2025}
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Context_Free_Grammar/document/root.tex
--- /dev/null
+++ thys/Context_Free_Grammar/document/root.tex
@@ -0,0 +1,39 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{amssymb}
+\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{literal}
+
+
+\begin{document}
+
+\title{Context-Free Grammars and Languages}
+\author{Tobias Nipkow, Markus Gscho{\ss}mann, Felix Krayer, Fabian Lehr, \\
+ Bruno Philipp, August Martin Stimpfle, Kaan Taskin, Akihisa Yamada}
+\maketitle
+
+\begin{abstract}
+This is a basic library of definitions and results about context-free grammars and languages.
+It includes context-free grammars and languages, parse trees, Chomsky normal form,
+pumping lemmas and the relationship of right-linear grammars to finite automata.
+\end{abstract}
+
+\tableofcontents
+
+\newpage
+
+%\section{Overview}
+
+% include generated text of all theories
+\input{session}
+
+\bibliographystyle{abbrvurl}
+\bibliography{root}
+
+\end{document}
diff -r 848f2c00e57f -r 29846bc89de2 thys/Diagonal_Ramsey/Big_Blue_Steps.thy
--- thys/Diagonal_Ramsey/Big_Blue_Steps.thy
+++ thys/Diagonal_Ramsey/Big_Blue_Steps.thy
@@ -410,9 +410,9 @@
     using W_def \<open>U \<subseteq> W\<close> by blast
   then have cardU_less_X: "card U < card X"
     by (meson \<open>X\<subseteq>V\<close> finV finite_subset psubset_card_mono)
-  with \<open>X\<subseteq>V\<close> have cardXU: "card (X-U) = card X - card U"
+  with \<open>X\<subseteq>V\<close> have cardXU: "card (X\<setminus>U) = card X - card U"
     by (meson \<open>U \<subset> X\<close> card_Diff_subset finV finite_subset psubset_imp_subset)
-  then have real_cardXU: "real (card (X-U)) = real (card X) - m"
+  then have real_cardXU: "real (card (X\<setminus>U)) = real (card X) - m"
     using \<open>card U = m\<close> cardU_less_X by linarith
   have [simp]: "m \<le> card X"
     using \<open>card U = m\<close> cardU_less_X nless_le by blast
@@ -429,57 +429,56 @@
   text \<open>First part of (10)\<close>
   have "card U * (\<mu> * card X - card U) = m * (\<mu> * (card X - card U)) - (1-\<mu>) * m2"
     using cardU_less_X by (simp add: \<open>card U = m\<close> algebra_simps numeral_2_eq_2)
-  also have "\<dots> \<le> real (card (Blue \<inter> all_edges_betw_un U (X-U)))"
+  also have "\<dots> \<le> real (card (Blue \<inter> all_edges_betw_un U (X\<setminus>U)))"
   proof -
-    have dfam: "disjoint_family_on (\<lambda>u. Blue \<inter> all_edges_betw_un {u} (X-U)) U"
+    have dfam: "disjoint_family_on (\<lambda>u. Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)) U"
       by (auto simp: disjoint_family_on_def all_edges_betw_un_def)
-    have "\<mu> * (card X - card U) \<le> card (Blue \<inter> all_edges_betw_un {u} (X-U)) + (1-\<mu>) * m" 
+    have "\<mu> * (card X - card U) \<le> card (Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)) + (1-\<mu>) * m" 
       if "u \<in> U" for u
     proof -
-      have NBU: "Neighbours Blue u \<inter> U = U - {u}"
+      have NBU: "Neighbours Blue u \<inter> U = U \<setminus> {u}"
         using \<open>clique U Blue\<close> Red_Blue_all singleton_not_edge that 
         by (force simp: Neighbours_def clique_def)
-      then have NBX_split: "(Neighbours Blue u \<inter> X) = (Neighbours Blue u \<inter> (X-U)) \<union> (U - {u})"
+      then have NBX_split: "(Neighbours Blue u \<inter> X) = (Neighbours Blue u \<inter> (X\<setminus>U)) \<union> (U \<setminus> {u})"
         using \<open>U \<subset> X\<close> by blast
-      moreover have "Neighbours Blue u \<inter> (X-U) \<inter> (U - {u}) = {}"
+      moreover have "Neighbours Blue u \<inter> (X\<setminus>U) \<inter> (U \<setminus> {u}) = {}"
         by blast
-      ultimately have "card(Neighbours Blue u \<inter> X) = card(Neighbours Blue u \<inter> (X-U)) + (m - Suc 0)"
+      ultimately have "card(Neighbours Blue u \<inter> X) = card(Neighbours Blue u \<inter> (X\<setminus>U)) + (m-1)"
         by (simp add: card_Un_disjoint finite_Neighbours \<open>finite U\<close> \<open>card U = m\<close> that)
-      then have "\<mu> * (card X) \<le> real (card (Neighbours Blue u \<inter> (X-U))) + real (m - Suc 0)"
+      then have "\<mu> * (card X) \<le> real (card (Neighbours Blue u \<inter> (X\<setminus>U))) + real (m - 1)"
         using W_def \<open>U \<subseteq> W\<close> bluish_def that by force
-      then have "\<mu> * (card X - card U) 
-                \<le> card (Neighbours Blue u \<inter> (X-U)) + real (m - Suc 0) - \<mu> *card U"
+      then have "\<mu> * (card X - card U) \<le> card (Neighbours Blue u \<inter> (X\<setminus>U)) + real (m-1) - \<mu> * card U"
         by (smt (verit) cardU_less_X nless_le of_nat_diff right_diff_distrib')
-      then have *: "\<mu> * (card X - card U) \<le> real (card (Neighbours Blue u \<inter> (X-U))) + (1-\<mu>)*m"
+      then have *: "\<mu> * (card X - card U) \<le> real (card (Neighbours Blue u \<inter> (X\<setminus>U))) + (1-\<mu>)*m"
         using assms by (simp add: \<open>card U = m\<close> left_diff_distrib)
       have "inj_on (\<lambda>x. {u,x}) (Neighbours Blue u \<inter> X)"
         by (simp add: doubleton_eq_iff inj_on_def)
-      moreover have "(\<lambda>x. {u,x}) ` (Neighbours Blue u \<inter> (X-U)) \<subseteq> Blue \<inter> all_edges_betw_un {u} (X-U)"
+      moreover have "(\<lambda>x. {u,x}) ` (Neighbours Blue u \<inter> (X\<setminus>U)) \<subseteq> Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)"
         using Blue_E by (auto simp: Neighbours_def all_edges_betw_un_def)
-      ultimately have "card (Neighbours Blue u \<inter> (X-U)) \<le> card (Blue \<inter> all_edges_betw_un {u} (X-U))"
+      ultimately have "card (Neighbours Blue u \<inter> (X\<setminus>U)) \<le> card (Blue \<inter> all_edges_betw_un {u} (X\<setminus>U))"
         by (metis NBX_split card_inj_on_le finite_Blue finite_Int inj_on_Un)
       with * show ?thesis
         by auto
     qed
     then have "(card U) * (\<mu> * real (card X - card U))
-             \<le> (\<Sum>x\<in>U. card (Blue \<inter> all_edges_betw_un {x} (X-U)) + (1-\<mu>) * m)"
+             \<le> (\<Sum>x\<in>U. card (Blue \<inter> all_edges_betw_un {x} (X\<setminus>U)) + (1-\<mu>) * m)"
       by (meson sum_bounded_below)
     then have "m * (\<mu> * (card X - card U))
-             \<le> (\<Sum>x\<in>U. card (Blue \<inter> all_edges_betw_un {x} (X-U))) + (1-\<mu>) * m2"
+             \<le> (\<Sum>x\<in>U. card (Blue \<inter> all_edges_betw_un {x} (X\<setminus>U))) + (1-\<mu>) * m2"
       by (simp add: sum.distrib power2_eq_square \<open>card U = m\<close> mult_ac)
-    also have "\<dots> \<le> card (\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X-U)) + (1-\<mu>) * m2"
+    also have "\<dots> \<le> card (\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)) + (1-\<mu>) * m2"
       by (simp add: dfam card_UN_disjoint' \<open>finite U\<close> flip: UN_simps)
     finally have "m * (\<mu> * (card X - card U)) 
-                \<le> card (\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X-U)) + (1-\<mu>) * m2" .
-    moreover have "(\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X-U)) = (Blue \<inter> all_edges_betw_un U (X-U))"
+                \<le> card (\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)) + (1-\<mu>) * m2" .
+    moreover have "(\<Union>u\<in>U. Blue \<inter> all_edges_betw_un {u} (X\<setminus>U)) = (Blue \<inter> all_edges_betw_un U (X\<setminus>U))"
       by (auto simp: all_edges_betw_un_def)
     ultimately show ?thesis
       by simp
   qed
-  also have "\<dots> \<le> edge_card Blue U (X-U)"
+  also have "\<dots> \<le> edge_card Blue U (X\<setminus>U)"
     by (simp add: edge_card_def)
-  finally have edge_card_XU: "edge_card Blue U (X-U) \<ge> card U * (\<mu> * card X - card U)" .
-  define \<sigma> where "\<sigma> \<equiv> blue_density U (X-U)"
+  finally have edge_card_XU: "edge_card Blue U (X\<setminus>U) \<ge> card U * (\<mu> * card X - card U)" .
+  define \<sigma> where "\<sigma> \<equiv> blue_density U (X\<setminus>U)"
   then have "\<sigma> \<ge> 0" by (simp add: gen_density_ge0)
   have "\<sigma> \<le> 1"
     by (simp add: \<sigma>_def gen_density_le1)
@@ -534,7 +533,7 @@
   then have "\<sigma> > 0"
     using \<open>0 < b\<close> \<open>0 \<le> \<sigma>\<close> less_eq_real_def by force
 
-  define \<Phi> where "\<Phi> \<equiv> \<Sum>v \<in> X-U. card (Neighbours Blue v \<inter> U) choose b"
+  define \<Phi> where "\<Phi> \<equiv> \<Sum>v \<in> X\<setminus>U. card (Neighbours Blue v \<inter> U) choose b"
   text \<open>now for the material between (10) and (11)\<close>
   have "\<sigma> * real m / 2 \<le> m"
     using \<open>\<sigma> \<le> 1\<close> \<open>m>0\<close> by auto
@@ -579,9 +578,9 @@
     by simp
   also have "\<dots> \<le> 2 * (\<sigma>^b) * exp(- (real b)2 / (\<sigma>*m)) * (card X - m)"
     by (intro mult_right_mono) (auto simp: \<open>0 \<le> \<sigma>\<close>)
-  finally have "\<mu>^b/2 * card X \<le> \<sigma>^b * exp(- of_nat (b2) / (\<sigma>*m)) * card (X-U)"
+  finally have "\<mu>^b/2 * card X \<le> \<sigma>^b * exp(- of_nat (b2) / (\<sigma>*m)) * card (X\<setminus>U)"
     by (simp add: \<open>card U = m\<close> cardXU real_cardXU)
-  also have "\<dots> \<le> 1/(m choose b) * ((\<sigma>*m) gchoose b) * card (X-U)"
+  also have "\<dots> \<le> 1/(m choose b) * ((\<sigma>*m) gchoose b) * card (X\<setminus>U)"
   proof (intro mult_right_mono)
     have "0 < real m gchoose b"
       by (metis \<open>b \<le> m\<close> binomial_gbinomial of_nat_0_less_iff zero_less_binomial_iff)
@@ -595,27 +594,27 @@
   also have "\<dots> \<le> 1/(m choose b) * \<Phi>"
     unfolding mult.assoc
   proof (intro mult_left_mono)
-    have eeq: "edge_card Blue U (X-U) = (\<Sum>i\<in>X-U. card (Neighbours Blue i \<inter> U))"
+    have eeq: "edge_card Blue U (X\<setminus>U) = (\<Sum>i\<in>X\<setminus>U. card (Neighbours Blue i \<inter> U))"
     proof (intro edge_card_eq_sum_Neighbours)
-      show "finite (X-U)"
+      show "finite (X\<setminus>U)"
         by (meson \<open>X\<subseteq>V\<close> finV finite_Diff finite_subset)
     qed (use disjnt_def Blue_E in auto)
-    have "(\<Sum>i\<in>X - U. card (Neighbours Blue i \<inter> U)) / (real (card X) - m) = blue_density U (X-U) * m"
+    have "(\<Sum>i\<in>X\<setminus>U. card (Neighbours Blue i \<inter> U)) / (real (card X) - m) = blue_density U (X\<setminus>U) * m"
       using \<open>m>0\<close> by (simp add: gen_density_def real_cardXU \<open>card U = m\<close> eeq divide_simps)
-    then have *: "(\<Sum>i\<in>X - U. real (card (Neighbours Blue i \<inter> U)) /R real (card (X-U))) = \<sigma> * m"
+    then have *: "(\<Sum>i\<in>X\<setminus>U. real (card (Neighbours Blue i \<inter> U)) /R real (card (X\<setminus>U))) = \<sigma> * m"
       by (simp add: \<sigma>_def divide_inverse_commute real_cardXU flip: sum_distrib_left)
-    have "mbinomial (\<Sum>i\<in>X - U. real (card (Neighbours Blue i \<inter> U)) /R (card (X-U))) b 
-         \<le> (\<Sum>i\<in>X - U. inverse (real (card (X-U))) * mbinomial (card (Neighbours Blue i \<inter> U)) b)"
+    have "mbinomial (\<Sum>i\<in>X\<setminus>U. real (card (Neighbours Blue i \<inter> U)) /R (card (X\<setminus>U))) b 
+         \<le> (\<Sum>i\<in>X\<setminus>U. inverse (real (card (X\<setminus>U))) * mbinomial (card (Neighbours Blue i \<inter> U)) b)"
     proof (rule convex_on_sum)
-      show "finite (X-U)"
+      show "finite (X\<setminus>U)"
         using cardU_less_X zero_less_diff by fastforce
       show "convex_on UNIV (\<lambda>a. mbinomial a b)"
         by (simp add: \<open>0 < b\<close> convex_mbinomial)
-      show "(\<Sum>i\<in>X - U. inverse (card (X-U))) = 1"
+      show "(\<Sum>i\<in>X\<setminus>U. inverse (card (X\<setminus>U))) = 1"
         using cardU_less_X cardXU by force
     qed (use \<open>U \<subset> X\<close> in auto)
     with ble 
-    show "(\<sigma>*m gchoose b) * card (X-U) \<le> \<Phi>"
+    show "(\<sigma>*m gchoose b) * card (X\<setminus>U) \<le> \<Phi>"
       unfolding * \<Phi>_def 
       by (simp add: cardU_less_X cardXU binomial_gbinomial divide_simps flip: sum_distrib_left sum_divide_distrib)
   qed auto
@@ -633,10 +632,10 @@
   have measure_eq: "measure M C = (if C \<subseteq> \<Omega> then card C / card \<Omega> else 0)" for C
     by (simp add: M_def fin\<Omega> measure_uniform_count_measure_if) 
 
-  define Int_NB where "Int_NB \<equiv> \<lambda>S. \<Inter>v\<in>S. Neighbours Blue v \<inter> (X-U)"
+  define Int_NB where "Int_NB \<equiv> \<lambda>S. \<Inter>v\<in>S. Neighbours Blue v \<inter> (X\<setminus>U)"
   have sum_card_NB: "(\<Sum>A\<in>\<Omega>. card (\<Inter>(Neighbours Blue ` A) \<inter> Y)) 
                      = (\<Sum>v\<in>Y. card (Neighbours Blue v \<inter> U) choose b)"
-    if "finite Y" "Y \<subseteq> X-U" for Y
+    if "finite Y" "Y \<subseteq> X\<setminus>U" for Y
     using that
   proof (induction Y)
     case (insert y Y)
@@ -650,10 +649,10 @@
           sum.subset_diff flip: insert.IH)
   qed auto
 
-  have "(\<Sum>x\<in>\<Omega>. card (if x = {} then UNIV else \<Inter> (Neighbours Blue ` x) \<inter> (X-U))) 
-        = (\<Sum>x\<in>\<Omega>. card (\<Inter> (Neighbours Blue ` x) \<inter> (X-U)))"
+  have "(\<Sum>x\<in>\<Omega>. card (if x = {} then UNIV else \<Inter> (Neighbours Blue ` x) \<inter> (X\<setminus>U))) 
+        = (\<Sum>x\<in>\<Omega>. card (\<Inter> (Neighbours Blue ` x) \<inter> (X\<setminus>U)))"
     unfolding \<Omega>_def nsets_def using \<open>0 < b\<close> by (force intro: sum.cong)
-  also have "\<dots> = (\<Sum>v\<in>X - U. card (Neighbours Blue v \<inter> U) choose b)"
+  also have "\<dots> = (\<Sum>v\<in>X\<setminus>U. card (Neighbours Blue v \<inter> U) choose b)"
     by (metis sum_card_NB \<open>X\<subseteq>V\<close> dual_order.refl finV finite_Diff rev_finite_subset)
   finally have "sum (card o Int_NB) \<Omega> = \<Phi>"
     by (simp add: \<Omega>_def \<Phi>_def Int_NB_def)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Diagonal_Ramsey/Book.thy
--- thys/Diagonal_Ramsey/Book.thy
+++ thys/Diagonal_Ramsey/Book.thy
@@ -251,7 +251,7 @@
                       - red_density X Y * card (Neighbours Red x \<inter> Y))"
 
 definition weight :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> real" where
-  "weight \<equiv> \<lambda>X Y x. \<Sum>y \<in> X-{x}. Weight X Y x y"
+  "weight \<equiv> \<lambda>X Y x. \<Sum>y \<in> X\<setminus>{x}. Weight X Y x y"
 
 definition p0 :: "real"
   where "p0 \<equiv> red_density X0 Y0"
@@ -260,7 +260,7 @@
   where "qfun \<equiv> \<lambda>h. p0 + qfun_base k h"
 
 lemma qfun_eq: "qfun \<equiv> \<lambda>h. p0 + ((1 + \<epsilon>)^h - 1) / k"
-  by (simp add: qfun_def qfun_base_def eps_def eps_def)
+  by (simp add: qfun_def qfun_base_def eps_def)
 
 definition hgt :: "real \<Rightarrow> nat"
   where "hgt \<equiv> \<lambda>p. LEAST h. p \<le> qfun h \<and> h>0"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Diagonal_Ramsey/Neighbours.thy
--- thys/Diagonal_Ramsey/Neighbours.thy
+++ thys/Diagonal_Ramsey/Neighbours.thy
@@ -174,7 +174,7 @@
 
 lemma edge_card_diff:
   assumes "Y\<subseteq>X" "disjnt X Z" "finite X" 
-  shows "edge_card C (X-Y) Z = edge_card C X Z - edge_card C Y Z"
+  shows "edge_card C (X\<setminus>Y) Z = edge_card C X Z - edge_card C Y Z"
 proof -
   have "(X\<setminus>Y) \<union> Y = X" "disjnt (X\<setminus>Y) Y"
     by (auto simp: Un_absorb2 assms disjnt_iff)
@@ -304,12 +304,12 @@
   by (smt (verit) card.infinite divide_le_eq_1 edge_card_le mult_eq_0_iff of_nat_le_0_iff of_nat_mono)
 
 lemma gen_density_le_1_minus: 
-  shows "gen_density C X Y \<le> 1 - gen_density (E-C) X Y"
+  shows "gen_density C X Y \<le> 1 - gen_density (E\<setminus>C) X Y"
 proof (cases "finite X \<and> finite Y")
   case True
-  have "C \<inter> all_edges_betw_un X Y \<union> (E - C) \<inter> all_edges_betw_un X Y = all_edges_betw_un X Y"
+  have "C \<inter> all_edges_betw_un X Y \<union> (E\<setminus>C) \<inter> all_edges_betw_un X Y = all_edges_betw_un X Y"
     by (auto simp: all_edges_betw_un_def)
-  with True have "(edge_card C X Y) + (edge_card (E - C) X Y) \<le> card (all_edges_betw_un X Y)"
+  with True have "(edge_card C X Y) + (edge_card (E\<setminus>C) X Y) \<le> card (all_edges_betw_un X Y)"
     unfolding edge_card_def
     by (metis Diff_Int_distrib2 Diff_disjoint card_Un_disjoint card_Un_le finite_Int finite_all_edges_betw_un)
   with True show ?thesis
@@ -318,16 +318,16 @@
 qed (auto simp: gen_density_def)
 
 lemma gen_density_lt1: 
-  assumes "{x,y} \<in> E-C" "x \<in> X" "y \<in> Y" "C \<subseteq> E"
+  assumes "{x,y} \<in> E\<setminus>C" "x \<in> X" "y \<in> Y" "C \<subseteq> E"
   shows "gen_density C X Y < 1"
 proof (cases "finite X \<and> finite Y")
   case True
-  then have "0 < gen_density (E - C) X Y"
+  then have "0 < gen_density (E\<setminus>C) X Y"
     using assms gen_density_gt0 by auto
-  have "gen_density C X Y \<le> 1 - gen_density (E - C) X Y"
+  have "gen_density C X Y \<le> 1 - gen_density (E\<setminus>C) X Y"
     by (intro gen_density_le_1_minus)
   then show ?thesis
-    using \<open>0 < gen_density (E - C) X Y\<close> by linarith
+    using \<open>0 < gen_density (E\<setminus>C) X Y\<close> by linarith
 qed (auto simp: gen_density_def)
 
 lemma gen_density_le_iff:
@@ -341,7 +341,7 @@
 lemma gen_density_below_avg_ge:
   assumes "disjnt X Z" "finite X" "Y\<subset>X" "finite Z" 
     and genY: "gen_density C Y Z \<le> gen_density C X Z"
-  shows "gen_density C (X-Y) Z \<ge> gen_density C X Z"
+  shows "gen_density C (X\<setminus>Y) Z \<ge> gen_density C X Z"
 proof -
   have "real (edge_card C Y Z) / card Y \<le> real (edge_card C X Z) / card X"
     using assms
diff -r 848f2c00e57f -r 29846bc89de2 thys/Diagonal_Ramsey/Red_Steps.thy
--- thys/Diagonal_Ramsey/Red_Steps.thy
+++ thys/Diagonal_Ramsey/Red_Steps.thy
@@ -306,7 +306,7 @@
       using W1 by (smt (verit) real_of_card sum_mono)
     finally show ?thesis .
   qed
-  have "weight X Y x \<le> real (card(X - {x})) * 1" for x
+  have "weight X Y x \<le> real (card(X \<setminus> {x})) * 1" for x
     unfolding weight_def by (meson DiffE abs_le_D1 sum_bounded_above W1)
   then have wgt_le_X1: "weight X Y x \<le> card X - 1" if "x \<in> X" for x
     using that card_Diff_singleton One_nat_def by (smt (verit, best)) 
@@ -315,13 +315,13 @@
     using not_many_bluish by (auto simp: m_of_def many_bluish_def XB_def)
   have "XB \<subseteq> X" "finite XB"
     using \<open>finite X\<close> by (auto simp: XB_def)
-  then have cv_non_XB: "\<And>y. y \<in> X - XB \<Longrightarrow> central_vertex X y"
+  then have cv_non_XB: "\<And>y. y \<in> X \<setminus> XB \<Longrightarrow> central_vertex X y"
     by (auto simp: central_vertex_def XB_def bluish_def)
   have "0 \<le> (\<Sum>y\<in>X. weight X Y y + Weight X Y y y)"
     by (fact ge0)
-  also have "\<dots> = (\<Sum>y\<in>XB. weight X Y y + Weight X Y y y) + (\<Sum>y\<in>X-XB. weight X Y y + Weight X Y y y)"
+  also have "\<dots> = (\<Sum>y\<in>XB. weight X Y y + Weight X Y y y) + (\<Sum>y\<in>X\<setminus>XB. weight X Y y + Weight X Y y y)"
     using sum.subset_diff [OF \<open>XB\<subseteq>X\<close>] by (smt (verit) X_def Xseq_subset_V finV finite_subset)
-  also have "\<dots> \<le> (\<Sum>y\<in>XB. weight X Y y + Weight X Y y y) + (\<Sum>y\<in>X-XB. weight X Y (cvx i) + 1)"
+  also have "\<dots> \<le> (\<Sum>y\<in>XB. weight X Y y + Weight X Y y y) + (\<Sum>y\<in>X\<setminus>XB. weight X Y (cvx i) + 1)"
     by (intro add_mono sum_mono w_maximal W1 order_refl cv_non_XB)
   also have "\<dots> = (\<Sum>y\<in>XB. weight X Y y + Weight X Y y y) + (card X - card XB) * (weight X Y (cvx i) + 1)"
     using \<open>XB\<subseteq>X\<close> \<open>finite XB\<close> by (simp add: card_Diff_subset)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Elementary_Ultrametric_Spaces/Elementary_Ultrametric_Spaces.thy
--- /dev/null
+++ thys/Elementary_Ultrametric_Spaces/Elementary_Ultrametric_Spaces.thy
@@ -0,0 +1,434 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF
+   Author: Benjamin Puyobro, Université Paris-Saclay, IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+   Author: Burkhart Wolff, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory Elementary_Ultrametric_Spaces
+  imports "HOL-Analysis.Elementary_Metric_Spaces"
+begin
+  (*>*)
+
+
+section \<open>Definition\<close>
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, NONE)\<close>
+  \<comment> \<open>To be able to use const\<open>dist\<close> out of the class\<open>metric_space\<close> class.\<close>
+
+class ultrametric_space = uniformity_dist + open_uniformity +
+  assumes dist_eq_0_iff [simp]: \<open>dist x y = 0 \<longleftrightarrow> x = y\<close>
+    and ultrametric_dist_triangle2: \<open>dist x y \<le> max (dist x z) (dist y z)\<close>
+begin
+
+subclass metric_space
+proof (unfold_locales)
+  show \<open>dist x y = 0 \<longleftrightarrow> x = y\<close> for x y by simp
+next
+  show \<open>dist x y \<le> dist x z + dist y z\<close> for x y z
+    by (rule order_trans[OF ultrametric_dist_triangle2, of _ z], simp)
+      (metis local.dist_eq_0_iff ultrametric_dist_triangle2 max.idem)
+qed
+
+end
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, SOME typ\<open>'a :: metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
+  \<comment> \<open>Back to normal.\<close>
+
+
+class complete_ultrametric_space = ultrametric_space +
+  assumes Cauchy_convergent: \<open>Cauchy X \<Longrightarrow> convergent X\<close>
+begin
+
+subclass complete_space by unfold_locales (fact Cauchy_convergent)
+
+end
+
+
+
+section \<open>Properties on Balls\<close>
+
+text \<open>In ultrametric space, balls satisfy very strong properties.\<close>
+
+context ultrametric_space begin
+
+lemma ultrametric_dist_triangle: \<open>dist x z \<le> max (dist x y) (dist y z)\<close>
+  using ultrametric_dist_triangle2 [of x z y] by (simp add: dist_commute)
+
+lemma ultrametric_dist_triangle3: \<open>dist x y \<le> max (dist a x) (dist a y)\<close>
+  using ultrametric_dist_triangle2 [of x y a] by (simp add: dist_commute)
+
+end
+
+
+
+subsection \<open>Balls are centered everywhere\<close>
+
+context fixes x :: \<open>'a :: ultrametric_space\<close> begin
+
+text \<open>The best way to do this would be to work in the context locale\<open>ultrametric_space\<close>.
+      Unfortunately, const\<open>ball\<close>, const\<open>cball\<close>, etc. are not defined inside
+      the context locale\<open>metric_space\<close> but through a sort constraint.\<close>
+
+lemma ultrametric_every_point_of_ball_is_centre :
+  \<open>ball y r = ball x r\<close> if \<open>y \<in> ball x r\<close> 
+proof (unfold set_eq_iff, rule allI)
+  from \<open>y \<in> ball x r\<close> have * : \<open>dist x y < r\<close> by simp
+  show \<open>z \<in> ball y r \<longleftrightarrow> z \<in> ball x r\<close> for z
+    using ultrametric_dist_triangle[of x z y]
+      "*" ultrametric_dist_triangle[of y z x]
+    by (intro iffI) (simp_all add: dist_commute)
+qed
+
+
+lemma ultrametric_every_point_of_cball_is_centre :
+  \<open>cball y r = cball x r\<close> if \<open>y \<in> cball x r\<close>
+proof (unfold set_eq_iff, rule allI)
+  from \<open>y \<in> cball x r\<close> have * : \<open>dist x y \<le> r\<close> by simp
+  show \<open>z \<in> cball y r \<longleftrightarrow> z \<in> cball x r\<close> for z
+    using ultrametric_dist_triangle[of x z y]
+      "*" ultrametric_dist_triangle[of y z x]
+    by (intro iffI) (simp_all add: dist_commute)
+qed
+
+end
+
+
+
+subsection \<open>Balls are ``clopen''\<close>
+
+text \<open>Balls are both open and closed.\<close>
+
+context fixes x :: \<open>'a :: ultrametric_space\<close> begin
+
+lemma ultrametric_open_cball [intro, simp] : \<open>open (cball x r)\<close> if \<open>0 < r\<close>
+proof (rule openI) 
+  fix y assume \<open>y \<in> cball x r\<close>
+  hence \<open>cball y r = cball x r\<close>
+    by (rule ultrametric_every_point_of_cball_is_centre)
+  hence \<open>ball y (r / 2) \<subseteq> cball x r\<close>
+    by (metis ball_subset_cball cball_divide_subset_numeral subset_trans)
+  moreover have \<open>0 < r / 2\<close> by (simp add: \<open>0 < r\<close>)
+  ultimately show \<open>\<exists>e>0. ball y e \<subseteq> cball x r\<close> by blast
+qed
+
+lemma \<open>closed (cball y r)\<close> by (fact closed_cball)
+
+
+lemma ultrametric_closed_ball [intro, simp]: \<open>closed (ball x r)\<close> if \<open>0 \<le> r\<close>
+proof (cases \<open>r = 0\<close>)
+  show \<open>r = 0 \<Longrightarrow> closed (ball x r)\<close> by simp
+next
+  assume \<open>r \<noteq> 0\<close>
+  with \<open>0 \<le> r\<close> have \<open>0 < r\<close> by simp
+  show \<open>closed (ball x r)\<close>
+  proof (unfold closed_def)
+    have \<open>- ball x r = (\<Union>z \<in> - ball x r. ball z r)\<close>
+    proof (intro subset_antisym subsetI)
+      from \<open>0 < r\<close> show \<open>z \<in> - ball x r \<Longrightarrow> z \<in> (\<Union>z\<in>- ball x r. ball z r)\<close> for z
+        by (meson UN_iff centre_in_ball)
+    next
+      show \<open>z \<in> (\<Union>z \<in> - ball x r. ball z r) \<Longrightarrow> z \<in> - ball x r\<close> for z
+        by simp (metis ComplD dist_commute mem_ball
+            ultrametric_every_point_of_ball_is_centre)
+    qed
+    show \<open>open (- ball x r)\<close> by (subst \<open>?this\<close>, rule open_Union, simp)
+  qed
+qed
+
+lemma ultrametric_open_sphere [intro, simp] : \<open>0 < r \<Longrightarrow> open (sphere x r)\<close>
+  by (fold cball_diff_eq_sphere) (simp add: open_Diff order_le_less)
+
+lemma closed_sphere [intro, simp] : \<open>closed (sphere y r)\<close>
+  by (metis open_ball cball_diff_eq_sphere closed_Diff closed_cball)
+
+end
+
+
+
+subsection \<open>Balls are disjoint or contained\<close>
+
+context fixes x :: \<open>'a :: ultrametric_space\<close> begin
+
+lemma ultrametric_ball_ball_disjoint_or_subset:
+  \<open>ball x r \<inter> ball y s = {} \<or> ball x r \<subseteq> ball y s \<or>
+   ball y s \<subseteq> ball x r\<close>
+proof (unfold disj_imp, intro impI)
+  assume \<open>ball x r \<inter> ball y s \<noteq> {}\<close> \<open>\<not> ball x r \<subseteq> ball y s\<close>
+  from \<open>ball x r \<inter> ball y s \<noteq> {}\<close>
+  obtain z where \<open>z \<in> ball x r\<close> and \<open>z \<in> ball y s\<close> by blast
+  with ultrametric_every_point_of_ball_is_centre
+  have \<open>ball x r = ball z r\<close> and \<open>ball y s = ball z s\<close> by auto
+  with \<open>\<not> ball x r \<subseteq> ball y s\<close> have \<open>s < r\<close> by auto
+  with \<open>ball x r = ball z r\<close> and \<open>ball y s = ball z s\<close>
+  show \<open>ball y s \<subseteq> ball x r\<close> by auto
+qed
+
+lemma ultrametric_ball_cball_disjoint_or_subset:
+  \<open>ball x r \<inter> cball y s = {} \<or> ball x r \<subseteq> cball y s \<or>
+   cball y s \<subseteq> ball x r\<close>
+proof (unfold disj_imp, intro impI)
+  assume \<open>ball x r \<inter> cball y s \<noteq> {}\<close> \<open>\<not> ball x r \<subseteq> cball y s\<close>
+  from \<open>ball x r \<inter> cball y s \<noteq> {}\<close>
+  obtain z where \<open>z \<in> ball x r\<close> and \<open>z \<in> cball y s\<close> by blast
+  with ultrametric_every_point_of_ball_is_centre
+    ultrametric_every_point_of_cball_is_centre
+  have \<open>ball x r = ball z r\<close> \<open>cball y s = cball z s\<close> by blast+
+  with \<open>\<not> ball x r \<subseteq> cball y s\<close> have \<open>s < r\<close> by auto
+  with \<open>ball x r = ball z r\<close> and \<open>cball y s = cball z s\<close>
+  show \<open>cball y s \<subseteq> ball x r\<close> by auto
+qed
+
+corollary ultrametric_cball_ball_disjoint_or_subset:
+  \<open>cball x r \<inter> ball y s = {} \<or> cball x r \<subseteq> ball y s \<or>
+   ball y s \<subseteq> cball x r\<close>
+  using Elementary_Ultrametric_Spaces.ultrametric_ball_cball_disjoint_or_subset by blast
+
+lemma ultrametric_cball_cball_disjoint_or_subset:
+  \<open>cball x r \<inter> cball y s = {} \<or> cball x r \<subseteq> cball y s \<or>
+   cball y s \<subseteq> cball x r\<close>
+proof (unfold disj_imp, intro impI)
+  assume \<open>cball x r \<inter> cball y s \<noteq> {}\<close> \<open>\<not> cball x r \<subseteq> cball y s\<close>
+  from \<open>cball x r \<inter> cball y s \<noteq> {}\<close>
+  obtain z where \<open>z \<in> cball x r\<close> and \<open>z \<in> cball y s\<close> by blast
+  with ultrametric_every_point_of_cball_is_centre
+  have \<open>cball x r = cball z r\<close> \<open>cball y s = cball z s\<close> by auto
+  with \<open>\<not> cball x r \<subseteq> cball y s\<close> have \<open>s < r\<close> by auto
+  with \<open>cball x r = cball z r\<close> and \<open>cball y s = cball z s\<close>
+  show \<open>cball y s \<subseteq> cball x r\<close> by auto
+qed
+
+end
+
+
+
+subsection \<open>Distance to a Ball\<close>
+
+context fixes a :: \<open>'a :: ultrametric_space\<close> begin
+
+lemma ultrametric_equal_distance_to_ball:
+  \<open>dist a y = dist a z\<close> if \<open>a \<notin> ball x r\<close> \<open>y \<in> ball x r\<close> \<open>z \<in> ball x r\<close>
+proof (rule order_antisym)
+  show \<open>dist a y \<le> dist a z\<close>
+    by (rule order_trans[OF ultrametric_dist_triangle[of a y z]], simp)
+      (metis dist_commute dual_order.strict_trans2 linorder_linear
+        mem_ball that ultrametric_every_point_of_ball_is_centre)
+next
+  show \<open>dist a z \<le> dist a y\<close>
+    by (rule order_trans[OF ultrametric_dist_triangle[of a z y]], simp)
+      (metis dist_commute dual_order.strict_trans2 linorder_linear
+        mem_ball that ultrametric_every_point_of_ball_is_centre)
+qed
+
+
+lemma ultrametric_equal_distance_to_cball:
+  \<open>dist a y = dist a z\<close> if \<open>a \<notin> cball x r\<close> \<open>y \<in> cball x r\<close> \<open>z \<in> cball x r\<close>
+proof (rule order_antisym)
+  show \<open>dist a y \<le> dist a z\<close>
+    by (rule order_trans[OF ultrametric_dist_triangle[of a y z]], simp)
+      (metis dist_commute dual_order.trans linorder_linear mem_cball
+        that ultrametric_every_point_of_cball_is_centre)
+next
+  show \<open>dist a z \<le> dist a y\<close>
+    by (rule order_trans[OF ultrametric_dist_triangle[of a z y]], simp)
+      (metis dist_commute dual_order.trans linorder_linear mem_cball
+        that ultrametric_every_point_of_cball_is_centre)
+qed
+
+end
+
+
+context fixes x :: \<open>'a :: ultrametric_space\<close> begin
+
+lemma ultrametric_equal_distance_between_ball_ball:
+  \<open>ball x r \<inter> ball y s = {} \<Longrightarrow>
+   \<exists>d. \<forall>a \<in> ball x r. \<forall>b \<in> ball y s. dist a b = d\<close>
+  by (metis disjoint_iff dist_commute ultrametric_equal_distance_to_ball)          
+
+lemma ultrametric_equal_distance_between_ball_cball:
+  \<open>ball x r \<inter> cball y s = {} \<Longrightarrow>
+   \<exists>d. \<forall>a \<in> ball x r. \<forall>b \<in> cball y s. dist a b = d\<close>
+  by (metis disjoint_iff dist_commute ultrametric_equal_distance_to_ball
+      ultrametric_equal_distance_to_cball)
+
+lemma ultrametric_equal_distance_between_cball_ball:
+  \<open>cball x r \<inter> ball y s = {} \<Longrightarrow>
+   \<exists>d. \<forall>a \<in> cball x r. \<forall>b \<in> ball y s. dist a b = d\<close>
+  by (metis disjoint_iff_not_equal dist_commute ultrametric_equal_distance_to_ball
+      ultrametric_equal_distance_to_cball)
+
+lemma ultrametric_equal_distance_between_cball_cball:
+  \<open>cball x r \<inter> cball y s = {} \<Longrightarrow>
+   \<exists>d. \<forall>a \<in> cball x r. \<forall>b \<in> cball y s. dist a b = d\<close>
+  by (metis disjoint_iff dist_commute ultrametric_equal_distance_to_cball)
+
+end
+
+
+
+section \<open>Additional Properties\<close>
+
+text \<open>Here are a few other interesting properties.\<close>
+
+subsection \<open>Cauchy Sequences\<close>
+
+lemma (in ultrametric_space) ultrametric_dist_triangle_generalized:
+  \<open>n < m \<Longrightarrow> dist (\<sigma> n) (\<sigma> m) \<le> (MAX l \<in> {n..m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close>
+proof (induct m)
+  show \<open>n < 0 \<Longrightarrow> dist (\<sigma> n) (\<sigma> 0) \<le> (MAX l\<in>{n..0 - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close> by simp
+next
+  case (Suc m)
+  show \<open>dist (\<sigma> n) (\<sigma> (Suc m)) \<le> (MAX l\<in>{n..Suc m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close>
+  proof (cases \<open>n = m\<close>)
+    show \<open>n = m \<Longrightarrow> dist (\<sigma> n) (\<sigma> (Suc m)) \<le> (MAX l\<in>{n..Suc m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close>
+      by simp
+  next
+    assume \<open>n \<noteq> m\<close>
+    with \<open>n < Suc m\<close> obtain m' where \<open>m = Suc m'\<close> \<open>n \<le> m'\<close> 
+      by (metis le_add1 less_Suc_eq less_imp_Suc_add)
+    have \<open>{n..Suc m'} = {n..m-1} \<union> {m}\<close>
+      by (simp add: \<open>m = Suc m'\<close> \<open>n \<le> m'\<close> atLeastAtMostSuc_conv le_Suc_eq)
+    have \<open>dist (\<sigma> n) (\<sigma> (Suc m)) \<le> max (dist (\<sigma> n) (\<sigma> m)) (dist (\<sigma> m) (\<sigma> (Suc m)))\<close>
+      by (simp add: ultrametric_dist_triangle)
+    also have \<open>\<dots> \<le> max ((MAX l\<in>{n..m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))) (dist (\<sigma> m) (\<sigma> (Suc m)))\<close>
+      using Suc.hyps Suc.prems \<open>n \<noteq> m\<close> by linarith
+    also have \<open>\<dots> = (MAX l\<in>{n..Suc m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close>
+      by (subst Max_Un[of _ \<open>((\<lambda>l. dist (\<sigma> l) (\<sigma> (Suc l))) ` {m})\<close>, simplified, symmetric])
+        (simp_all add: \<open>m = Suc m'\<close> \<open>n \<le> m'\<close> \<open>{n..Suc m'} = {n..m-1} \<union> {m}\<close>)
+    finally show \<open>dist (\<sigma> n) (\<sigma> (Suc m)) \<le> (MAX l\<in>{n..Suc m - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close> .
+  qed
+qed
+
+
+lemma (in ultrametric_space) ultrametric_Cauchy_iff:
+  \<open>Cauchy \<sigma> \<longleftrightarrow> (\<lambda>n. dist (\<sigma> (Suc n)) (\<sigma> n)) \<longlonglongrightarrow> 0\<close>
+proof (rule iffI)
+  assume \<open>Cauchy \<sigma>\<close>
+  show \<open>(\<lambda>n. dist (\<sigma> (Suc n)) (\<sigma> n)) \<longlonglongrightarrow> 0\<close>
+  proof (unfold lim_sequentially, intro allI impI)
+    fix \<epsilon> :: real
+    assume \<open>0 < \<epsilon>\<close>
+    from \<open>Cauchy \<sigma>\<close>[unfolded Cauchy_altdef, rule_format, OF \<open>0 < \<epsilon>\<close>]
+    show \<open>\<exists>N. \<forall>n\<ge>N. dist (dist (\<sigma> (Suc n)) (\<sigma> n)) 0 < \<epsilon>\<close>
+      by (auto simp add: dist_commute)
+  qed
+next
+  assume convergent : \<open>(\<lambda>n. dist (\<sigma> (Suc n)) (\<sigma> n)) \<longlonglongrightarrow> 0\<close>
+  show \<open>Cauchy \<sigma>\<close>
+  proof (unfold Cauchy_altdef2, intro allI impI)
+    fix \<epsilon> :: real
+    assume \<open>0 < \<epsilon>\<close>
+    from convergent[unfolded lim_sequentially, rule_format, OF \<open>0 < \<epsilon>\<close>]
+    obtain N where * : \<open>N \<le> n \<Longrightarrow> dist (\<sigma> (Suc n)) (\<sigma> n) < \<epsilon>\<close> for n
+      by (simp add: dist_real_def) blast
+
+    have \<open>N < n \<Longrightarrow> dist (\<sigma> n) (\<sigma> N) < \<epsilon>\<close> for n
+    proof (subst dist_commute, rule le_less_trans)
+      show \<open>N < n \<Longrightarrow> dist (\<sigma> N) (\<sigma> n) \<le> (MAX l\<in>{N..n - 1}. dist (\<sigma> l) (\<sigma> (Suc l)))\<close>
+        by (fact ultrametric_dist_triangle_generalized)
+    next
+      show \<open>N < n \<Longrightarrow> (MAX l\<in>{N..n - 1}. dist (\<sigma> l) (\<sigma> (Suc l))) < \<epsilon>\<close>
+        by simp (metis "*" atLeastAtMost_iff dist_commute)
+    qed
+    with \<open>0 < \<epsilon>\<close> have \<open>N \<le> n \<Longrightarrow> dist (\<sigma> n) (\<sigma> N) < \<epsilon>\<close> for n
+      by (cases \<open>N = n\<close>) simp_all 
+    thus \<open>\<exists>N. \<forall>n\<ge>N. dist (\<sigma> n) (\<sigma> N) < \<epsilon>\<close> by blast
+  qed
+qed
+
+
+
+subsection \<open>Isosceles Triangle Principle\<close>
+
+lemma (in ultrametric_space) ultrametric_isosceles_triangle_principle : 
+  \<open>dist x z = max (dist x y) (dist y z)\<close> if \<open>dist x y \<noteq> dist y z\<close>
+proof (rule order_antisym)
+  show \<open>dist x z \<le> max (dist x y) (dist y z)\<close>
+    by (fact ultrametric_dist_triangle)
+next
+  from \<open>dist x y \<noteq> dist y z\<close> linorder_less_linear
+  have \<open>dist x y < dist y z \<or> dist y z < dist x y\<close> by blast
+  with ultrametric_dist_triangle[of y z x]
+    ultrametric_dist_triangle[of x y z]
+  show \<open>max (dist x y) (dist y z) \<le> dist x z\<close>
+    by (elim disjE) (simp_all add: dist_commute)
+qed
+
+
+
+subsection \<open>Distance to a convergent Sequence\<close>
+
+lemma ultrametric_dist_to_convergent_sequence_is_eventually_const :
+  fixes \<sigma> :: \<open>nat \<Rightarrow> 'a :: ultrametric_space\<close>
+  assumes \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close> and \<open>x \<noteq> \<Sigma>\<close>
+  shows \<open>\<exists>N. \<forall>n\<ge>N. dist (\<sigma> n) x = dist \<Sigma> x\<close>
+proof -
+  from \<open>x \<noteq> \<Sigma>\<close> have \<open>0 < dist x \<Sigma>\<close> by simp
+  then obtain \<epsilon> where \<open>0 < \<epsilon>\<close> \<open>ball x \<epsilon> \<inter> ball \<Sigma> \<epsilon> = {}\<close>
+    by (metis centre_in_ball disjoint_iff mem_ball order_less_le
+        ultrametric_every_point_of_ball_is_centre)
+
+  from \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close> \<open>0 < \<epsilon>\<close> obtain N where \<open>N \<le> n \<Longrightarrow> \<sigma> n \<in> ball \<Sigma> \<epsilon>\<close> for n
+    by (auto simp add: dist_commute lim_sequentially)
+  with \<open>0 < \<epsilon>\<close> \<open>ball x \<epsilon> \<inter> ball \<Sigma> \<epsilon> = {}\<close> have \<open>N \<le> n \<Longrightarrow> dist (\<sigma> n) x = dist \<Sigma> x\<close> for n
+    by (metis centre_in_ball dist_commute ultrametric_equal_distance_between_ball_ball)
+  thus \<open>\<exists>N. \<forall>n\<ge>N. dist (\<sigma> n) x = dist \<Sigma> x\<close> by blast
+qed
+
+
+
+subsection \<open>Diameter\<close>
+
+lemma ultrametric_diameter : \<open>diameter S = (SUP y \<in> S. dist x y)\<close>
+  if \<open>bounded S\<close> and \<open>x \<in> S\<close> for x :: \<open>'a :: ultrametric_space\<close>
+proof -
+  from \<open>x \<in> S\<close> have \<open>S \<noteq> {}\<close> by blast
+  show \<open>diameter S = (SUP y \<in> S. dist x y)\<close>
+  proof (rule order_antisym)
+    from diameter_bounded_bound[OF \<open>bounded S\<close> \<open>x \<in> S\<close>]
+    have \<open>y \<in> S \<Longrightarrow> dist x y \<le> diameter S\<close> for y .
+    thus \<open>(SUP y \<in> S. dist x y) \<le> diameter S\<close>
+      by (rule cSUP_least[OF \<open>S \<noteq> {}\<close>])
+  next
+    have \<open>bdd_above (dist x ` S)\<close>
+      by (meson bdd_above.I2 bounded_any_center \<open>bounded S\<close>)
+    have \<open>y \<in> S \<Longrightarrow> z \<in> S \<Longrightarrow> dist y z \<le> max (dist x y) (dist x z)\<close> for y z
+      by (metis dist_commute ultrametric_dist_triangle)
+    also have \<open>y \<in> S \<Longrightarrow> z \<in> S \<Longrightarrow>
+               max (dist x y) (dist x z) \<le> (SUP y \<in> S. dist x y)\<close> for y z
+      by (cases \<open>dist x y \<le> dist x z\<close>)
+        (simp_all add: cSUP_upper2[OF \<open>bdd_above (dist x ` S)\<close>])
+    finally have * : \<open>y \<in> S \<Longrightarrow> z \<in> S \<Longrightarrow> dist y z \<le> (SUP y \<in> S. dist x y)\<close> for y z .
+    have \<open>(SUP (y, z) \<in> S \<times> S. dist y z) \<le> (SUP y \<in> S. dist x y)\<close>
+      by (rule cSUP_least) (use "*" in \<open>auto simp add: \<open>S \<noteq> {}\<close>\<close>)
+    thus \<open>diameter S \<le> Sup (dist x ` S)\<close>
+      by (simp add: diameter_def \<open>S \<noteq> {}\<close>)
+  qed
+qed
+
+
+
+subsection \<open>Totally disconnected\<close>
+
+lemma ultrametric_totally_disconnected :
+  \<open>\<exists>x. S = {x}\<close> if \<open>S \<noteq> {}\<close> \<open>connected S\<close>
+for S :: \<open>'a :: ultrametric_space set\<close>
+proof -
+  from \<open>S \<noteq> {}\<close> obtain x where \<open>x \<in> S\<close> by blast
+  have \<open>ball x r \<inter> S = {} \<or> - ball x r \<inter> S = {}\<close> if \<open>0 < r\<close> for r
+    by (rule \<open>connected S\<close>[unfolded connected_def, simplified, rule_format])
+      (simp_all, use order_less_imp_le that ultrametric_closed_ball in blast)
+  with \<open>x \<in> S\<close> have \<open>0 < r \<Longrightarrow> - ball x r \<inter> S = {}\<close> for r
+    by (metis centre_in_ball disjoint_iff)
+  hence \<open>0 < r \<Longrightarrow> y \<in> S \<Longrightarrow> dist x y < r\<close> for r y
+    by (auto simp add: disjoint_iff)
+  hence \<open>y \<in> S \<Longrightarrow> dist x y = 0\<close> for y
+    by (metis dist_self order_less_irrefl zero_less_dist_iff)
+  hence \<open>y \<in> S \<Longrightarrow> y = x\<close> for y by simp
+  with \<open>x \<in> S\<close> show \<open>\<exists>x. S = {x}\<close> by blast
+qed
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Elementary_Ultrametric_Spaces/ROOT
--- /dev/null
+++ thys/Elementary_Ultrametric_Spaces/ROOT
@@ -0,0 +1,10 @@
+chapter AFP
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF
+   Author: Benjamin Puyobro, Université Paris-Saclay, IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+   Author: Burkhart Wolff, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+session Elementary_Ultrametric_Spaces = "HOL-Analysis" +
+  options [timeout = 300]
+  theories Elementary_Ultrametric_Spaces (global)
+  document_files "root.tex"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Elementary_Ultrametric_Spaces/document/root.tex
--- /dev/null
+++ thys/Elementary_Ultrametric_Spaces/document/root.tex
@@ -0,0 +1,87 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+
+\usepackage{amsmath}
+
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{Definition and Elementary Properties of Ultrametric Spaces}
+\author{
+Benoît Ballenghien\footnote{Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF\\\url{https://orcid.org/0009-0000-4941-187X}} \and
+Benjamin Puyobro\footnote{Université Paris-Saclay, IRT SystemX, CNRS, ENS Paris-Saclay, LMF\\\url{https://orcid.org/0009-0006-0294-9043}\\
+This work has been supported by the French government under the ``France 2030''
+program, as part of the SystemX Technological Research Institute within the CVH project.} \and
+Burkhart Wolff\footnote{Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF\\\url{https://orcid.org/0000-0002-9648-7663}}}
+
+\maketitle
+
+\abstract{
+An ultrametric space is a metric space in which the triangle inequality is strengthened
+by using the maximum instead of the sum. More formally, such a space is equipped with
+a real-valued application $dist$, called distance, verifying the four following conditions.
+\begin{align*}
+    & dist \ x \ y \ge 0\\
+    & dist \ x \ y = dist \ y \ x\\
+    & dist \ x \ y = 0 \longleftrightarrow x = y\\
+    & dist \ x \ z \le max \ (dist \ x \ y) \ (dist \ y \ z)
+\end{align*}
+In this entry, we present an elementary formalization of these spaces relying on axiomatic type classes.
+The connection with standard metric spaces is obtained through a subclass relationship,
+and fundamental properties of ultrametric spaces are formally established.
+}
+
+\newpage
+
+\tableofcontents
+
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Farey_Sequences/Farey_Ford.thy
--- /dev/null
+++ thys/Farey_Sequences/Farey_Ford.thy
@@ -0,0 +1,1202 @@
+section \<open>Farey Sequences and Ford Circles\<close>
+
+theory Farey_Ford
+  imports "HOL-Analysis.Analysis" "HOL-Number_Theory.Totient" "HOL-Library.Sublist"
+
+begin
+
+subsection \<open>Farey sequences\<close>
+
+lemma sorted_two_sublist:
+  fixes x:: "'a::order"
+  assumes sorted: "sorted_wrt (<) l"
+  shows "sublist [x, y] l \<longleftrightarrow> x < y \<and> x \<in> set l \<and> y \<in> set l \<and> (\<forall>z \<in> set l. z \<le> x \<or> z \<ge> y)"
+proof (cases "x < y \<and> x \<in> set l \<and> y \<in> set l")
+  case True
+  then obtain xs us where us: "l = xs @ [x] @ us"
+    by (metis append_Cons append_Nil in_set_conv_decomp_first)
+  with True assms have "y \<in> set us"
+    by (fastforce simp add: sorted_wrt_append)
+  then obtain ys zs where yz: "l = xs @ [x] @ ys @ [y] @ zs"
+    by (metis split_list us append_Cons append_Nil)
+  have "sublist [x, y] l \<longleftrightarrow> ys = []"
+    using sorted yz 
+    apply (simp add: sublist_def sorted_wrt_append)
+    by (metis (mono_tags, opaque_lifting) append_Cons_eq_iff append_Nil assms sorted_wrt.simps(2)
+        sorted_wrt_append less_irrefl)
+  also have "\<dots> = (\<forall>z \<in> set l. z \<le> x \<or> z \<ge> y)"
+    using sorted yz
+    apply (simp add: sublist_def sorted_wrt_append)
+    by (metis Un_iff empty_iff less_le_not_le list.exhaust list.set(1) list.set_intros(1))
+  finally show ?thesis
+    using True by blast 
+next
+  case False
+  then show ?thesis
+    by (metis list.set_intros(1) set_subset_Cons sorted_wrt.simps(2) sorted_wrt_append sublist_def
+        set_mono_sublist sorted subset_iff)
+qed
+
+(* not clear to do with these transp lemmas, are they of general use?*)
+lemma transp_add1_int:
+  assumes "\<And>n::int. R (f n) (f (1 + n))"
+    and "n < n'"
+    and "transp R"
+  shows "R (f n) (f n')"
+proof -
+  have "R (f n) (f (1 + n + int k))" for k
+    by (induction k) (use assms in \<open>auto elim!: transpE\<close>)
+  then show ?thesis
+    by (metis add.commute assms(2) zle_iff_zadd zless_imp_add1_zle)
+qed
+
+lemma refl_transp_add1_int:
+  assumes "\<And>n::int. R (f n) (f (1 + n))"
+      and "n \<le> n'"
+      and "reflp R" "transp R"
+    shows "R (f n) (f n')"
+  by (metis assms order_le_less reflpE transp_add1_int)
+
+lemma transp_Suc:
+  assumes "\<And>n. R (f n) (f (Suc n))"
+      and "n < n'"
+      and "transp R"
+    shows "R (f n) (f n')"
+proof -
+  have "R (f n) (f (1 + n + k))" for k
+  by (induction k) (use assms in \<open>auto elim!: transpE\<close>)
+  then show ?thesis
+    by (metis add_Suc_right add_Suc_shift assms(2) less_natE plus_1_eq_Suc)
+qed
+
+lemma refl_transp_Suc:
+  assumes "\<And>n. R (f n) (f (Suc n))"
+      and "n \<le> n'"
+      and "reflp R" "transp R"
+    shows "R (f n) (f n')"
+  by (metis assms dual_order.order_iff_strict reflpE transp_Suc)
+
+lemma coprime_unimodular_int:
+  fixes a b::int
+  assumes "coprime a b" "a>1" "b>1"
+  obtains x y where "a*x - b*y = 1" "0 < x" "x < b" "0 < y" "y < a"
+proof -
+  obtain u v where 1: "a * u + b * v = 1"
+    by (metis \<open>coprime a b\<close> cong_iff_lin coprime_iff_invertible_int)
+  define k where "k \<equiv> u div b"
+  define x where "x \<equiv> u - k*b"
+  define y where "y \<equiv> -(v + k*a)"
+  show thesis
+  proof
+    show *: "a * x - b * y = 1" 
+      using 1 by (simp add: x_def y_def algebra_simps)
+    have "u \<noteq> k*b" "b>0"
+      using assms "*"  by (auto simp: k_def x_def y_def zmult_eq_neg1_iff) 
+    moreover have "u - k*b \<ge> 0"
+      by (simp add: k_def \<open>b>0\<close> minus_div_mult_eq_mod)
+    ultimately show "x > 0"
+      by (fastforce simp: x_def)
+    show "x < b"
+      by (simp add: \<open>0 < b\<close> k_def minus_div_mult_eq_mod x_def)
+    have "a*x > 1"
+      by (metis \<open>0 < x\<close> \<open>a>1\<close> int_one_le_iff_zero_less less_1_mult linorder_not_less
+          mult.right_neutral nle_le)
+    then have "\<not> b * y \<le> 0"
+      using "*" by linarith
+    then show "y > 0"
+      by (simp add: \<open>0 < b\<close> mult_less_0_iff order_le_less)
+    show "y < a"
+      using "*" \<open>0 < x\<close> \<open>x < b\<close>
+      by (smt (verit, best) \<open>a>1\<close> mult.commute mult_less_le_imp_less)
+  qed
+qed
+
+subsection \<open>Farey Fractions\<close>
+
+type_synonym farey = rat
+
+definition num_farey :: "farey \<Rightarrow> int" 
+  where "num_farey \<equiv> \<lambda>x. fst (quotient_of x)" 
+
+definition denom_farey :: "farey \<Rightarrow> int"
+  where "denom_farey \<equiv> \<lambda>x. snd (quotient_of x)" 
+
+definition farey :: "[int,int] \<Rightarrow> farey" 
+  where "farey \<equiv> \<lambda>a b. max 0 (min 1 (Fract a b))"
+
+lemma farey01 [simp]: "0 \<le> farey a b" "farey a b \<le> 1"
+  by (auto simp: min_def max_def farey_def)
+
+lemma farey_0 [simp]: "farey 0 n = 0"
+  by (simp add: farey_def rat_number_collapse)
+
+lemma farey_1 [simp]: "farey 1 1 = 1"
+  by (auto simp: farey_def rat_number_expand)
+
+lemma num_farey_nonneg: "x \<in> {0..1} \<Longrightarrow> num_farey x \<ge> 0"
+  by (cases x) (simp add: num_farey_def quotient_of_Fract zero_le_Fract_iff)
+
+lemma num_farey_le_denom: "x \<in> {0..1} \<Longrightarrow> num_farey x \<le> denom_farey x"
+  by (cases x) (simp add: Fract_le_one_iff denom_farey_def num_farey_def quotient_of_Fract)
+
+lemma denom_farey_pos: "denom_farey x > 0"
+  by (simp add: denom_farey_def quotient_of_denom_pos')
+
+lemma coprime_num_denom_farey [intro]: "coprime (num_farey x) (denom_farey x)"
+  by (simp add: denom_farey_def num_farey_def quotient_of_coprime)
+
+lemma rat_of_farey_conv_num_denom:
+  "x = rat_of_int (num_farey x) / rat_of_int (denom_farey x)"
+  by (simp add: denom_farey_def num_farey_def quotient_of_div)
+
+lemma num_denom_farey_eqI:
+  assumes "x = of_int a / of_int b" "b > 0" "coprime a b"
+  shows   "num_farey x = a" "denom_farey x = b"
+  using Fract_of_int_quotient assms quotient_of_Fract
+  by (auto simp: num_farey_def denom_farey_def)
+
+lemma farey_cases [cases type, case_names farey]:
+  assumes "x \<in> {0..1}"
+  obtains a b where "0\<le>a" "a\<le>b" "coprime a b" "x = Fract a b"
+  by (metis Fract_of_int_quotient Rat_cases assms num_denom_farey_eqI num_farey_le_denom num_farey_nonneg)
+
+lemma rat_of_farey: "\<lbrakk>x = of_int a / of_int b; x \<in> {0..1}\<rbrakk> \<Longrightarrow> x = farey a b"
+  by (simp add: Fract_of_int_quotient farey_def max_def)
+
+lemma farey_num_denom_eq [simp]: "x \<in> {0..1} \<Longrightarrow> farey (num_farey x) (denom_farey x) = x"
+  using rat_of_farey rat_of_farey_conv_num_denom by fastforce
+
+lemma farey_eqI:
+  assumes "num_farey x = num_farey y" "denom_farey x = denom_farey y"
+  shows   "x=y"
+  by (metis Rat.of_int_def assms rat_of_farey_conv_num_denom)
+
+lemma
+  assumes "coprime a b" "0\<le>a" "a<b"
+  shows num_farey_eq [simp]: "num_farey (farey a b) = a"
+  and denom_farey_eq [simp]: "denom_farey (farey a b) = b"
+  using Fract_less_one_iff quotient_of_Fract zero_le_Fract_iff
+  using assms num_denom_farey_eqI rat_of_farey by force+
+
+lemma
+  assumes "0 \<le> a" "a \<le> b" "0<b"
+  shows num_farey: "num_farey (farey a b) = a div (gcd a b)" 
+    and denom_farey: "denom_farey (farey a b) = b div (gcd a b)"
+proof -
+  have "0 \<le> Fract a b" "Fract a b \<le> 1"
+    using assms by (auto simp: Fract_le_one_iff zero_le_Fract_iff)
+  with assms show "num_farey (farey a b) = a div (gcd a b)" "denom_farey (farey a b) = b div (gcd a b)"
+    by (auto simp: num_farey_def denom_farey_def farey_def quotient_of_Fract Rat.normalize_def Let_def)
+qed
+
+lemma
+  assumes "coprime a b" "0<b"
+  shows num_farey_Fract [simp]: "num_farey (Fract a b) = a"
+  and denom_farey_Fract [simp]: "denom_farey (Fract a b) = b"
+  using Fract_of_int_quotient assms num_denom_farey_eqI by force+
+
+lemma num_farey_0 [simp]: "num_farey 0 = 0"
+  and denom_farey_0 [simp]: "denom_farey 0 = 1"
+  and num_farey_1 [simp]: "num_farey 1 = 1"
+  and denom_farey_1 [simp]: "denom_farey 1 = 1"
+  by (auto simp: num_farey_def denom_farey_def)
+
+lemma num_farey_neq_denom: "denom_farey x \<noteq> 1 \<Longrightarrow> num_farey x \<noteq> denom_farey x"
+  by (metis denom_farey_0 div_0 div_self num_farey_1 rat_of_farey_conv_num_denom)
+
+lemma num_farey_0_iff [simp]: "num_farey x = 0 \<longleftrightarrow> x = 0"
+  unfolding num_farey_def
+  by (metis div_0 eq_fst_iff of_int_0 quotient_of_div rat_zero_code)
+
+lemma denom_farey_le1_cases:
+  assumes "denom_farey x \<le> 1" "x \<in> {0..1}"
+  shows "x = 0 \<or> x = 1"
+proof -
+  consider "num_farey x = 0" | "num_farey x = 1" "denom_farey x = 1"
+    using assms num_farey_le_denom [of x] num_farey_nonneg [of x] by linarith
+  then show ?thesis
+    by (metis denom_farey_1 farey_eqI num_farey_0_iff num_farey_1)
+qed
+
+definition mediant :: "farey \<Rightarrow> farey \<Rightarrow> farey" where 
+  "mediant \<equiv> \<lambda>x y. Fract (fst (quotient_of x) + fst (quotient_of y)) 
+                         (snd (quotient_of x) + snd (quotient_of y))"
+
+lemma mediant_eq_Fract:
+  "mediant x y = Fract (num_farey x + num_farey y) (denom_farey x + denom_farey y)"
+  by (simp add: denom_farey_def num_farey_def mediant_def)
+
+lemma mediant_eq_farey:
+  assumes "x \<in> {0..1}" "y \<in> {0..1}"
+  shows "mediant x y = farey (num_farey x + num_farey y) (denom_farey x + denom_farey y)"
+proof -
+  have "0 \<le> num_farey x + num_farey y"
+    using assms num_farey_nonneg by auto
+  moreover have "num_farey x + num_farey y \<le> denom_farey x + denom_farey y"
+    by (meson add_mono assms num_farey_le_denom)
+  ultimately show ?thesis
+    by (simp add: add_pos_pos denom_farey_pos Fract_of_int_quotient rat_of_farey mediant_eq_Fract)
+qed
+
+
+definition farey_unimodular :: "farey \<Rightarrow> farey \<Rightarrow> bool" where
+  "farey_unimodular x y \<longleftrightarrow>
+     denom_farey x * num_farey y - num_farey x * denom_farey y = 1"
+
+lemma farey_unimodular_imp_less:
+  assumes "farey_unimodular x y"
+  shows   "x < y"
+  using assms
+  by (auto simp: farey_unimodular_def rat_less_code denom_farey_def num_farey_def)
+
+lemma denom_mediant: "denom_farey (mediant x y) \<le> denom_farey x + denom_farey y"
+  using quotient_of_denom_pos' [of x] quotient_of_denom_pos' [of y]
+  by (simp add: mediant_eq_Fract denom_farey_def num_farey_def quotient_of_Fract normalize_def Let_def int_div_le_self)
+
+lemma unimodular_imp_both_coprime:
+  fixes a:: "'a::{algebraic_semidom,comm_ring_1}"
+  assumes "b*c - a*d = 1" 
+  shows   "coprime a b" "coprime c d"
+  using mult.commute by (metis assms coprimeI dvd_diff dvd_mult2)+
+
+lemma unimodular_imp_coprime:
+  fixes a:: "'a::{algebraic_semidom,comm_ring_1}"
+  assumes "b*c - a*d = 1" 
+  shows   "coprime (a+c) (b+d)"
+proof (rule coprimeI)
+  fix k 
+  assume k: "k dvd (a+c)" "k dvd (b+d)"
+  moreover have "(b+d)*c = (a+c)*d + 1"
+    using assms by (simp add: algebra_simps)
+  ultimately show "is_unit k"
+    by (metis add_diff_cancel_left' dvd_diff dvd_mult2)
+qed
+
+definition fareys :: "int \<Rightarrow> rat list"
+  where "fareys n \<equiv> sorted_list_of_set {x \<in> {0..1}. denom_farey x \<le> n}"
+
+lemma strict_sorted_fareys: "sorted_wrt (<) (fareys n)"
+  by (simp add: fareys_def)
+
+lemma farey_set_UN_farey: "{x \<in> {0..1}. denom_farey x \<le> n} = (\<Union>b \<in> {1..n}. \<Union>a \<in> {0..b}. {farey a b})"
+proof -
+  have "\<exists>b \<in> {1..n}. \<exists>a \<in> {0..b}. x = farey a b"
+    if "denom_farey x \<le> n" "x \<in> {0..1}" for x :: farey
+    unfolding Bex_def
+    using that denom_farey_pos int_one_le_iff_zero_less num_farey_le_denom num_farey_nonneg
+    by fastforce
+  moreover have "\<And>b a::int. 0 \<le> b \<Longrightarrow> b div gcd a b \<le> b"
+    by (metis div_0 int_div_le_self nless_le)
+  ultimately show ?thesis
+    by (auto simp: denom_farey) (meson order_trans)
+qed
+
+lemma farey_set_UN_farey': "{x \<in> {0..1}. denom_farey x \<le> n} = (\<Union>b \<in> {1..n}. \<Union>a \<in> {0..b}. if coprime a b then {farey a b} else {})"
+proof -
+  have "\<exists>aa ba. farey a b = farey aa ba \<and> 0 \<le> aa \<and> aa \<le> ba \<and> 1 \<le> ba \<and> ba \<le> n \<and> coprime aa ba"
+    if "1 \<le> b" and "b \<le> n" and "0 \<le> a" and "a \<le> b" for a b
+  proof -
+    let ?a = "a div gcd a b"
+    let ?b = "b div gcd a b"
+    have "coprime ?a ?b"
+      by (metis div_gcd_coprime not_one_le_zero \<open>b\<ge>1\<close>)
+    moreover have "farey a b = farey ?a ?b"
+      using Fract_coprime farey_def by presburger
+    moreover have "?a \<le> ?b \<and> ?b \<le> n"
+      by (smt (verit, best) gcd_pos_int int_div_le_self that zdiv_mono1)
+    ultimately show ?thesis
+      using that by (metis denom_farey denom_farey_pos div_int_pos_iff gcd_ge_0_int int_one_le_iff_zero_less)
+  qed
+  then show ?thesis
+    unfolding farey_set_UN_farey
+    by (fastforce split: if_splits)
+qed
+
+lemma farey_set_UN_Fract: "{x \<in> {0..1}. denom_farey x \<le> n} = (\<Union>b \<in> {1..n}. \<Union>a \<in> {0..b}. {Fract a b})"
+  unfolding farey_set_UN_farey
+  by (simp add: Fract_of_int_quotient farey_def)
+
+lemma farey_set_UN_Fract': "{x \<in> {0..1}. denom_farey x \<le> n} = (\<Union>b \<in> {1..n}. \<Union>a \<in> {0..b}. if coprime a b then {Fract a b} else {})"
+  unfolding farey_set_UN_farey'
+  by (simp add: Fract_of_int_quotient farey_def)
+
+lemma finite_farey_set: "finite {x \<in> {0..1}. denom_farey x \<le> n}"
+  unfolding farey_set_UN_farey by blast
+
+lemma denom_in_fareys_iff: "x \<in> set (fareys n) \<longleftrightarrow> denom_farey x \<le> int n \<and> x \<in> {0..1}"
+  using finite_farey_set by (auto simp: fareys_def)
+
+lemma denom_fareys_leI: "x \<in> set (fareys n) \<Longrightarrow> denom_farey x \<le> n"
+  using finite_farey_set by (auto simp: fareys_def)
+
+lemma denom_fareys_leD: "\<lbrakk>denom_farey x \<le> int n; x \<in> {0..1}\<rbrakk> \<Longrightarrow> x \<in> set (fareys n)"
+  using denom_in_fareys_iff by blast
+
+lemma fareys_increasing_1: "set (fareys n) \<subseteq> set (fareys (1 + n))"
+  using farey_set_UN_farey by (force simp: fareys_def)
+
+definition fareys_new :: "int \<Rightarrow> rat set" where
+  "fareys_new n \<equiv> {Fract a n| a. coprime a n \<and> a \<in> {0..n}}"
+
+lemma fareys_new_0 [simp]: "fareys_new 0 = {}"
+  by (auto simp: fareys_new_def)
+
+lemma fareys_new_1 [simp]: "fareys_new 1 = {0,1}"
+proof -
+  have "Fract a 1 = 1"
+    if a: "Fract a 1 \<noteq> 0" "0 \<le> a" "a \<le> 1" for a
+    by (metis One_rat_def int_one_le_iff_zero_less nless_le order_antisym_conv
+        rat_number_collapse(1) that) 
+  moreover have "\<exists>a. Fract a 1 = 0 \<and> 0 \<le> a \<and> a \<le> 1"
+    using rat_number_expand(1) by auto
+  moreover have "\<exists>a. Fract a 1 = 1 \<and> 0 \<le> a \<and> a \<le> 1"
+    using One_rat_def by fastforce
+  ultimately show ?thesis
+    by (auto simp: fareys_new_def)
+qed
+
+lemma fareys_new_not01:
+  assumes "n>1"
+  shows "0 \<notin> (fareys_new n)" "1 \<notin> (fareys_new n)"
+  using assms by (simp_all add: Fract_of_int_quotient fareys_new_def)
+
+lemma inj_num_farey: "inj_on num_farey (fareys_new n)"
+proof (cases "n=1")
+  case True
+  then show ?thesis
+    using fareys_new_1 by auto
+next
+  case False
+  then show ?thesis
+  proof -
+    have "Fract a n = Fract a' n"
+      if "coprime a n" "0 \<le> a" "a \<le> n"
+        and "coprime a' n" "0 \<le> a'" "a' \<le> n"
+        and "num_farey (Fract a n) = num_farey (Fract a' n)"
+      for a a'
+    proof -
+      from that
+      obtain "a < n" "a' < n"
+        using False by force+
+      with that show ?thesis
+        by auto
+    qed
+    with False show ?thesis
+      by (auto simp: inj_on_def fareys_new_def)
+  qed
+qed
+ 
+lemma finite_fareys_new [simp]: "finite (fareys_new n)"
+  by (auto simp: fareys_new_def)
+
+lemma card_fareys_new:
+  assumes "n > 1"
+  shows "card (fareys_new (int n)) = totient n"
+proof -
+  have "bij_betw num_farey (fareys_new n) (int ` totatives n)"
+  proof -
+    have "\<exists>a>0. a \<le> n \<and> coprime a n \<and> num_farey x = int a"
+      if x: "x \<in> fareys_new n" for x
+    proof -
+      obtain a where a: "x = Fract a (int n)" "coprime a (int n)" "0 \<le> a" "a \<le> int n"
+        using x by (auto simp: fareys_new_def)
+      then have "a > 0"
+        using assms less_le_not_le by fastforce
+      moreover have "coprime (nat a) n"
+        by (metis a(2,3) coprime_int_iff nat_0_le)
+      ultimately have "num_farey x = int (nat a)"
+        using a num_farey by auto
+      then show ?thesis
+        using \<open>0 < a\<close> \<open>coprime (nat a) n\<close> a(4) nat_le_iff zero_less_nat_eq by blast
+    qed
+    moreover have "\<exists>x\<in>fareys_new n. int a = num_farey x"
+      if "0 < a" "a \<le> n" "coprime a n" for a 
+    proof -
+      have \<section>: "coprime (int a) (int n)" "0 \<le> (int a)" "(int a) \<le> int n"
+        using that by auto
+      then have "Fract (int a) (int n) = Fract (int a) (int n)"
+        using Fract_of_int_quotient assms rat_of_farey by auto
+      with \<section> have "Fract (int a) (int n) \<in> fareys_new n"
+        by (auto simp: fareys_new_def)
+      then have "int a = num_farey (Fract (int a) n)"
+        using \<open>coprime (int a) (int n)\<close> assms by auto
+      then show ?thesis
+        using \<open>Fract (int a) (int n) \<in> fareys_new (int n)\<close> by blast
+    qed
+    ultimately show ?thesis
+      by (auto simp add: totatives_def bij_betw_def inj_num_farey comp_inj_on image_iff)
+  qed
+  then show ?thesis
+    unfolding totient_def by (metis bij_betw_same_card bij_betw_of_nat)
+qed
+
+lemma disjoint_fareys_plus1: 
+  assumes "n > 0"
+  shows "disjnt (set (fareys n)) (fareys_new (1 + n))"
+proof -
+  have False
+    if \<section>: "0 \<le> a" "a \<le> 1 + n" "coprime a (1 + n)"
+      "1 \<le> d" "d \<le> n" "Fract a (1 + n) = Fract c d" "0 \<le> c" "c \<le> d" "coprime c d"
+    for a c d
+  proof (cases "c<d")
+    case True
+    have alen: "a \<le> n"
+      using nle_le that by fastforce
+    have "d = denom_farey (Fract c d)"
+      using that by force
+    also have "\<dots> = 1 + n"
+      using denom_farey_Fract that by fastforce
+    finally show False
+      using that(5) by fastforce
+  next
+    case False
+    with \<open>c \<le> d\<close> have "c=d" by auto
+    with that have "d=1" by force
+    with that have "1 + n = 1"
+      by (metis One_rat_def \<open>c = d\<close> assms denom_farey_1 denom_farey_Fract le_imp_0_less
+          order_less_le)
+    then show ?thesis
+      using \<open>c = d\<close> \<section> assms by auto
+  qed
+  then show ?thesis
+  unfolding fareys_def farey_set_UN_Fract' fareys_new_def disjnt_iff
+    by auto
+qed
+
+
+lemma set_fareys_plus1: "set (fareys (1 + n)) = set (fareys n) \<union> fareys_new (1 + n)"
+proof -
+  have "\<exists>b\<ge>1. b \<le> n \<and> (\<exists>a\<ge>0. a \<le> b \<and> coprime a b \<and> Fract c d = Fract a b)"
+    if "Fract c d \<notin> fareys_new (1 + n)"
+      and "coprime c d" "1 \<le> d" "d \<le> 1 + n" "0 \<le> c" "c \<le> d"
+    for c d
+  proof (cases "d = 1 + n")
+    case True
+    with that show ?thesis
+      by (auto simp: fareys_new_def)
+  qed (use that in auto)
+  moreover have "\<exists>d\<ge>1. d \<le> 1 + n \<and> (\<exists>c\<ge>0. c \<le> d \<and> coprime c d \<and> x = Fract c d)"
+    if "x \<in> fareys_new (1 + n)" for x
+    using that nle_le by (fastforce simp add: fareys_new_def)
+  ultimately show ?thesis
+    unfolding fareys_def farey_set_UN_Fract' by fastforce
+qed
+
+lemma length_fareys_Suc: 
+  assumes "n>0"
+  shows "length (fareys (1 + int n)) = length (fareys n) + totient (Suc n)"
+proof -
+  have "length (fareys (1 + int n)) = card (set (fareys (1 + int n)))"
+    by (metis fareys_def finite_farey_set sorted_list_of_set.sorted_key_list_of_set_unique)
+  also have "\<dots> = card (set (fareys n)) + card (fareys_new(1 + int n))"
+    using disjoint_fareys_plus1 assms by (simp add: set_fareys_plus1 card_Un_disjnt)
+  also have "\<dots> = card (set (fareys n)) + totient (Suc n)"
+    using assms card_fareys_new by force
+  also have "\<dots> = length (fareys n) + totient (Suc n)"
+    using fareys_def finite_farey_set by auto
+  finally show ?thesis .
+qed
+
+
+lemma fareys_0 [simp]: "fareys 0 = []"
+  unfolding fareys_def farey_set_UN_farey
+  by simp
+
+lemma fareys_1 [simp]: "fareys 1 = [0, 1]"
+proof -
+  have "{x \<in> {0..1}. denom_farey x \<le> 1} = {0,1}"
+    using denom_farey_le1_cases by auto
+  then show ?thesis
+    by (simp add: fareys_def)
+qed
+
+lemma fareys_2 [simp]: "fareys 2 = [0, farey 1 2, 1]"
+proof -
+  have \<section>: "denom_farey x \<le> 2 \<longleftrightarrow> denom_farey x = 1 \<or> denom_farey x = 2" for x
+    using denom_farey_pos [of x] by auto
+  have "{x \<in> {0..1}. denom_farey x \<le> 2} = {farey 0 1, farey 1 2, farey 1 1}"
+  proof -
+    have "x = farey 1 1"
+      if "x \<noteq> farey 0 1" "x \<in> {0..1}" "denom_farey x = 1" for x
+      using that denom_farey_le1_cases order.eq_iff rat_of_farey by auto
+    moreover have False
+      if "x \<noteq> farey 0 1" "x \<noteq> farey 1 2" "denom_farey x = 2" "x \<in> {0..1}" for x
+      using that num_farey_neq_denom
+      by (smt (verit) farey_0 farey_num_denom_eq num_farey_le_denom num_farey_nonneg)
+    moreover have "denom_farey (farey 1 1) = 1"
+       by (simp add: Fract_of_int_quotient farey_def)
+    ultimately show ?thesis
+      by (auto simp: farey_set_UN_farey \<section>)
+  qed
+  also have "\<dots> = {0, 1/2, 1::rat}"
+    by (simp add: farey_def Fract_of_int_quotient)
+  finally show ?thesis
+    by (simp add: fareys_def farey_def Fract_of_int_quotient)
+qed
+
+lemma length_fareys_1: 
+  shows "length (fareys 1) = 1 + totient 1"
+  by simp
+
+lemma length_fareys: "n>0 \<Longrightarrow> length (fareys n) = 1 + (\<Sum>k=1..n. totient k)"
+proof (induction n)
+  case (Suc n)
+  then show ?case
+    by (cases "n=0") (auto simp add: length_fareys_Suc)
+qed auto
+
+lemma subseq_fareys_1: "subseq (fareys n) (fareys (1 + n))"
+  by (metis fareys_increasing_1 strict_sorted_fareys sorted_subset_imp_subseq strict_sorted_imp_sorted)
+
+lemma monotone_fareys: "monotone (\<le>) subseq fareys"
+  using refl_transp_add1_int [OF _ _ subseq_order.reflp_on_le subseq_order.transp_on_le]
+  by (metis monotoneI subseq_fareys_1)
+
+lemma farey_unimodular_0_1 [simp, intro]: "farey_unimodular 0 1"
+  by (auto simp: farey_unimodular_def)
+
+text \<open>Apostol's Theorem 5.2 for integers\<close>
+lemma mediant_lies_betw_int:
+  fixes a b c d::int
+  assumes "rat_of_int a / of_int b < of_int c / of_int d" "b>0" "d>0"
+  shows "rat_of_int a / of_int b < (of_int a + of_int c) / (of_int b + of_int d)" 
+        "(rat_of_int a + of_int c) / (of_int b + of_int d) < of_int c / of_int d"
+    using assms by (simp_all add: field_split_simps)
+
+text \<open>Apostol's Theorem 5.2\<close>
+theorem mediant_inbetween:
+  fixes x y::farey
+  assumes "x < y"
+  shows "x < mediant x y" "mediant x y < y"
+  using assms mediant_lies_betw_int Fract_of_int_quotient
+  by (metis denom_farey_pos mediant_eq_Fract of_int_add rat_of_farey_conv_num_denom)+
+
+lemma sublist_fareysD:
+  assumes "sublist [x,y] (fareys n)"
+  obtains "x \<in> set (fareys n)" "y \<in> set (fareys n)"
+  by (meson assms list.set_intros set_mono_sublist subsetD)
+
+text \<open>Adding the denominators of two consecutive Farey fractions\<close>
+lemma sublist_fareys_add_denoms:
+  fixes a b c d::int
+  defines "x \<equiv> Fract a b"
+  defines "y \<equiv> Fract c d"
+  assumes sub: "sublist [x,y] (fareys (int n))" and "b>0" "d>0" "coprime a b" "coprime c d"
+  shows "b + d > n"
+proof (rule ccontr)
+  have \<section>: "x < y" "\<forall>z \<in> set (fareys n). z \<le> x \<or> z \<ge> y"
+    using sorted_two_sublist strict_sorted_fareys sub by blast+
+  assume "\<not> int n < b + d"
+  with assms have "denom_farey (mediant x y) \<le> int n"
+    by (metis denom_farey_Fract denom_mediant dual_order.trans leI)
+  then have "mediant x y \<in> set (fareys n)"
+    by (metis sub atLeastAtMost_iff denom_in_fareys_iff farey01 mediant_eq_farey
+        sublist_fareysD)
+  moreover have "x < mediant x y" "mediant x y < y"
+    by (simp_all add: mediant_inbetween \<open>x < y\<close>)
+  ultimately show False
+    using \<section> x_def y_def by fastforce
+qed
+
+subsection \<open>Apostol's Theorems 5.3--5.5\<close>
+
+theorem consec_subset_fareys:
+  fixes a b c d::int
+  assumes abcd: "0 \<le> Fract a b" "Fract a b < Fract c d" "Fract c d \<le> 1"
+    and consec: "b*c - a*d = 1"
+    and max: "max b d \<le> n" "n < b+d" 
+    and "b>0" 
+  shows "sublist [Fract a b, Fract c d] (fareys n)"
+proof (rule ccontr)
+  assume con: "\<not> ?thesis"
+  have "d > 0"
+    using max by force
+  have "coprime a b" "coprime c d"
+    using consec unimodular_imp_both_coprime by blast+
+  with \<open>b > 0\<close> \<open>d > 0\<close> have "denom_farey (Fract a b) = b" "denom_farey (Fract c d) = d"
+    by auto
+  moreover have "b\<le>n" "d\<le>n"
+    using max by auto
+  ultimately have ab: "Fract a b \<in> set (fareys n)" and cd: "Fract c d \<in> set (fareys n)"
+    using abcd finite_farey_set by (auto simp: fareys_def)
+  then obtain xs us where us: "fareys n = xs @ [Fract a b] @ us"
+    using abcd by (metis append_Cons append_Nil split_list)
+  have "Fract c d \<in> set us"
+    using abcd cd strict_sorted_fareys [of n]
+    by (fastforce simp add: us sorted_wrt_append)
+  then obtain ys zs where yz: "fareys n = xs @ [Fract a b] @ ys @ [Fract c d] @ zs"
+    by (metis split_list us append_Cons append_Nil)
+  with con have "ys \<noteq> []"
+    by (metis Cons_eq_append_conv sublist_appendI)
+  then obtain h k where hk: "coprime h k" "Fract h k \<in> set ys"  "k>0"
+    by (metis Rat_cases list.set_sel(1))
+  then have hk_fareys: "Fract h k \<in> set (fareys n)" 
+    by (auto simp: yz)
+  have less: "Fract a b < Fract h k" "Fract h k < Fract c d" 
+    using hk strict_sorted_fareys [of n] by (auto simp add: yz sorted_wrt_append)
+  with \<open>b > 0\<close> \<open>d > 0\<close> hk have *: "k*a < h*b" "d*h < c*k"
+    by (simp_all add: Fract_of_int_quotient mult.commute divide_simps flip: of_int_mult)
+  have "k \<le> n"
+    using hk by (metis hk_fareys denom_fareys_leI denom_farey_Fract)
+  have "k = (b*c - a*d)*k"
+    by (simp add: consec)
+  also have "\<dots> = b*(c*k - d*h) + d*(b*h - a*k)"
+    by (simp add: algebra_simps)
+  finally have k: "k = b * (c * k - d * h) + d * (b * h - a * k)" .  
+  moreover have "c*k - d*h \<ge> 1" "b*h - a*k \<ge> 1"
+    using \<open>b > 0\<close> \<open>d > 0\<close> * by (auto simp: mult.commute)
+  ultimately have "b * (c * k - d * h) + d * (b * h - a * k) \<ge> b+d"
+    by (metis \<open>b > 0\<close> \<open>d > 0\<close> add_mono mult.right_neutral mult_left_mono
+        order_le_less)
+  then show False
+    using \<open>k \<le> n\<close> max k by force
+qed
+
+lemma farey_unimodular_mediant:
+  assumes "farey_unimodular x y"
+  shows "farey_unimodular x (mediant x y)" "farey_unimodular (mediant x y) y"
+  using assms quotient_of_denom_pos' [of x] quotient_of_denom_pos' [of y]
+  unfolding farey_unimodular_def
+  by (auto simp: mediant_eq_Fract denom_farey_def num_farey_def quotient_of_Fract unimodular_imp_coprime algebra_simps)
+
+text \<open>Apostol's Theorem 5.4\<close>
+theorem mediant_unimodular:
+  fixes a b c d::int
+  assumes abcd: "0 \<le> Fract a b" "Fract a b < Fract c d" "Fract c d \<le> 1"
+    and consec: "b*c - a*d = 1"
+    and 0: "b>0" "d>0" 
+  defines "h \<equiv> a+c"
+  defines "k \<equiv> b+d"
+  obtains "Fract a b < Fract h k" "Fract h k < Fract c d" "coprime h k"
+          "b*h - a*k = 1"  "c*k - d*h = 1"
+proof
+  show "Fract a b < Fract h k" "Fract h k < Fract c d"
+    using abcd 0
+    by (simp_all add: Fract_of_int_quotient h_def k_def distrib_left distrib_right divide_simps)
+  show "coprime h k"
+    by (simp add: consec unimodular_imp_coprime h_def k_def)
+  show "b * h - a * k = 1"
+    by (simp add: consec distrib_left h_def k_def)
+  show "c * k - d * h = 1"
+    by (simp add: consec h_def distrib_left k_def mult.commute)
+qed
+
+text \<open>Apostol's Theorem 5.5, first part: "Each fraction in @{term"F(n+1)"} which is not in @{term"F(n)"}
+      is the mediant of a pair of consecutive fractions in @{term"F(n)"}\<close>
+
+lemma get_consecutive_parents:
+  fixes m n::int
+  assumes "coprime m n" "0<m" "m<n"
+  obtains a b c d where "m = a+c" "n = b+d" "b*c - a*d = 1" "a\<ge>0" "b>0" "c>0" "d>0" "a<b" "c\<le>d"
+proof (cases "m=1")
+  case True
+  show ?thesis
+  proof
+    show "m = 0 + 1" "n = 1 + (n-1)"
+      by (auto simp: True)
+  qed (use True \<open>m<n\<close> in auto)
+next
+  case False
+  then obtain d c where *: "n*c - m*d = 1" "0 < d" "d < n" "0 < c" "c < m"
+    using coprime_unimodular_int
+ [of n m] coprime_commute assms by (smt (verit) coprime_commute)
+  then have **: "n * (c - d) + (n - m) * d = 1"
+    by (metis mult_diff_mult)
+  show ?thesis
+  proof
+    show "c\<le>d"
+      using * ** \<open>m<n\<close> by (smt (verit) mult_le_0_iff)
+    show "(n-d) * c - (m-c) * d = 1"
+      using * by (simp add: algebra_simps)
+    with * \<open>m<n\<close> show "m-c < n-d"
+      by (smt (verit, best) mult_mono)
+  qed (use * in auto)
+qed
+
+theorem fareys_new_eq_mediant:
+  assumes "x \<in> fareys_new n" "n>1"
+  obtains a b c d where 
+    "sublist [Fract a b, Fract c d] (fareys (n-1))" 
+    "x = mediant (Fract a b) (Fract c d)" "coprime a b" "coprime c d" "a\<ge>0" "b>0" "c>0" "d>0" 
+proof -
+  obtain m where m: "coprime m n" "0 \<le> m" "m \<le> n" "x = Fract m n"
+    using assms nless_le zero_less_imp_eq_int by (force simp: fareys_new_def)
+  moreover
+  have "x \<noteq> 0" "x \<noteq> 1"
+    using assms fareys_new_not01 by auto
+  with m have "0<m" "m<n"
+    using \<open>n>1\<close> of_nat_le_0_iff by fastforce+
+  ultimately
+  obtain a b c d where 
+    abcd: "m = a+c" "n = b+d" "b*c - a*d = 1" "a\<ge>0" "b>0" "c>0" "d>0" "a<b" "c\<le>d"
+    by (metis get_consecutive_parents)
+  show thesis
+  proof
+    have "Fract a b < Fract c d"
+      using abcd mult.commute[of b c] by force
+    with consec_subset_fareys
+    show "sublist [Fract a b, Fract c d] (fareys (n-1))"
+      using Fract_le_one_iff abcd zero_le_Fract_iff by auto
+    show "x = mediant (Fract a b) (Fract c d)"
+      using abcd \<open>x = Fract m n\<close> mediant_eq_Fract unimodular_imp_both_coprime by fastforce
+    show "coprime a b" "coprime c d"
+      using \<open>b * c - a * d = 1\<close> unimodular_imp_both_coprime by blast+
+  qed (use abcd in auto)
+qed
+
+text \<open>Apostol's Theorem 5.5, second part: "Moreover, if @{term"a/b<c/d"} are consecutive in any @{term"F(n)"},
+then they satisfy the unimodular relation @{term"bc - ad = 1"}.\<close>
+theorem consec_imp_unimodular:
+  assumes "sublist [Fract a b, Fract c d] (fareys (int n))" "b>0" "d>0" "coprime a b" "coprime c d"
+  shows  "b*c - a*d = 1"
+  using assms 
+proof (induction n arbitrary: a b c d)
+  case 0
+  then show ?case
+    by auto
+next
+  case (Suc n)
+  show ?case
+  proof (cases "n=0")
+    case True
+    with Suc.prems have "Fract a b = 0" "Fract c d = 1"
+      by (auto simp add: sublist_Cons_right)
+    with Suc.prems obtain "a=0" "b=1" "c=1" "d=1"
+      by (auto simp: Fract_of_int_quotient)
+    then show ?thesis
+      by auto
+  next
+    case False
+    have "Fract a b < Fract c d" 
+      and ab: "Fract a b \<in> set (fareys (1 + int n))" and cd: "Fract c d \<in> set (fareys (1 + int n))"
+      using strict_sorted_fareys [of "Suc n"] Suc.prems
+      by (auto simp add: sublist_def sorted_wrt_append)
+    have con: "z \<le> Fract a b \<or> Fract c d \<le> z" if "z \<in> set (fareys (int n))" for z
+    proof -
+      have "z \<in> set (fareys (1 + int n))"
+        using fareys_increasing_1 that by blast
+      with Suc.prems strict_sorted_fareys [of "1 + int n"] show ?thesis
+        by (fastforce simp add: sublist_def sorted_wrt_append)
+    qed
+    show ?thesis
+    proof (cases "Fract a b \<in> set (fareys n) \<and> Fract c d \<in> set (fareys n)")
+      case True
+      then have "sublist [Fract a b, Fract c d] (fareys n)"
+        using con \<open>Fract a b < Fract c d\<close> sorted_two_sublist strict_sorted_fareys by blast
+      then show ?thesis
+        by (simp add: Suc) 
+    next
+      case False
+      have notboth: False if \<section>: "Fract a b \<in> fareys_new (1+n)" "Fract c d \<in> fareys_new (1+n)"
+      proof -
+        obtain a' b' c' d' where eq':
+          "sublist [Fract a' b', Fract c' d'] (fareys n)" "Fract a b = mediant (Fract a' b') (Fract c' d')"
+          by (smt (verit) \<open>n\<noteq>0\<close> \<section> fareys_new_eq_mediant of_nat_Suc of_nat_le_0_iff plus_1_eq_Suc)
+        then have abcd': "Fract a' b' \<in> set (fareys n)" "Fract c' d' \<in> set (fareys n)"
+          by (auto simp: sublist_def)
+        have con': "z \<le> Fract a' b' \<or> Fract c' d' \<le> z" if "z \<in> set (fareys n)" for z
+          by (metis eq'(1) sorted_two_sublist strict_sorted_fareys that)
+        have "Fract a' b' < Fract c' d'"
+          using eq'(1) sorted_two_sublist [OF strict_sorted_fareys] by blast         
+        then obtain A: "Fract a' b' < Fract a b" "Fract a b < Fract c' d'"
+          using eq'(2) mediant_inbetween by presburger
+          obtain a'' b'' c'' d'' where eq'':
+          "sublist [Fract a'' b'', Fract c'' d''] (fareys n)" "Fract c d = mediant (Fract a'' b'') (Fract c'' d'')"
+          by (smt (verit) \<section> \<open>n\<noteq>0\<close> fareys_new_eq_mediant of_nat_Suc of_nat_le_0_iff plus_1_eq_Suc)
+        then have abcd'': "Fract a'' b'' \<in> set (fareys n)" "Fract c'' d'' \<in> set (fareys n)"
+          by (auto simp: sublist_def)
+        then have "Fract c'' d'' \<in> set (fareys (1 + int n))"
+          using fareys_increasing_1 by blast
+        have con'': "z \<le> Fract a'' b'' \<or> Fract c'' d'' \<le> z" if "z \<in> set (fareys n)" for z
+          using sorted_two_sublist [OF strict_sorted_fareys] eq''(1) that by blast
+        then obtain "Fract a'' b'' < Fract c'' d''" "Fract a'' b'' < Fract c d" "Fract c d < Fract c'' d''"
+          by (metis eq'' sorted_two_sublist [OF strict_sorted_fareys] mediant_inbetween)
+        with A show False
+          using con' con'' abcd' abcd'' con \<open>Fract a b < Fract c d\<close>
+          by (metis eq'(2) eq''(2) dual_order.strict_trans1 not_less_iff_gr_or_eq)
+      qed
+      consider "Fract a b \<in> fareys_new (1+n)" | "Fract c d \<in> fareys_new (1+n)"
+        using False set_fareys_plus1 [of n]
+        by (metis (mono_tags, opaque_lifting) Suc.prems(1) Suc_eq_plus1_left UnE list.set_intros(1) of_nat_Suc set_mono_sublist
+            set_subset_Cons subset_iff)
+      then show ?thesis
+      proof cases
+        case 1
+        then obtain a' b' c' d' where eq:
+          "sublist [Fract a' b', Fract c' d'] (fareys n)" 
+          "Fract a b = mediant (Fract a' b') (Fract c' d')" "coprime a' b'" "coprime c' d'" "b'>0" "d'>0" 
+          by (smt (verit) \<open>n\<noteq>0\<close> fareys_new_eq_mediant of_nat_Suc of_nat_le_0_iff plus_1_eq_Suc)
+        then have abcd': "Fract a' b' \<in> set (fareys n)" "Fract c' d' \<in> set (fareys n)"
+          by (auto simp: sublist_def)
+        have con': "z \<le> Fract a' b' \<or> Fract c' d' \<le> z" if "z \<in> set (fareys n)" for z
+          using eq(1) sorted_two_sublist strict_sorted_fareys that by blast
+        obtain "Fract a' b' < Fract c' d'" "Fract a b < Fract c' d'"
+          using eq by (simp add: mediant_inbetween(2) sorted_two_sublist strict_sorted_fareys)
+        then have "Fract c d \<le> Fract c' d'"
+          using con abcd' linorder_not_less by blast
+        moreover have "Fract c' d' \<le> Fract c d" 
+          if "Fract c d \<in> set (fareys n)"
+          by (metis con' \<open>Fract a b < Fract c d\<close> \<open>Fract a' b' < Fract c' d'\<close> eq(2) order.trans linorder_not_less mediant_inbetween(1)
+              nless_le that)
+        ultimately have "Fract c' d' = Fract c d"
+          using notboth "1" cd set_fareys_plus1 by auto
+        with Suc.prems obtain "c' = c" "d' = d"
+          by (metis \<open>0 < d'\<close> \<open>coprime c' d'\<close> denom_farey_Fract num_farey_Fract)
+        then have uni: "b'*c - a'*d = 1"
+          using Suc eq by blast
+        then obtain "a = a' + c" "b = b' + d"
+          using eq Suc.prems apply (simp add: mediant_eq_Fract)
+          by (metis \<open>c' = c\<close> \<open>d' = d\<close> denom_farey_Fract num_farey_Fract pos_add_strict
+              unimodular_imp_coprime)
+        with uni show ?thesis
+          by (auto simp: algebra_simps)
+      next
+        case 2
+        then obtain a' b' c' d' where eq:
+          "sublist [Fract a' b', Fract c' d'] (fareys n)" 
+          "Fract c d = mediant (Fract a' b') (Fract c' d')" "coprime a' b'" "coprime c' d'" "b'>0" "d'>0" 
+          by (smt (verit) \<open>n\<noteq>0\<close> fareys_new_eq_mediant of_nat_Suc of_nat_le_0_iff plus_1_eq_Suc)
+        then have abcd': "Fract a' b' \<in> set (fareys n)" "Fract c' d' \<in> set (fareys n)"
+          by (auto simp: sublist_def)
+        have con': "z \<le> Fract a' b' \<or> Fract c' d' \<le> z" if "z \<in> set (fareys n)" for z
+          using eq(1) sorted_two_sublist strict_sorted_fareys that by blast
+        obtain "Fract a' b' < Fract c' d'" "Fract a' b' < Fract c d"
+          using eq mediant_inbetween
+          by (metis sorted_two_sublist strict_sorted_fareys)
+        then have "Fract a' b' \<le> Fract a b"
+          using con abcd' linorder_not_less by blast
+        moreover have "Fract a b \<le> Fract a' b'" 
+          if "Fract a b \<in> set (fareys n)"
+          by (metis \<open>Fract a b < Fract c d\<close> \<open>Fract a' b' < Fract c' d'\<close> con' order.strict_trans2 eq(2) mediant_inbetween(2)
+              not_less_iff_gr_or_eq that)
+        ultimately have "Fract a' b' = Fract a b"
+          using notboth 2 ab set_fareys_plus1 by auto
+        with Suc.prems obtain "a' = a" "b' = b"
+          by (metis \<open>0 < b'\<close> \<open>coprime a' b'\<close> denom_farey_Fract num_farey_Fract)
+        then have uni: "b*c' - a*d' = 1"
+          using Suc.IH Suc.prems eq by blast
+        then obtain "c = a + c'" "d = b + d'"
+          using eq Suc.prems apply (simp add: mediant_eq_Fract)
+          by (metis \<open>a' = a\<close> \<open>b' = b\<close> denom_farey_Fract num_farey_Fract pos_add_strict
+              unimodular_imp_coprime)
+        with uni show ?thesis
+          by (auto simp: algebra_simps)
+      qed
+    qed
+  qed
+qed
+
+subsection \<open>Ford circles\<close>
+
+definition Ford_center :: "rat \<Rightarrow> complex" where
+  "Ford_center r \<equiv> (\<lambda>(h,k). Complex (h/k) (1/(2 * k^2))) (quotient_of r)"
+
+definition Ford_radius :: "rat \<Rightarrow> real" where
+  "Ford_radius r \<equiv> (\<lambda>(h,k). 1/(2 * k^2)) (quotient_of r)"
+
+definition Ford_tan :: "[rat,rat] \<Rightarrow> bool" where
+  "Ford_tan r s \<equiv> dist (Ford_center r) (Ford_center s) = Ford_radius r + Ford_radius s"
+
+definition Ford_circle :: "rat \<Rightarrow> complex set" where
+  "Ford_circle r \<equiv> sphere (Ford_center r) (Ford_radius r)"
+
+lemma Im_Ford_center [simp]: "Im (Ford_center r) = Ford_radius r"
+  by (auto simp: Ford_center_def Ford_radius_def split: prod.splits)
+
+lemma Ford_radius_nonneg: "Ford_radius r \<ge> 0"
+  by (simp add: Ford_radius_def split: prod.splits)
+
+lemma two_Ford_tangent:
+  assumes r: "(a,b) = quotient_of r" and s: "(c,d) = quotient_of s"
+  shows "(dist (Ford_center r) (Ford_center s))^2 - (Ford_radius r + Ford_radius s)^2 
+       = ((a*d - b*c)^2 - 1) / (b*d)^2"
+proof -
+  obtain 0: "b > 0" "d > 0"
+    by (metis assms quotient_of_denom_pos)
+  have 1: "dist (Ford_center r) (Ford_center s) ^ 2 = (a/b - c/d)^2 + (1/(2*b^2) - 1/(2*d^2)) ^ 2"
+    using assms by (force simp: Ford_center_def dist_norm complex_norm complex_diff split: prod.splits)
+  have 2: "(Ford_radius r + Ford_radius s) ^ 2 = (1/(2*b^2) + 1/(2*d^2)) ^ 2"
+    using assms by (force simp: Ford_radius_def split: prod.splits)
+  show ?thesis
+    using 0 unfolding 1 2 by (simp add: field_simps eval_nat_numeral)
+qed
+
+text \<open>Apostol's Theorem 5.6\<close>
+lemma two_Ford_tangent_iff:
+  assumes r: "(a,b) = quotient_of r" and s: "(c,d) = quotient_of s"
+  shows "Ford_tan r s \<longleftrightarrow> \<bar>b * c - a * d\<bar> = 1"
+proof -
+  obtain 0: "b > 0" "d > 0"
+    by (metis assms quotient_of_denom_pos)
+  have "Ford_tan r s \<longleftrightarrow> dist (Ford_center r) (Ford_center s) ^ 2 = (Ford_radius r + Ford_radius s) ^ 2"
+    using Ford_radius_nonneg by (simp add: Ford_tan_def)
+  also have "\<dots> \<longleftrightarrow> ((a*d - b*c)^2 - 1) / (b*d)^2 = 0"
+    using two_Ford_tangent [OF assms] by (simp add: diff_eq_eq)
+  also have "\<dots> \<longleftrightarrow> \<bar>b * c - a * d\<bar> = 1"
+    using 0 by (simp add: abs_square_eq_1 abs_minus_commute flip: of_int_mult of_int_diff)
+  finally show ?thesis .
+qed
+
+text \<open>Also Apostol's Theorem 5.6: Distinct Ford circles do not overlap\<close>
+lemma Ford_no_overlap:
+  assumes "r \<noteq> s"
+  shows "dist (Ford_center r) (Ford_center s) \<ge> Ford_radius r + Ford_radius s"
+proof -
+  obtain a b c d where r: "(a,b) = quotient_of r" and s: "(c,d) = quotient_of s"
+                 and "b>0" "d>0"
+    by (metis quotient_of_denom_pos surj_pair)
+  moreover have "a \<noteq> c \<or> b \<noteq> d"
+    using assms r s quotient_of_inject by force
+  ultimately have "a * d \<noteq> c * b"
+    by (metis Fract_of_int_quotient assms eq_rat(1) less_irrefl quotient_of_div)
+  then have "(a*d - b*c)^2 \<ge> (1::int)"
+    by (simp add: mult.commute int_one_le_iff_zero_less)
+  then have "((a*d - b*c)^2 - 1) / (b*d)^2 \<ge> (0::real)"
+    by (simp add: divide_simps mult_less_0_iff flip: of_int_mult of_int_power)
+  then show ?thesis
+    using two_Ford_tangent [OF r s]
+    by (metis (no_types, lifting) ge_iff_diff_ge_0 of_int_1 of_int_diff of_int_mult 
+              of_int_power power2_le_imp_le zero_le_dist)
+qed
+
+lemma Ford_aux1:
+  assumes "a\<noteq>0"
+  shows "cmod (Complex (b / (a * (a2 + b2))) (1 / (2 * a2) - inverse (a2 + b2))) = 1 / (2 * a2)"
+       (is "cmod ?z = ?r")
+proof -
+  have "(2 * a2) * cmod ?z = cmod ((2 * a2) * ?z)"
+    by (simp add: norm_mult power2_eq_square)
+  also have "\<dots> = cmod (Complex (2*a*b / (a2 + b2)) (1 - (2 * a2) / (a2 + b2)))"
+    unfolding complex_of_real_mult_Complex inverse_eq_divide
+    using \<open>a\<noteq>0\<close> by (simp add: power2_eq_square mult.assoc right_diff_distrib)
+  also have "\<dots> = cmod (Complex (2*a*b) ((a2 + b2) - (2 * a2)) / (a2 + b2))"
+    unfolding Complex_divide_complex_of_real diff_divide_distrib
+    using assms by force
+  also have "\<dots> = cmod (Complex (2*a*b) ((a2 + b2) - (2 * a2))) / (a2 + b2)"
+    by (smt (verit) norm_divide norm_of_real not_sum_power2_lt_zero)
+  also have "\<dots> = sqrt ((a2 + b2)^2) / (a2 + b2)"
+    unfolding power2_eq_square complex_norm
+    by (simp add: algebra_simps)
+  also have "\<dots> = 1"
+    using assms by auto
+  finally show ?thesis
+    by (metis inverse_eq_divide inverse_unique)
+qed
+
+lemma Ford_aux2:
+  assumes "a\<noteq>0"
+  shows "cmod (Complex (a / (b * (b2 + a2)) - 1 / (a * b)) (1 / (2 * a2) - inverse (b2 + a2))) = 1 / (2 * a2)"
+       (is "cmod ?z = ?r")
+proof -
+  have "a / (b * (b2 + a2)) - 1 / (a * b) = -b / (a * (b2 + a2))"
+    by (simp add: divide_simps power2_eq_square)
+  then have "cmod ?z = cmod (Complex (b / (a * (a2 + b2))) (1 / (2 * a2) - inverse (a2 + b2)))"
+    by (simp add: cmod_neg_real add.commute)
+  also have "\<dots> = 1 / (2 * a2)"
+    using Ford_aux1 assms by simp
+  finally show ?thesis .
+qed
+
+definition Radem_trans :: "rat \<Rightarrow> complex \<Rightarrow> complex" where
+  "Radem_trans \<equiv> \<lambda>r \<tau>. let (h,k) = quotient_of r in - \<i> * of_int k ^ 2 * (\<tau> - of_rat r)"
+
+text \<open>Theorem 5.8 first part\<close>
+lemma Radem_trans_image: "Radem_trans r ` Ford_circle r = sphere (of_rat (1/2)) (1/2)"
+proof -
+  obtain h k where r: "quotient_of r = (h,k)" and "k>0" and req: "r = of_int h / of_int k"
+    using quotient_of_denom_pos quotient_of_div by fastforce
+  have "Radem_trans r ` Ford_circle r = ((*) (-\<i> * of_int k ^ 2)) ` (\<lambda>\<tau>. \<tau> - of_rat r) ` Ford_circle r"
+    by (simp add: Radem_trans_def r image_comp)
+  also have "\<dots> = ((*) (-\<i> * of_int k ^ 2)) ` sphere (Ford_center r - of_rat r) (Ford_radius r)"
+    by (simp add: Ford_circle_def flip: sphere_translation_subtract)
+  also have "\<dots> = sphere (- \<i> * (of_int k)2 * (Ford_center r - r))
+                          (cmod (- \<i> * (of_int k)2) * Ford_radius r)"
+    using \<open>k>0\<close> by (intro sphere_cscale) auto
+  also have "\<dots> = sphere (of_rat (1/2)) (1/2)"
+  proof -
+    have "(- \<i> * (of_int k)2 * (Ford_center r - r)) = 1/2"
+      using \<open>k>0\<close>
+      apply (simp add: Ford_center_def r algebra_simps Complex_eq)
+      by (simp add: of_rat_divide req)
+    moreover 
+    have "(cmod (- \<i> * (of_int k)2) * Ford_radius r) = 1/2"
+      using \<open>k>0\<close>
+      by (simp add: norm_mult norm_power Ford_radius_def r)
+    ultimately show ?thesis
+      by (metis of_rat_1 of_rat_divide of_rat_numeral_eq)
+  qed
+  finally show ?thesis .
+qed
+
+locale three_Ford =
+  fixes N::nat
+  fixes h1 k1 h k h2 k2::int
+  assumes sub1: "sublist [Fract h1 k1, Fract h k] (fareys (int N))"
+  assumes sub2: "sublist [Fract h k, Fract h2 k2] (fareys (int N))"
+  assumes coprime: "coprime h1 k1" "coprime h k" "coprime h2 k2"
+  assumes k_pos: "k1 > 0" "k > 0" "k2 > 0"
+
+begin
+
+definition "r1 \<equiv> Fract h1 k1"
+definition "r \<equiv> Fract h k"
+definition "r2 \<equiv> Fract h2 k2"
+
+lemma N_pos: "N>0"
+  using sub1 gr0I by force
+
+lemma r_eq_quotient:
+  "(h1,k1) = quotient_of r1" "(h,k) = quotient_of r" "(h2,k2) = quotient_of r2"
+  by (simp_all add: coprime k_pos quotient_of_Fract r1_def r_def r2_def)
+
+lemma r_eq_divide:
+  "r1 = of_int h1 / of_int k1" "r = of_int h / of_int k" "r2 = of_int h2/ of_int k2"
+  by (simp_all add: Fract_of_int_quotient of_rat_divide r1_def r2_def r_def)
+
+lemma collapse_r:
+  "real_of_int h1 / of_int k1 = of_rat r1" 
+  "real_of_int h / of_int k = of_rat r" "real_of_int h2/ of_int k2 = of_rat r2"
+  by (simp_all add: of_rat_divide r_eq_divide)
+
+lemma unimod1: "k1*h - h1*k = 1"
+  and unimod2: "k*h2 - h*k2 = 1"
+  using consec_imp_unimodular coprime k_pos sub1 sub2 by blast+
+
+lemma r_less: "r1 < r" "r < r2"
+  using r1_def r_def r2_def sub1 sub2 sorted_two_sublist [OF strict_sorted_fareys] by auto
+
+lemma r01:
+  obtains "r1 \<in> {0..1}" "r \<in> {0..1}" "r2 \<in> {0..1}"
+  by (metis denom_in_fareys_iff r1_def r2_def r_def sub1 sub2 sublist_fareysD)
+
+lemma atMost_N:
+  obtains "k1 \<le> N" "k \<le> N" "k2 \<le> N"
+  by (metis denom_farey_def denom_in_fareys_iff prod.sel(2) r1_def r2_def r_def r_eq_quotient sub1 sub2
+      sublist_fareysD)
+
+lemma greaterN1: "k1 + k > N"
+  using sublist_fareys_add_denoms coprime k_pos sub1 by blast
+
+lemma greaterN2: "k + k2 > N"
+  using sublist_fareys_add_denoms coprime k_pos sub2 by blast
+
+definition "alpha1 \<equiv> Complex (h/k - k1 / of_int(k * (k2 + k12))) (inverse (of_int (k2 + k12)))"
+definition "alpha2 \<equiv> Complex (h/k + k2 / of_int(k * (k2 + k22))) (inverse (of_int (k2 + k22)))"
+
+definition "zed1 \<equiv> Complex (k2) (k*k1) / ((k2 + k12))"
+definition "zed2 \<equiv> Complex (k2) (- k*k2) / ((k2 + k22))"
+
+text \<open>Apostol's Theorem 5.7\<close>
+lemma three_Ford_tangent:
+  obtains "alpha1 \<in> Ford_circle r" "alpha1 \<in> Ford_circle r1"
+          "alpha2 \<in> Ford_circle r" "alpha2 \<in> Ford_circle r2"
+proof
+  show "alpha1 \<in> Ford_circle r"
+    using k_pos Ford_aux1 r_eq_quotient
+    by (force simp: alpha1_def Ford_circle_def Ford_center_def dist_norm complex_diff 
+        Ford_radius_def split: prod.splits)
+  have 1: "real_of_int h1 / real_of_int k1 = real_of_int h / real_of_int k - 1 / (k1*k)"
+    using unimod1 k_pos
+    by (simp add: divide_simps) (simp add: algebra_simps flip: of_int_mult of_int_diff)
+  show "alpha1 \<in> Ford_circle r1"
+    using k_pos Ford_aux2 r_eq_quotient
+    by (force simp: alpha1_def Ford_circle_def Ford_center_def dist_norm complex_diff 1 Ford_radius_def split: prod.splits)
+  show "alpha2 \<in> Ford_circle r"
+    using k_pos Ford_aux1 [of k k2] cmod_neg_real r_eq_quotient
+    by (force simp add: alpha2_def Ford_circle_def Ford_center_def dist_norm complex_diff Ford_radius_def split: prod.splits)
+  have 2: "real_of_int h / real_of_int k = real_of_int h2 / real_of_int k2 - 1 / (k*k2)"
+    using unimod2 k_pos
+    by (simp add: divide_simps) (simp add: algebra_simps flip: of_int_mult of_int_diff)
+  show "alpha2 \<in> Ford_circle r2"
+    using k_pos Ford_aux2 [of k2 k] cmod_neg_real r_eq_quotient
+    apply (simp add: alpha2_def Ford_circle_def Ford_center_def dist_norm complex_diff 2 Ford_radius_def split: prod.splits)
+    by (smt (verit) mult.commute prod.sel)
+qed
+
+text \<open>Theorem 5.8 second part, for alpha1\<close>
+lemma Radem_trans_alpha1: "Radem_trans r alpha1 = zed1"
+proof -
+  have "Radem_trans r alpha1 = ((*) (-\<i> * of_int k ^ 2)) ((\<lambda>\<tau>. \<tau> - of_rat r) alpha1)"
+    by (metis Radem_trans_def prod.simps(2) r_eq_quotient(2))
+  also have "\<dots> = ((*) (-\<i> * of_int k ^ 2)) (Complex (- k1 / of_int(k * (k2 + k12))) (inverse (of_int (k2 + k12))))"
+    using k_pos by (simp add: alpha1_def r_def of_rat_rat Complex_eq)
+  also have "\<dots> = zed1"
+    unfolding complex_eq_iff by (simp add: zed1_def inverse_eq_divide power2_eq_square)
+  finally show ?thesis .
+qed
+
+text \<open>Theorem 5.8 second part, for alpha2\<close>
+lemma Radem_trans_alpha2: "Radem_trans r alpha2 = zed2"
+proof -
+  have "Radem_trans r alpha2 = ((*) (-\<i> * of_int k ^ 2)) ((\<lambda>\<tau>. \<tau> - of_rat r) alpha2)"
+    by (metis Radem_trans_def prod.simps(2) r_eq_quotient(2))
+  also have "\<dots> = ((*) (-\<i> * of_int k ^ 2)) (Complex (k2 / of_int(k * (k2 + k22))) (inverse (of_int (k2 + k22))))"
+    using k_pos by (simp add: alpha2_def r_def of_rat_rat Complex_eq)
+  also have "\<dots> = zed2"
+    unfolding complex_eq_iff by (simp add: zed2_def inverse_eq_divide power2_eq_square)
+  finally show ?thesis .
+qed
+
+text \<open>Theorem 5.9, for zed1\<close>
+lemma cmod_zed1: "cmod zed1 = k / sqrt (k2 + k12)"
+proof -
+  have "cmod zed1 ^ 2 = (k^4 + k2 * k12) / (k2 + k12)^2"
+    by (simp add: zed1_def cmod_def divide_simps)
+  also have "\<dots> = (of_int k) ^ 2 / (k2 + k12)"
+    by (simp add: eval_nat_numeral divide_simps) argo
+  finally have "cmod zed1 ^ 2 = (of_int k) ^ 2 / (k2 + k12)" .
+  with k_pos real_sqrt_divide show ?thesis
+    unfolding cmod_def by force
+qed
+
+text \<open>Theorem 5.9, for zed2\<close>
+lemma cmod_zed2: "cmod zed2 = k / sqrt (k2 + k22)"
+proof -
+  have "cmod zed2 ^ 2 = (k^4 + k2 * k22) / (k2 + k22)^2"
+    by (simp add: zed2_def cmod_def divide_simps)
+  also have "\<dots> = (of_int k) ^ 2 / (k2 + k22)"
+    by (simp add: eval_nat_numeral divide_simps) argo
+  finally have "cmod zed2 ^ 2 = (of_int k) ^ 2 / (k2 + k22)" .
+  with k_pos real_sqrt_divide show ?thesis
+    unfolding cmod_def by force
+qed
+
+
+text \<open>The last part of theorem 5.9\<close>
+
+lemma RMS_calc:
+  assumes "k' > 0" "k' + k > int N"
+  shows "k / sqrt (k2 + k'2) < sqrt 2 * k/N"
+proof -
+  have \<section>: "(k + k')/2 \<le> sqrt ((k2 + k'2) / 2)"
+    using sum_squared_le_sum_of_squares_2 by simp
+  have "N / sqrt 2 < (N+1) / sqrt 2"
+    by (simp add: divide_strict_right_mono)
+  also have "\<dots> \<le> (k + k') / sqrt 2"
+    using assms by (simp add: divide_simps)
+  also have "\<dots> = (k + k')/2 * sqrt 2"
+    by (metis nonzero_divide_eq_eq real_div_sqrt times_divide_eq_right zero_le_numeral
+        zero_neq_numeral)
+  also have "\<dots> \<le> sqrt (k2 + k'2)"
+    using \<section> by (simp add: le_divide_eq real_sqrt_divide)
+  finally have 1: "real N / sqrt 2 < sqrt (real_of_int (k2 + k'2))" .
+  with N_pos k_pos not_sum_power2_lt_zero show ?thesis
+    by (force simp add: cmod_zed1 mult.commute divide_simps)
+qed
+
+lemma on_chord_bounded_cmod:
+  assumes "z \<in> closed_segment zed1 zed2"
+  shows "cmod z < sqrt 2 * k/N"
+proof -
+  have "cmod z \<le> max (cmod zed1) (cmod zed2)"
+    using segment_furthest_le [OF assms, of 0] by auto
+  moreover
+  obtain "cmod zed1 < sqrt 2 * k/N" "cmod zed2 < sqrt 2 * k/N"
+    using RMS_calc cmod_zed1 cmod_zed2 greaterN1 greaterN2 k_pos by force
+  ultimately show ?thesis
+    using assms cmod_zed1 cmod_zed2 by linarith
+qed
+
+end
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Farey_Sequences/ROOT
--- /dev/null
+++ thys/Farey_Sequences/ROOT
@@ -0,0 +1,11 @@
+chapter AFP
+
+session Farey_Sequences (AFP) = "HOL-Analysis" +
+  options [timeout = 300, document_build = pdflatex]
+  sessions "HOL-Library" "HOL-Number_Theory"
+  theories
+    Farey_Ford
+  document_files
+    "root.tex"
+    "root.bib"
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Farey_Sequences/document/root.bib
--- /dev/null
+++ thys/Farey_Sequences/document/root.bib
@@ -0,0 +1,17 @@
+%% This BibTeX bibliography file was created using BibDesk.
+%% https://bibdesk.sourceforge.io/
+
+%% Created for Larry Paulson at 2025-05-15 14:39:22 +0100 
+
+
+%% Saved with string encoding Unicode (UTF-8) 
+
+
+
+@book{apostol-modular-functions,
+	author = {Tom M. Apostol},
+	date-added = {2025-05-15 14:39:00 +0100},
+	date-modified = {2025-05-15 14:39:00 +0100},
+	publisher = {Springer},
+	title = {Modular Functions and Dirichlet Series in Number Theory},
+	year = {1990}}
diff -r 848f2c00e57f -r 29846bc89de2 thys/Farey_Sequences/document/root.tex
--- /dev/null
+++ thys/Farey_Sequences/document/root.tex
@@ -0,0 +1,53 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+\usepackage{amsmath,amssymb}
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{it}
+
+%%Euro-style date: 20 September 1955
+\def\today{\number\day\space\ifcase\month\or
+January\or February\or March\or April\or May\or June\or
+July\or August\or September\or October\or November\or December\fi
+\space\number\year}
+
+\begin{document}
+
+\title{Farey Sequences and Ford Circles}
+\author{Lawrence C. Paulson}
+\maketitle
+
+\begin{abstract}
+
+The sequence $F_n$ of \emph{Farey fractions} of order~$n$ 
+has the form 
+$$\frac{0}{1}, \frac{1}{n}, \frac{1}{n-1}, \ldots, \frac{n-1}{n}, \frac{1}{1}$$
+where the fractions appear in numerical order and have denominators at most $n$.
+The transformation from $F_n$ to $F_{n+1}$ can be effected by combining adjacent elements of 
+the sequence~$F_n$, using an operation called the \emph{mediant}.
+Adjacent (reduced) fractions $(a/b) < (c/d)$ satisfy the \emph{unimodular}
+relation $bc - ad = 1$ and their mediant is $\frac{a+c}{b+d}$.
+A \emph{Ford circle} is specified by a rational number, and interesting consequences follow
+in the case of Ford circles obtained from some Farey sequence~$F_n$.
+The formalised material is drawn from Apostol's \emph{Modular Functions and Dirichlet Series in Number Theory}
+\cite{apostol-modular-functions}.
+
+\end{abstract}
+
+\newpage
+
+% include generated text of all theories
+\input{session}
+
+\paragraph*{Acknowledgements}
+Manual Eberl set up the initial Farey development.
+
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
diff -r 848f2c00e57f -r 29846bc89de2 thys/Finite_Automata_HF/Finite_Automata_HF.thy
--- thys/Finite_Automata_HF/Finite_Automata_HF.thy
+++ thys/Finite_Automata_HF/Finite_Automata_HF.thy
@@ -11,6 +11,10 @@
   the powerset construction mapping NFAs to DFAs. Left and right languages; minimal DFAs.
   Brzozowski's minimization algorithm. Uniqueness up to isomorphism of minimal DFAs.\<close>
 
+text\<open>Initially this formalization was based on automata whose state type was always \<open>hf\<close>.
+In a revision, it was generalized: many of the constructions are now based on automata
+with a polymorphic state type \<open>'s\<close>.\<close>
+
 section\<open>Deterministic Finite Automata\<close>
 
 text\<open>Right invariance is the key property for equivalence relations on states of DFAs.\<close>
@@ -19,14 +23,17 @@
 
 subsection\<open>Basic Definitions\<close>
 
+text \<open>We try to make as much as possible polymorphic in the state space \<open>'s\<close>
+and independent of type typ\<open>hf\<close> to increase reusability.\<close>
+
 text\<open>First, the record for DFAs\<close>
-record 'a dfa = states :: "hf set"
-                init   :: "hf"
-                final  :: "hf set"
-                nxt    :: "hf \<Rightarrow> 'a \<Rightarrow> hf"
+record ('a,'s) dfa = states :: "'s set"
+                init   :: "'s"
+                final  :: "'s set"
+                nxt    :: "'s \<Rightarrow> 'a \<Rightarrow> 's"
 
 locale dfa =
-  fixes M :: "'a dfa"
+  fixes M :: "('a,'s) dfa"
   assumes init [simp]: "init M \<in> states M"
       and final:       "final M \<subseteq> states M"
       and nxt:         "\<And>q x. q \<in> states M \<Longrightarrow> nxt M q x \<in> states M"
@@ -37,7 +44,7 @@
   using final finite_subset finite by blast
 
 text\<open>Transition function for a given starting state and word.\<close>
-primrec nextl :: "[hf, 'a list] \<Rightarrow> hf" where
+primrec nextl :: "['s, 'a list] \<Rightarrow> 's" where
     "nextl q []     = q"
   | "nextl q (x#xs) = nextl (nxt M q x) xs"
 
@@ -45,7 +52,7 @@
   "language \<equiv> {xs. nextl (init M) xs \<in> final M}"
 
 text\<open>The left language WRT a state q is the set of words that lead to q.\<close>
-definition left_lang :: "hf \<Rightarrow> 'a list set"  where
+definition left_lang :: "'s \<Rightarrow> 'a list set"  where
   "left_lang q \<equiv> {u. nextl (init M) u = q}"
 
 text\<open>Part of Prop 1 of
@@ -57,7 +64,7 @@
   unfolding left_lang_def by auto
 
 text\<open>The right language WRT a state q is the set of words that go from q to F.\<close>
-definition right_lang :: "hf \<Rightarrow> 'a list set"  where
+definition right_lang :: "'s \<Rightarrow> 'a list set"  where
   "right_lang q \<equiv> {u. nextl q u \<in> final M}"
 
 lemma language_eq_right_lang: "language = right_lang (init M)"
@@ -114,7 +121,7 @@
 
 subsection \<open>Minimisation via Accessibility\<close>
 
-definition accessible :: "hf set"  where
+definition accessible :: "'s set"  where
   "accessible \<equiv> {q. left_lang q \<noteq> {}}"
 
 lemma accessible_imp_states: "q \<in> accessible \<Longrightarrow> q \<in> states M"
@@ -126,7 +133,7 @@
 lemma inj_on_left_lang: "inj_on left_lang accessible"
   by (auto simp: inj_on_def left_lang_def accessible_def)
 
-definition path_to :: "hf \<Rightarrow> 'a list"  where
+definition path_to :: "'s \<Rightarrow> 'a list"  where
   "path_to q \<equiv> SOME u. u \<in> left_lang q"
 
 lemma path_to_left_lang: "q \<in> accessible \<Longrightarrow> path_to q \<in> left_lang q"
@@ -137,7 +144,7 @@
   using path_to_left_lang  unfolding left_lang_def
   by auto
 
-definition Accessible_dfa :: "'a dfa" where
+definition Accessible_dfa :: "('a,'s) dfa" where
   "Accessible_dfa = \<lparr>dfa.states = accessible,
                      init  = init M,
                      final = final M \<inter> accessible,
@@ -229,7 +236,7 @@
 subsection\<open>An Equivalence Relation on States\<close>
 
 text\<open>Collapsing map on states. Two states are equivalent if they yield identical outcomes\<close>
-definition eq_right_lang :: "(hf \<times> hf) set" where
+definition eq_right_lang :: "('s \<times> 's) set" where
   "eq_right_lang \<equiv> {(u,v). u \<in> states M \<and> v \<in> states M \<and> right_lang u = right_lang v}"
 
 lemma equiv_eq_right_lang: "equiv (states M) eq_right_lang"
@@ -238,7 +245,16 @@
 lemma eq_right_lang_finite_index: "finite (states M // eq_right_lang)"
   by (metis finite_imageI finite proj_image)
 
-definition Collapse_dfa :: "'a dfa" where
+end
+
+text \<open>Now we need to specialize to typ\<open>hf\<close> states.\<close>
+
+type_synonym 'a dfa_hf = "('a,hf) dfa"
+
+locale dfa_hf = dfa M for M :: "'a dfa_hf"
+begin
+
+definition Collapse_dfa :: "'a dfa_hf" where
   "Collapse_dfa = \<lparr>dfa.states = HF ` (states M // eq_right_lang),
                    init       = HF (eq_right_lang `` {init M}),
                    final      = {HF (eq_right_lang `` {q}) | q. q \<in> final M},
@@ -311,8 +327,8 @@
 
 subsection \<open>Isomorphisms Between DFAs\<close>
 
-locale dfa_isomorphism = M: dfa M + N: dfa N for M :: "'a dfa" and N :: "'a dfa" +
-  fixes h :: "hf \<Rightarrow> hf"
+locale dfa_isomorphism = M: dfa M + N: dfa N for M :: "('a,'sm) dfa" and N :: "('a,'sn) dfa" +
+  fixes h :: "'sm \<Rightarrow> 'sn"
   assumes h: "bij_betw h (states M) (states N)"
       and init  [simp]: "h (init M) = init N"
       and final [simp]: "h ` final M = final N"
@@ -360,10 +376,74 @@
 
 end
 
+text\<open>In order to transition between \<open>'s\<close> and \<open>hf\<close>, we use bijections:\<close>
+
+lemma inj_on_finite_hf:
+  \<open>finite S \<Longrightarrow> \<exists>f:: 's \<Rightarrow> hf.  inj_on f S\<close>
+by (meson comp_inj_on finite_imp_inj_to_nat_seg inj_ord_of)
+
+lemma bij_betw_finite_hf:
+  \<open>finite S \<Longrightarrow> \<exists>f:: 's \<Rightarrow> hf.  bij_betw f S (f ` S)\<close>
+by (meson inj_on_finite_hf inj_on_imp_bij_betw)
+
+text\<open>There is always an isomorphism between a typ\<open>('a,'s) dfa\<close> and a typ\<open>'a dfa_hf\<close>:\<close>
+
+context dfa
+begin
+
+definition bij_s_hf :: "'s \<Rightarrow> hf" where
+"bij_s_hf = (SOME f :: 's \<Rightarrow> hf. bij_betw f (dfa.states M) (f ` dfa.states M))"
+
+lemma bij_betw_bij_s_hf: "bij_betw bij_s_hf (dfa.states M) (bij_s_hf ` dfa.states M)"
+unfolding bij_s_hf_def using bij_betw_finite_hf[OF finite]
+by (metis (no_types, lifting) someI_ex)
+
+abbreviation bij_s_hf_M :: "'a dfa_hf" where
+"bij_s_hf_M \<equiv> \<lparr> dfa.states = bij_s_hf ` dfa.states M,
+         dfa.init  = bij_s_hf (dfa.init M),
+         dfa.final = bij_s_hf ` dfa.final M,
+         dfa.nxt   = (\<lambda>q x. bij_s_hf (dfa.nxt M (the_inv_into (dfa.states M) bij_s_hf q) x)) \<rparr>"
+
+lemma dfa_bij_s_hf_M: "dfa bij_s_hf_M"
+proof (unfold_locales, goal_cases)
+  case 1
+  then show ?case by simp
+next
+  case 2
+  then show ?case using final by auto
+next
+  case (3 q x)
+  then show ?case apply (simp add: bij_betw_def)
+    by (meson bij_betwE bij_betw_bij_s_hf bij_betw_the_inv_into nxt)
+next
+  case 4
+  then show ?case using finite by simp
+qed
+
+interpretation M_iso: dfa_isomorphism M bij_s_hf_M bij_s_hf
+proof(intro dfa_isomorphism.intro dfa_axioms dfa_bij_s_hf_M dfa_isomorphism_axioms.intro, goal_cases)
+  case 1
+  then show ?case using bij_betw_bij_s_hf by simp
+next
+  case 2
+  then show ?case by simp
+next
+  case 3
+  then show ?case by simp
+next
+  case (4 q x)
+  then show ?case
+    by (metis bij_betw_def bij_betw_bij_s_hf dfa.select_convs(4) the_inv_into_f_f)
+qed
+
+lemmas L_M_eq_L_bij_s_hf_M = M_iso.language
+
+end
+
 section\<open>The Myhill-Nerode theorem: three characterisations of a regular language\<close>
 
 definition regular :: "'a list set \<Rightarrow> bool" where
-  "regular L \<equiv> \<exists>M. dfa M \<and> dfa.language M = L"
+  "regular L \<equiv> \<exists>M :: 'a dfa_hf. dfa M \<and> dfa.language M = L"
 
 definition MyhillNerode :: "'a list set \<Rightarrow> ('a list * 'a list) set \<Rightarrow> bool" where
   "MyhillNerode L R \<equiv> equiv UNIV R \<and> right_invariant R \<and> finite (UNIV//R) \<and> (\<exists>A. L = R``A)"
@@ -401,6 +481,11 @@
   by some Myhill-Nerode relation, @{term R}\<close>
 context dfa
 begin
+
+  lemma regular_dfa: "regular language"
+  unfolding regular_def L_M_eq_L_bij_s_hf_M using dfa_bij_s_hf_M
+  by blast
+
   lemma MN_eq_nextl: "MyhillNerode language eq_nextl"
     unfolding MyhillNerode_def
     using language_eq_nextl
@@ -451,7 +536,7 @@
   lemma finix: "finite (UNIV//R)"
     using h bij_betw_finite finite_hfset by blast
 
-  definition DFA :: "'a dfa" where
+  definition DFA :: "'a dfa_hf" where
     "DFA = \<lparr>states = h ` (UNIV//R),
             init  = h (R `` {[]}),
             final = {h (R `` {u}) | u. u \<in> A},
@@ -506,7 +591,7 @@
 
 theorem MN_imp_dfa:
   assumes "MyhillNerode L R"
-  obtains M where "dfa M" "dfa.language M = L" "card (states M) = card (UNIV//R)"
+  obtains M where "dfa_hf M" "dfa.language M = L" "card (states M) = card (UNIV//R)"
 proof -
   from assms obtain A
     where eqR: "equiv UNIV R"
@@ -521,23 +606,22 @@
   interpret MN: MyhillNerode_dfa L R A "?n" h
     by (simp add: MyhillNerode_dfa_def eqR riR L h)
   show ?thesis
-    by (auto simp: MN.language intro: that MN.dfa MN.language MN.card_states)
+    using MN.card_states MN.dfa MN.language dfa_hf.intro that by blast
   qed
 
 corollary MN_imp_regular:
   assumes "MyhillNerode L R"  shows "regular L"
-  using MN_imp_dfa [OF assms] unfolding regular_def
-  by auto
+  unfolding regular_def by (metis dfa_hf_def MN_imp_dfa[OF assms])
 
 corollary eq_app_right_finite_index_imp_dfa:
   assumes "finite (UNIV // eq_app_right L)"
   obtains M where
-    "dfa M" "dfa.language M = L" "card (states M) = card (UNIV // eq_app_right L)"
+    "dfa_hf M" "dfa.language M = L" "card (states M) = card (UNIV // eq_app_right L)"
   using MN_eq_app_right MN_imp_dfa assms by blast
 
 text\<open>Step 3\<close>
 corollary L3_1: "finite (UNIV // eq_app_right L) \<Longrightarrow> regular L"
-  using eq_app_right_finite_index_imp_dfa regular_def by blast
+  unfolding regular_def by (metis eq_app_right_finite_index_imp_dfa dfa_hf_def)
 
 
 section\<open>Non-Deterministic Finite Automata\<close>
@@ -546,14 +630,14 @@
 
 subsection\<open>Basic Definitions\<close>
 
-record 'a nfa = states :: "hf set"
-                init   :: "hf set"
-                final  :: "hf set"
-                nxt    :: "hf \<Rightarrow> 'a \<Rightarrow> hf set"
-                eps    :: "(hf * hf) set"
+record ('a,'s) nfa = states :: "'s set"
+                init   :: "'s set"
+                final  :: "'s set"
+                nxt    :: "'s \<Rightarrow> 'a \<Rightarrow> 's set"
+                eps    :: "('s * 's) set"
 
 locale nfa =
-  fixes M :: "'a nfa"
+  fixes M :: "('a,'s) nfa"
   assumes init: "init M \<subseteq> states M"
       and final: "final M \<subseteq> states M"
       and nxt:   "\<And>q x. q \<in> states M \<Longrightarrow> nxt M q x \<subseteq> states M"
@@ -563,7 +647,7 @@
 lemma subset_states_finite [intro,simp]: "Q \<subseteq> states M \<Longrightarrow> finite Q"
   by (simp add: finite_subset finite)
 
-definition epsclo :: "hf set \<Rightarrow> hf set" where
+definition epsclo :: "'s set \<Rightarrow> 's set" where
   "epsclo Q \<equiv> states M \<inter> (\<Union>q\<in>Q. {q'. (q,q') \<in> (eps M)*})"
 
 lemma epsclo_eq_Image: "epsclo Q = states M \<inter> (eps M)* `` Q"
@@ -603,7 +687,7 @@
   by (metis finite_subset finite nxt)
 
 text\<open>Transition function for a given starting state and word.\<close>
-primrec nextl :: "[hf set, 'a list] \<Rightarrow> hf set" where
+primrec nextl :: "['s set, 'a list] \<Rightarrow> 's set" where
     "nextl Q []     = epsclo Q"
   | "nextl Q (x#xs) = nextl (\<Union>q \<in> epsclo Q. nxt M q x) xs"
 
@@ -611,7 +695,7 @@
   "language \<equiv> {xs. nextl (init M) xs \<inter> final M \<noteq> {}}"
 
 text\<open>The right language WRT a state q is the set of words that go from q to F.\<close>
-definition right_lang :: "hf \<Rightarrow> 'a list set"  where
+definition right_lang :: "'s \<Rightarrow> 'a list set"  where
   "right_lang q \<equiv> {u. nextl {q} u \<inter> final M \<noteq> {}}"
 
 lemma nextl_epsclo [simp]: "nextl (epsclo Q) xs = nextl Q xs"
@@ -647,24 +731,33 @@
 lemma nextl_UN: "nextl (\<Union>i\<in>I. f i) xs = (\<Union>i\<in>I. nextl (f i) xs)"
   by (induct xs arbitrary: f) auto
 
+end
+
 subsection\<open>The Powerset Construction\<close>
 
-definition Power_dfa :: "'a dfa" where
-  "Power_dfa = \<lparr>dfa.states = {HF (epsclo q) | q. q \<in> Pow (states M)},
-                     init  = HF (epsclo (init M)),
-                     final = {HF (epsclo Q) | Q. Q \<subseteq> states M \<and> Q \<inter> final M \<noteq> {}},
-                     nxt   = \<lambda>Q x. HF(\<Union>q \<in> epsclo (hfset Q). epsclo (nxt M q x))\<rparr>"
+text \<open>First: The construction of a typ\<open>('a,'s set) dfa\<close> from an typ\<open>('a,'s) nfa\<close>.
+Is not used later on but provides an easy means of showing regularity of some language
+by constructing an NFA without having to use \<open>hf\<close>.\<close>
+
+context nfa
+begin
 
-lemma states_Power_dfa [simp]: "dfa.states Power_dfa = HF ` epsclo ` Pow (states M)"
+definition Power_dfa :: "('a,'s set) dfa" where
+  "Power_dfa = \<lparr>dfa.states = {epsclo q | q. q \<in> Pow (states M)},
+                     init  = epsclo (init M),
+                     final = {epsclo Q | Q. Q \<subseteq> states M \<and> Q \<inter> final M \<noteq> {}},
+                     nxt   = \<lambda>Q x. \<Union>q \<in> epsclo Q. epsclo (nxt M q x)\<rparr>"
+
+lemma states_Power_dfa [simp]: "dfa.states Power_dfa = epsclo ` Pow (states M)"
   by (auto simp add: Power_dfa_def)
 
-lemma init_Power_dfa [simp]: "dfa.init Power_dfa = HF (epsclo (nfa.init M))"
+lemma init_Power_dfa [simp]: "dfa.init Power_dfa = epsclo (nfa.init M)"
   by (simp add: Power_dfa_def)
 
-lemma final_Power_dfa [simp]: "dfa.final Power_dfa = {HF (epsclo Q) | Q. Q \<subseteq> states M \<and> Q \<inter> final M \<noteq> {}}"
+lemma final_Power_dfa [simp]: "dfa.final Power_dfa = {epsclo Q | Q. Q \<subseteq> states M \<and> Q \<inter> final M \<noteq> {}}"
   by (simp add: Power_dfa_def)
 
-lemma nxt_Power_dfa [simp]: "dfa.nxt Power_dfa = (\<lambda>Q x. HF(\<Union>q \<in> epsclo (hfset Q). epsclo (nxt M q x)))"
+lemma nxt_Power_dfa [simp]: "dfa.nxt Power_dfa = (\<lambda>Q x. \<Union>q \<in> epsclo Q. epsclo (nxt M q x))"
   by (simp add: Power_dfa_def)
 
 interpretation Power: dfa Power_dfa
@@ -678,8 +771,6 @@
   fix q a
   show "q \<in> dfa.states Power_dfa \<Longrightarrow> dfa.nxt Power_dfa q a \<in> dfa.states Power_dfa"
     apply (auto simp: nxt)
-    apply (subst inj_on_image_mem_iff [OF inj_on_HF])
-    apply (auto simp: rev_finite_subset [OF finite] nxt)
     apply (metis Pow_iff epsclo_UN epsclo_idem epsclo_subset image_eqI)
     done
 next
@@ -687,12 +778,90 @@
     by (force simp: finite)
 qed
 
-corollary dfa_Power: "dfa Power_dfa"
+text\<open>The Power DFA accepts the same language as the NFA.\<close>
+theorem Power_language [simp]: "Power.language = language"
+proof -
+  { fix u
+    have "(Power.nextl (dfa.init Power_dfa) u) = nextl (init M) u"
+    proof (induct u rule: List.rev_induct)
+      case Nil show ?case
+        using Power.nextl.simps
+        by (auto simp: hinsert_def)
+    next
+      case (snoc x u) then show ?case
+        by (simp add: init finite_nextl nextl_state [THEN subsetD])
+    qed
+    then have "u \<in> Power.language \<longleftrightarrow> u \<in> language"
+      apply (auto simp add: Power.language_def language_def disjoint_iff_not_equal)
+      apply (metis Int_iff epsclo_increasing subsetCE)
+      apply (metis epsclo_nextl nextl_state)
+      done
+  }
+  then show ?thesis
+    by blast
+qed
+
+text\<open>Every language accepted by a NFA is also accepted by a DFA.\<close>
+corollary imp_regular: "regular language"
+  by (metis Power.regular_dfa Power_language)
+
+end
+
+text\<open>As above, outside the locale\<close>
+corollary nfa_imp_regular:
+  assumes "nfa M" "nfa.language M = L"
+    shows "regular L"
+using assms nfa.imp_regular by blast
+
+
+type_synonym 'a nfa_hf = "('a,hf) nfa"
+
+text \<open>The construction of a typ\<open>'a dfa_hf\<close> from an typ\<open>'a nfa_hf\<close>.
+Very little can be reused from the generic \<open>'s\<close>-based construction above.\<close>
+
+locale nfa_hf = nfa M for M :: "'a nfa_hf"
+begin
+
+definition Power_dfa_hf :: "'a dfa_hf" where
+  "Power_dfa_hf = \<lparr>dfa.states = HF ` dfa.states Power_dfa,
+                     init  = HF (dfa.init Power_dfa),
+                     final = HF ` dfa.final Power_dfa,
+                     nxt   = \<lambda>Q. HF o dfa.nxt Power_dfa (hfset Q)\<rparr>"
+
+lemma states_Power_dfa [simp]: "dfa.states Power_dfa_hf = HF ` epsclo ` Pow (states M)"
+  by (auto simp add: Power_dfa_hf_def)
+
+lemma init_Power_dfa [simp]: "dfa.init Power_dfa_hf = HF (dfa.init Power_dfa)"
+  by (simp add: Power_dfa_hf_def)
+
+lemma final_Power_dfa [simp]: "dfa.final Power_dfa_hf = {HF (epsclo Q) | Q. Q \<subseteq> states M \<and> Q \<inter> final M \<noteq> {}}"
+  by (auto simp add: Power_dfa_hf_def)
+
+lemma nxt_Power_dfa [simp]: "dfa.nxt Power_dfa_hf = (\<lambda>Q x. HF(\<Union>q \<in> epsclo (hfset Q). epsclo (nxt M q x)))"
+  by (simp add: Power_dfa_hf_def o_def)
+
+interpretation Power: dfa Power_dfa_hf
+proof (unfold_locales, goal_cases)
+  case 1 thus ?case
+    by (simp add: init)
+next
+  case 2 thus ?case
+    by auto
+next
+  case (3 q a) thus ?case
+    unfolding dfa_def nxt_Power_dfa
+    by (metis Pow_iff epsclo_UN epsclo_idem epsclo_subset imageI states_Power_dfa)
+next
+  case 4 thus ?case
+    by (auto simp: finite)
+qed
+
+corollary dfa_Power: "dfa Power_dfa_hf"
   by unfold_locales
 
 lemma nextl_Power_dfa:
-     "qs \<in> dfa.states Power_dfa
-     \<Longrightarrow> dfa.nextl Power_dfa qs u = HF (\<Union>q \<in> hfset qs. nextl {q} u)"
+     "qs \<in> dfa.states Power_dfa_hf
+     \<Longrightarrow> dfa.nextl Power_dfa_hf qs u = HF (\<Union>q \<in> hfset qs. nextl {q} u)"
   apply (induct u rule: List.rev_induct)
   apply (auto simp: finite_nextl inj_on_HF [THEN inj_on_eq_iff])
   apply (metis Int_empty_left Int_insert_left_if1 epsclo_increasing epsclo_subset subsetD singletonI)
@@ -701,7 +870,7 @@
 
 text\<open>Part of Prop 4 of Jean-Marc Champarnaud, A. Khorsi and T. Paranthoën (2002)\<close>
 lemma Power_right_lang:
-     "qs \<in> dfa.states Power_dfa \<Longrightarrow> Power.right_lang qs = (\<Union>q \<in> hfset qs. right_lang q)"
+     "qs \<in> dfa.states Power_dfa_hf \<Longrightarrow> Power.right_lang qs = (\<Union>q \<in> hfset qs. right_lang q)"
 using epsclo_increasing
 apply (auto simp: Power.right_lang_def right_lang_def nextl_Power_dfa
                   inj_on_HF [THEN inj_on_eq_iff] finite_nextl, blast)
@@ -716,7 +885,7 @@
 theorem Power_language [simp]: "Power.language = language"
 proof -
   { fix u
-    have "(Power.nextl (dfa.init Power_dfa) u) = HF (nextl (init M) u)"
+    have "(Power.nextl (dfa.init Power_dfa_hf) u) = HF (nextl (init M) u)"
     proof (induct u rule: List.rev_induct)
       case Nil show ?case
         using Power.nextl.simps
@@ -735,18 +904,8 @@
     by blast
 qed
 
-text\<open>Every language accepted by a NFA is also accepted by a DFA.\<close>
-corollary imp_regular: "regular language"
-  using Power_language dfa_Power regular_def by blast
-
 end
 
-text\<open>As above, outside the locale\<close>
-corollary nfa_imp_regular:
-  assumes "nfa M" "nfa.language M = L"
-    shows "regular L"
-using assms nfa.imp_regular by blast
-
 
 section\<open>Closure Properties for Regular Languages\<close>
 
@@ -754,8 +913,8 @@
 
 theorem regular_empty:  "regular {}"
 proof -
-  interpret D: dfa "\<lparr>dfa.states = {0}, init = 0, final = {}, nxt = \<lambda>q x. q\<rparr>"
-    by (auto simp: dfa_def)
+  interpret D: dfa_hf "\<lparr>dfa.states = {0}, init = 0, final = {}, nxt = \<lambda>q x. q\<rparr>"
+    by (auto simp: dfa_hf_def dfa_def)
   have "D.language = {}"
     by (simp add: D.language_def)
   then show ?thesis
@@ -767,8 +926,8 @@
 
 theorem regular_nullstr:  "regular {[]}"
 proof -
-  interpret N: nfa "\<lparr>states = {0}, init = {0}, final = {0}, nxt = \<lambda>q x. {}, eps = {} \<rparr>"
-    by (auto simp: nfa_def)
+  interpret N: nfa_hf "\<lparr>states = {0}, init = {0}, final = {0}, nxt = \<lambda>q x. {}, eps = {} \<rparr>"
+    by (auto simp: nfa_hf_def nfa_def)
   have "\<And>u. 0 \<in> N.nextl {0} u \<Longrightarrow> u = []"
     by (rule list.exhaust, auto)
   then have "N.language = {[]}"
@@ -783,9 +942,9 @@
 proof -
   let ?N = "\<lparr>states = {0,1}, init = {0}, final = {1},
              nxt = \<lambda>q x. if q=0 \<and> x=a then {1} else {},
-             eps = {}\<rparr>"
-  interpret N: nfa ?N
-    by (auto simp: nfa_def)
+             eps = {}\<rparr> :: 'a nfa_hf"
+  interpret N: nfa_hf ?N
+    by (auto simp: nfa_def nfa_hf_def)
   have [intro]: "\<And>u. 1 \<in> N.nextl {1} u \<Longrightarrow> u = []"
       by (rule list.exhaust) auto
   { fix u
@@ -803,8 +962,9 @@
 theorem regular_Compl:
   assumes S: "regular S" shows "regular (-S)"
 proof -
-  obtain MS  where M: "dfa MS" and lang: "dfa.language MS = S"
-    using S by (auto simp: regular_def)
+  obtain MS where M: "dfa_hf MS" and lang: "dfa.language MS = S"
+    using S by (auto simp: regular_def dfa_hf_def)
+  note M = dfa_hf.axioms[OF M]
   interpret ST: dfa "\<lparr>dfa.states= dfa.states MS,
                      init= dfa.init MS, final= dfa.states MS - dfa.final MS,
                      nxt= \<lambda>q x. dfa.nxt MS q x\<rparr>"
@@ -830,8 +990,9 @@
 theorem regular_Int:
   assumes S: "regular S" and T: "regular T" shows "regular (S \<inter> T)"
 proof -
-  obtain MS MT where M: "dfa MS" "dfa MT" and lang: "dfa.language MS = S" "dfa.language MT = T"
-    using S T by (auto simp: regular_def)
+  obtain MS MT where M: "dfa_hf MS" "dfa_hf MT" and lang: "dfa.language MS = S" "dfa.language MT = T"
+    using S T by (auto simp: regular_def dfa_hf_def)
+  note M = dfa_hf.axioms[OF M(1)] dfa_hf.axioms[OF M(2)]
   interpret ST: dfa "\<lparr>dfa.states = {\<langle>q1,q2\<rangle> | q1 q2. q1 \<in> dfa.states MS \<and> q2 \<in> dfa.states MT},
                      init       = \<langle>dfa.init MS, dfa.init MT\<rangle>,
                      final      = {\<langle>q1,q2\<rangle> | q1 q2. q1 \<in> dfa.final MS \<and> q2 \<in> dfa.final MT},
@@ -872,8 +1033,9 @@
 theorem regular_conc:
   assumes S: "regular S" and T: "regular T" shows "regular (S @@ T)"
 proof -
-  obtain MS MT where M: "dfa MS" "dfa MT" and lang: "dfa.language MS = S" "dfa.language MT = T"
-    using S T by (auto simp: regular_def)
+  obtain MS MT where M: "dfa_hf MS" "dfa_hf MT" and lang: "dfa.language MS = S" "dfa.language MT = T"
+    using S T by (auto simp: regular_def dfa_hf_def)
+  note M = dfa_hf.axioms[OF M(1)] dfa_hf.axioms[OF M(2)]
   note [simp] = dfa.init dfa.nxt dfa.nextl.simps dfa.nextl_snoc
   let ?ST = "\<lparr>nfa.states = Inl ` (dfa.states MS) \<union> Inr ` (dfa.states MT),
                   init  = {Inl (dfa.init MS)},
@@ -931,7 +1093,7 @@
     using lang unfolding conc_def
     by blast
   then show ?thesis
-    by (metis ST.imp_regular)
+    using ST.nfa_axioms nfa_hf_def nfa_imp_regular by blast
 qed
 
 lemma regular_word: "regular {u}"
@@ -957,8 +1119,9 @@
 theorem regular_star:
   assumes S: "regular S" shows "regular (star S)"
 proof -
-  obtain MS where M: "dfa MS" and lang: "dfa.language MS = S"
-    using S by (auto simp: regular_def)
+  obtain MS where M: "dfa_hf MS" and lang: "dfa.language MS = S"
+    using S by (auto simp: regular_def dfa_hf_def)
+  note M = dfa_hf.axioms[OF M]
   note [simp] = dfa.init [OF M] dfa.nextl.simps [OF M] dfa.nextl_snoc [OF M]
   obtain q0 where q0: "q0 \<notin> dfa.states MS" using dfa.finite [OF M]
     by (metis hdomain_not_mem hfset_HF hmem_def)
@@ -967,7 +1130,7 @@
   have [simp]: "\<And>q x. q \<in> dfa.states MS \<Longrightarrow> q0 \<noteq> dfa.nxt MS q x"
     using M dfa.nxt q0 by fastforce
   have [simp]: "\<And>q u. q \<in> dfa.states MS \<Longrightarrow> q0 \<noteq> dfa.nextl MS q u"
-    using M dfa.nextl_state q0 by blast
+    using M dfa.nextl_state q0 by metis
   let ?ST = "\<lparr>nfa.states = insert q0 (dfa.states MS),
                    init  = {q0},
                    final = {q0},
@@ -1048,12 +1211,12 @@
     using lang
     by (auto simp: ST.language_def star_dfa_in_ST)
   then show ?thesis
-    by (metis ST.imp_regular)
+    using ST.nfa_axioms nfa_hf.intro nfa_imp_regular by blast
 qed
 
 subsection\<open>The Reversal of a Regular Language\<close>
 
-definition Reverse_nfa :: "'a dfa \<Rightarrow> 'a nfa" where
+definition Reverse_nfa :: "'a dfa_hf \<Rightarrow> 'a nfa_hf" where
   "Reverse_nfa MS = \<lparr>nfa.states = dfa.states MS,
                          init  = dfa.final MS,
                          final = {dfa.init MS},
@@ -1076,7 +1239,7 @@
 lemma eps_Reverse_nfa [simp]: "eps (Reverse_nfa MS) = {}"
   by (simp add: Reverse_nfa_def)
 
-context dfa
+context dfa_hf
 begin
 
   lemma nfa_Reverse_nfa: "nfa (Reverse_nfa M)"
@@ -1130,13 +1293,12 @@
 corollary regular_Reverse:
   assumes S: "regular S" shows "regular (rev ` S)"
 proof -
-  obtain MS where MS: "dfa MS" "dfa.language MS = S"
-    using S by (auto simp: regular_def)
-  then interpret dfa "MS"
+  obtain MS where MS: "dfa_hf MS" "dfa.language MS = S"
+    using S by (auto simp: regular_def dfa_hf_def)
+  then interpret dfa_hf "MS"
     by simp
   show ?thesis
-    using nfa_Reverse_nfa nfa_imp_regular language_Reverse_nfa MS
-    by blast
+    by (metis MS(2) language_Reverse_nfa nfa_Reverse_nfa nfa_imp_regular)
 qed
 
 
@@ -1148,7 +1310,7 @@
 
 section\<open>Brzozowski's Minimization Algorithm\<close>
 
-context dfa
+context dfa_hf
   begin
 
 subsection\<open>More about the relation @{term eq_app_right}\<close>
@@ -1272,40 +1434,40 @@
   qed
 
   lemma states_PR [simp]:
-       "dfa.states (nfa.Power_dfa (Reverse_nfa M)) = HF ` Pow (dfa.states M)"
+       "dfa.states (nfa_hf.Power_dfa_hf (Reverse_nfa M)) = HF ` Pow (dfa.states M)"
     by (rule set_eqI)
-       (auto simp: nfa.states_Power_dfa nfa_Reverse_nfa image_iff Bex_def)
+       (auto simp: nfa_hf.states_Power_dfa nfa_Reverse_nfa  nfa_hf.intro Bex_def)
 
   lemma inj_on_right_lang_PR:
     assumes "dfa.states M = accessible"
-      shows "inj_on (dfa.right_lang (nfa.Power_dfa (Reverse_nfa M)))
-                    (dfa.states (nfa.Power_dfa (Reverse_nfa M)))"
+      shows "inj_on (dfa.right_lang (nfa_hf.Power_dfa_hf (Reverse_nfa M)))
+                    (dfa.states (nfa_hf.Power_dfa_hf (Reverse_nfa M)))"
   proof (rule inj_onI)
     fix q1 q2
-    assume *: "q1 \<in> dfa.states (nfa.Power_dfa (Reverse_nfa M))"
-              "q2 \<in> dfa.states (nfa.Power_dfa (Reverse_nfa M))"
-              "dfa.right_lang (nfa.Power_dfa (Reverse_nfa M)) q1 =
-               dfa.right_lang (nfa.Power_dfa (Reverse_nfa M)) q2"
+    assume *: "q1 \<in> dfa.states (nfa_hf.Power_dfa_hf (Reverse_nfa M))"
+              "q2 \<in> dfa.states (nfa_hf.Power_dfa_hf (Reverse_nfa M))"
+              "dfa.right_lang (nfa_hf.Power_dfa_hf (Reverse_nfa M)) q1 =
+               dfa.right_lang (nfa_hf.Power_dfa_hf (Reverse_nfa M)) q2"
     then have "hfset q1 \<subseteq> accessible \<and> hfset q2 \<subseteq> accessible"
       using assms rev_finite_subset [OF finite]
       by force
     with * show "q1 = q2"
-      apply (simp add: nfa_Reverse_nfa nfa.Power_right_lang right_lang_Reverse
-                       image_UN [symmetric] inj_image_eq_iff)
+      apply (simp add: nfa_Reverse_nfa nfa_hf.Power_right_lang right_lang_Reverse
+                       image_UN [symmetric] inj_image_eq_iff nfa_hf_def)
       apply (metis HF_hfset le_sup_iff left_lang_UN)
       done
   qed
 
-  abbreviation APR :: "'x dfa \<Rightarrow> 'x dfa" where
-    "APR X \<equiv> dfa.Accessible_dfa (nfa.Power_dfa (Reverse_nfa X))"
+  abbreviation APR :: "'x dfa_hf \<Rightarrow> 'x dfa_hf" where
+    "APR X \<equiv> dfa.Accessible_dfa (nfa_hf.Power_dfa_hf (Reverse_nfa X))"
 
   theorem minimal_APR:
     assumes "dfa.states M = accessible"
       shows "dfa.minimal (APR M)"
     proof -
       have PR: "dfa (APR M)"
-               "dfa (nfa.Power_dfa (Reverse_nfa M))"
-        by (auto simp: dfa.dfa_Accessible nfa.dfa_Power nfa_Reverse_nfa)
+               "dfa (nfa_hf.Power_dfa_hf (Reverse_nfa M))"
+        by (auto simp add: dfa.dfa_Accessible nfa_Reverse_nfa nfa_hf.dfa_Power nfa_hf.intro)
       then show ?thesis
         apply (simp add: dfa.minimal_def dfa.states_Accessible_dfa dfa.Accessible_accessible)
         apply (simp add: inj_on_def dfa.Accessible_right_lang_eq)
@@ -1313,39 +1475,38 @@
         done
     qed
 
-  definition Brzozowski :: "'a dfa" where
+  definition Brzozowski :: "'a dfa_hf" where
     "Brzozowski \<equiv> APR (APR M)"
 
-  lemma dfa_Brzozowski: "dfa Brzozowski"
-    by (simp add: Brzozowski_def dfa.dfa_Accessible dfa.nfa_Reverse_nfa
-                  nfa.dfa_Power nfa_Reverse_nfa)
+  lemma dfa_Brzozowski: "dfa_hf Brzozowski"
+    by (simp add: Brzozowski_def dfa.dfa_Accessible dfa_hf.nfa_Reverse_nfa
+                  nfa_hf.dfa_Power nfa_Reverse_nfa dfa_hf_def nfa_hf_def)
 
   theorem language_Brzozowski: "dfa.language Brzozowski = language"
-    by (simp add: Brzozowski_def dfa.Accessible_language nfa.Power_language
-          dfa.dfa_Accessible dfa.nfa_Reverse_nfa nfa.dfa_Power nfa_Reverse_nfa
-          dfa.language_Reverse_nfa image_image)
+    by (simp add: Brzozowski_def dfa.Accessible_language nfa_hf.Power_language
+          dfa.dfa_Accessible dfa_hf.nfa_Reverse_nfa nfa_hf.dfa_Power nfa_Reverse_nfa
+          dfa_hf.language_Reverse_nfa image_image dfa_hf_def nfa_hf_def)
 
   theorem minimal_Brzozowski: "dfa.minimal Brzozowski"
   unfolding Brzozowski_def
-  proof (rule dfa.minimal_APR)
-    show "dfa (APR M)"
-      by (simp add: dfa.dfa_Accessible nfa.dfa_Power nfa_Reverse_nfa)
+  proof (rule dfa_hf.minimal_APR)
+    show "dfa_hf (APR M)"
+      by (simp add: dfa.dfa_Accessible nfa_hf.dfa_Power nfa_Reverse_nfa dfa_hf_def nfa_hf_def)
   next
     show "dfa.states (APR M) = dfa.accessible (APR M)"
-      by (simp add: dfa.Accessible_accessible dfa.states_Accessible_dfa nfa.dfa_Power nfa_Reverse_nfa)
+      by (simp add: dfa.Accessible_accessible dfa.states_Accessible_dfa nfa_hf.dfa_Power nfa_Reverse_nfa nfa_hf_def)
   qed
 
   end
 
 lemma index_f_cong:
-     "\<lbrakk>dfa.language M = dfa.language N; dfa M; dfa N\<rbrakk> \<Longrightarrow> dfa.index_f M = dfa.index_f N"
-  by (simp add: dfa.index_f_def dfa.min_states_def)
+     "\<lbrakk>dfa.language M = dfa.language N; dfa M; dfa N\<rbrakk> \<Longrightarrow> dfa_hf.index_f M = dfa_hf.index_f N"
+  by (simp add: dfa_hf.index_f_def dfa_hf.min_states_def dfa_hf_def)
 
 theorem minimal_imp_isomorphic:
-     "\<lbrakk>dfa.language M = dfa.language N; dfa.minimal M; dfa.minimal N; dfa M; dfa N\<rbrakk>
+     "\<lbrakk>dfa.language M = dfa.language N; dfa.minimal M; dfa.minimal N; dfa_hf M; dfa_hf N\<rbrakk>
       \<Longrightarrow> \<exists>h. dfa_isomorphism M N h"
-  by (metis dfa_isomorphism.sym dfa_isomorphism.trans
-            dfa.minimal_imp_isomorphic_to_canonical index_f_cong)
+  by (metis dfa_hf.minimal_imp_isomorphic_to_canonical dfa_hf_def dfa_isomorphism.sym
+    dfa_isomorphism.trans index_f_cong)
 
 end
-
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/After_Operator_Non_Too_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/After_Operator_Non_Too_Destructive.thy
@@ -0,0 +1,202 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Non too Destructiveness of After\<close>
+
+(*<*)
+theory After_Operator_Non_Too_Destructive
+  imports Process_Restriction_Space
+    "HOL-CSP_OpSem.After_Trace_Operator" 
+begin
+  (*>*)
+
+
+subsection \<open>Equality\<close>
+
+lemma initials_restriction_processptick: \<open>(P \<down> n)0 = (if n = 0 then UNIV else P0)\<close>
+  by (cases n, solves simp)
+    (auto simp add: initials_def T_restriction_processptick,
+      metis append.right_neutral append_eq_conv_conj drop_Nil drop_Suc_Cons)
+
+
+lemma (in After) restriction_processptick_After:
+  \<open>P after e \<down> n = (if ev e \<in> P0 then (P \<down> Suc n) after e else \<Psi> P e \<down> n)\<close>
+proof (split if_split, intro conjI impI)
+  show \<open>ev e \<notin> P0 \<Longrightarrow> P after e \<down> n = \<Psi> P e \<down> n\<close> by (simp add: not_initial_After)
+next
+  assume \<open>ev e \<in> P0\<close>
+  show \<open>P after e \<down> n = (P \<down> Suc n) after e\<close> (is \<open>?lhs = ?rhs\<close>)
+  proof (subst Process_eq_spec, safe)
+    show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+      by (elim D_restriction_processptickE)
+        (simp_all add: \<open>ev e \<in> P0\<close> After_projs
+          initials_restriction_processptick D_restriction_processptick,
+          meson Cons_eq_appendI eventptick.disc(1) length_Cons tickFree_Cons_iff)
+  next
+    show \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+      by (auto simp add: D_After initials_restriction_processptick \<open>ev e \<in> P0\<close>
+          D_restriction_processptick T_After Cons_eq_append_conv)
+  next
+    show \<open>(t, X) \<in> \<F> ?lhs \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for t X
+      by (elim F_restriction_processptickE)
+        (simp_all add: \<open>ev e \<in> P0\<close> After_projs
+          initials_restriction_processptick F_restriction_processptick,
+          meson Cons_eq_appendI eventptick.disc(1) length_Cons tickFree_Cons_iff)
+  next
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+      by (auto simp add: F_After initials_restriction_processptick \<open>ev e \<in> P0\<close>
+          F_restriction_processptick T_After Cons_eq_append_conv)
+  qed
+qed
+
+lemma (in AfterExt) restriction_processptick_Aftertick:
+  \<open>P after\<checkmark> e \<down> n =
+   (case e of \<checkmark>(r) \<Rightarrow> \<Omega> P r \<down> n | ev x \<Rightarrow> if e \<in> P0 then (P \<down> Suc n) after\<checkmark> e else \<Psi> P x \<down> n)\<close>
+  by (simp add: Aftertick_def restriction_processptick_After split: eventptick.split )
+
+lemma (in AfterExt) restriction_processptick_Aftertrace:
+  \<open>t \<in> \<T> P \<Longrightarrow> tF t \<Longrightarrow> P after\<T> t \<down> n = (P \<down> (n + length t)) after\<T> t\<close>
+proof (induct t arbitrary: n rule: rev_induct)
+  show \<open>P after\<T> [] \<down> n = (P \<down> (n + length [])) after\<T> []\<close> for n by simp
+next
+  fix e t n
+  assume   hyp : \<open>t \<in> \<T> P \<Longrightarrow> tF t \<Longrightarrow> P after\<T> t \<down> n = (P \<down> (n + length t)) after\<T> t\<close> for n
+  assume prems : \<open>t @ [e] \<in> \<T> P\<close> \<open>tickFree (t @ [e])\<close>
+  from prems(2) obtain a where \<open>e = ev a\<close> by (cases e) simp_all
+  with initials_Aftertrace[OF prems(1)] have \<open>ev a \<in> (P after\<T> t)0\<close> by simp
+  from prems is_processT3_TR_append have \<open>t \<in> \<T> P\<close> \<open>tF t\<close> by auto
+  from hyp[OF this] have \<open>P after\<T> t \<down> Suc n = (P \<down> (Suc n + length t)) after\<T> t\<close> .
+  thus \<open>P after\<T> (t @ [e]) \<down> n = (P \<down> (n + length (t @ [e]))) after\<T> (t @ [e])\<close>
+    by (simp add: Aftertrace_snoc restriction_processptick_Aftertick
+        \<open>e = ev a\<close> \<open>ev a \<in> (P after\<T> t)0\<close>)
+qed
+
+
+
+subsection \<open>Non too Destructiveness\<close>
+
+lemma (in After) non_too_destructive_on_After :
+  \<open>non_too_destructive_on (\<lambda>P. P after e) {P. ev e \<in> P0}\<close>
+  by (auto intro!: non_too_destructive_onI simp add: restriction_processptick_After)
+
+lemma (in AfterExt) non_too_destructive_on_Aftertick :
+  \<open>non_too_destructive_on (\<lambda>P. P after\<checkmark> e) {P. e \<in> P0}\<close>
+  if \<open>\<And>r. e = \<checkmark>(r) \<Longrightarrow> non_too_destructive_on \<Omega> {P. \<checkmark>(r) \<in> P0}\<close>
+proof (intro non_too_destructive_onI, clarify)
+  fix P Q n assume * : \<open>P \<down> Suc n = Q \<down> Suc n\<close> \<open>e \<in> P0\<close> \<open>e \<in> Q0\<close>
+  show \<open>P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close>
+  proof (cases e)
+    from "*" show \<open>e = ev a \<Longrightarrow> P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close> for a
+      by (simp add: restriction_processptick_Aftertick)
+  next
+    fix r assume \<open>e = \<checkmark>(r)\<close>
+    with "*"(2, 3) have \<open>P \<in> {P. \<checkmark>(r) \<in> P0}\<close> \<open>Q \<in> {P. \<checkmark>(r) \<in> P0}\<close> by auto
+    from non_too_destructive_onD[OF that[simplified \<open>e = \<checkmark>(r)\<close>, OF refl] this "*"(1)]
+    have \<open>\<Omega> P \<down> n = \<Omega> Q \<down> n\<close> .
+    with \<open>e = \<checkmark>(r)\<close> show \<open>P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close>
+      by (simp add: restriction_processptick_Aftertick) (metis restriction_fun_def)
+  qed
+qed
+
+
+
+lemma (in After) non_too_destructive_After :
+  \<open>non_too_destructive (\<lambda>P. P after e)\<close> if * : \<open>non_too_destructive_on \<Psi> {P. ev e \<notin> P0}\<close>
+proof (rule non_too_destructiveI)
+  fix P Q :: \<open>('a, 'r) processptick\<close> and n
+  assume \<open>P \<down> Suc n = Q \<down> Suc n\<close>
+  hence \<open>ev e \<in> P0 \<and> ev e \<in> Q0 \<or> ev e \<notin> P0 \<and> ev e \<notin> Q0\<close>
+    by (metis initials_restriction_processptick nat.distinct(1))
+  thus \<open>P after e \<down> n = Q after e \<down> n\<close>
+  proof (elim disjE conjE)
+    show \<open>ev e \<in> P0 \<Longrightarrow> ev e \<in> Q0 \<Longrightarrow> P after e \<down> n = Q after e \<down> n\<close>
+      by (simp add: restriction_processptick_After \<open>P \<down> Suc n = Q \<down> Suc n\<close>)
+  next
+    assume \<open>ev e \<notin> P0\<close> \<open>ev e \<notin> Q0\<close>
+    hence \<open>P after e = \<Psi> P e\<close> \<open>Q after e = \<Psi> Q e\<close>
+      by (simp_all add: not_initial_After)
+    from \<open>P \<down> Suc n = Q \<down> Suc n\<close> have \<open>\<Psi> P \<down> n = \<Psi> Q \<down> n\<close>
+      by (intro "*"[THEN non_too_destructive_onD, of P Q])
+        (simp_all add: \<open>ev e \<notin> P0\<close> \<open>ev e \<notin> Q0\<close>)
+    with \<open>P after e = \<Psi> P e\<close> \<open>Q after e = \<Psi> Q e\<close>
+    show \<open>P after e \<down> n = Q after e \<down> n\<close> 
+      by (metis restriction_fun_def)
+  qed
+qed
+
+
+lemma (in AfterExt) non_too_destructive_Aftertick :
+  \<open>non_too_destructive (\<lambda>P. P after\<checkmark> e)\<close>
+  if * : \<open>\<And>a. e = ev a \<Longrightarrow> non_too_destructive_on \<Psi> {P. ev a \<notin> P0}\<close>
+    \<open>\<And>r. e = \<checkmark>(r) \<Longrightarrow> non_too_destructive (\<lambda>P. \<Omega> P r)\<close>
+proof (rule non_too_destructiveI)
+  show \<open>P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close> if \<open>P \<down> Suc n = Q \<down> Suc n\<close> for P Q n
+  proof (cases e)
+    show \<open>P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close> if \<open>e = ev a\<close> for a
+      by (simp add: Aftertick_def \<open>e = ev a\<close>)
+        (fact non_too_destructive_After[OF "*"(1)[simplified \<open>e = ev a\<close>, OF refl],
+            THEN non_too_destructiveD, OF \<open>P \<down> Suc n = Q \<down> Suc n\<close>])
+  next
+    fix r assume \<open>e = \<checkmark>(r)\<close>
+    from "*"(2)[unfolded \<open>e = \<checkmark>(r)\<close>, OF refl,
+        THEN non_too_destructiveD, OF \<open>P \<down> Suc n = Q \<down> Suc n\<close>]
+    have \<open>\<Omega> P r \<down> n = \<Omega> Q r \<down> n\<close> .
+    thus \<open>P after\<checkmark> e \<down> n = Q after\<checkmark> e \<down> n\<close>
+      by (simp add: Aftertick_def \<open>e = \<checkmark>(r)\<close>)
+  qed
+qed
+
+
+lemma (in AfterExt) restriction_shift_Aftertrace :
+  \<open>restriction_shift (\<lambda>P. P after\<T> t) (- int (length t))\<close>
+  if \<open>non_too_destructive \<Psi>\<close> \<open>non_too_destructive \<Omega>\<close>
+    \<comment> \<open>We could imagine more precise assumptions, but is it useful?\<close>
+proof (induct t)
+  case Nil show ?case by (simp add: restriction_shiftI)
+next
+  case (Cons e t)
+  have \<open>non_too_destructive (\<lambda>P. P after\<checkmark> e)\<close>
+    by (rule non_too_destructive_Aftertick)
+      (simp add: \<open>non_too_destructive \<Psi>\<close>,
+        metis non_too_destructive_fun_iff \<open>non_too_destructive \<Omega>\<close>)
+  hence * : \<open>restriction_shift (\<lambda>P. P after\<checkmark> e) (- 1)\<close>
+    unfolding non_too_destructive_def
+      non_too_destructive_on_def restriction_shift_def .
+  from restriction_shift_comp_restriction_shift[OF Cons.hyps "*"]
+  show ?case by simp
+qed
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/CSP_Restriction_Adm.thy
--- /dev/null
+++ thys/HOL-CSP_RS/CSP_Restriction_Adm.thy
@@ -0,0 +1,240 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Admissibility\<close>
+
+(*<*)
+theory CSP_Restriction_Adm
+  imports Process_Restriction_Space After_Operator_Non_Too_Destructive
+    "HOL-CSP_OpSem.Operational_Semantics_Laws"
+begin 
+  (*>*)
+
+
+named_theorems restriction_adm_processptick_simpset
+
+subsection \<open>Belonging\<close>
+
+lemma restriction_adm_in_D [restriction_adm_processptick_simpset] :
+  \<open>adm\<down> (\<lambda>x. t \<in> \<D> (f x))\<close>
+  and restriction_adm_notin_D [restriction_adm_processptick_simpset] :
+  \<open>adm\<down> (\<lambda>x. t \<notin> \<D> (f x))\<close>
+  and restriction_adm_in_F [restriction_adm_processptick_simpset] :
+  \<open>adm\<down> (\<lambda>x. (t, X) \<in> \<F> (f x))\<close>
+  and restriction_adm_notin_F [restriction_adm_processptick_simpset] :
+  \<open>adm\<down> (\<lambda>x. (t, X) \<notin> \<F> (f x))\<close> if \<open>cont\<down> f\<close>
+for f :: \<open>'b :: restriction \<Rightarrow> ('a, 'r) processptick\<close>
+proof (all \<open>rule restriction_adm_subst[OF \<open>cont\<down> f\<close>]\<close>)
+  have * : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>n \<ge> n0. \<sigma> n \<down> Suc (length t) = \<Sigma> \<down> Suc (length t)\<close>
+    for \<sigma> and \<Sigma> :: \<open>('a, 'r) processptick\<close>
+    by (metis restriction_tendstoD)
+
+  show \<open>adm\<down> (\<lambda>x. t \<in> \<D> x)\<close> \<open>adm\<down> (\<lambda>x. t \<notin> \<D> x)\<close>
+    by (rule restriction_admI,
+        metis (no_types) "*" D_restriction_processptick_Suc_length_iff_D dual_order.refl)+
+
+  show \<open>adm\<down> (\<lambda>x. (t, X) \<in> \<F> x)\<close> \<open>adm\<down> (\<lambda>x. (t, X) \<notin> \<F> x)\<close>
+    by (rule restriction_admI,
+        metis (no_types) "*" F_restriction_processptick_Suc_length_iff_F dual_order.refl)+
+qed
+
+
+corollary restriction_adm_in_T [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> adm\<down> (\<lambda>x. t \<in> \<T> (f x))\<close>
+  and restriction_adm_notin_T [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> adm\<down> (\<lambda>x. t \<notin> \<T> (f x))\<close>
+  by (fact restriction_adm_in_F[of f t \<open>{}\<close>, simplified T_F_spec])
+    (fact restriction_adm_notin_F[of f t \<open>{}\<close>, simplified T_F_spec])
+
+
+corollary restriction_adm_in_initials [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> adm\<down> (\<lambda>x. e \<in> (f x)0)\<close>
+  and restriction_adm_notin_initials [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> adm\<down> (\<lambda>x. e \<notin> (f x)0)\<close>
+  by (simp_all add: initials_def restriction_adm_in_T restriction_adm_notin_T)
+
+
+
+subsection \<open>Refining\<close>
+
+corollary restriction_adm_leF [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<sqsubseteq>F g x)\<close>
+  by (simp add: failure_refine_def subset_iff restriction_adm_processptick_simpset)
+
+corollary restriction_adm_leD [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<sqsubseteq>D g x)\<close> 
+  by (simp add: divergence_refine_def subset_iff restriction_adm_processptick_simpset)
+
+corollary restriction_adm_leT [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<sqsubseteq>T g x)\<close>
+  by (simp add: trace_refine_def subset_iff restriction_adm_processptick_simpset)
+
+corollary restriction_adm_leFD [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<sqsubseteq>FD g x)\<close>
+  by (simp add: failure_divergence_refine_def restriction_adm_processptick_simpset)
+
+corollary restriction_adm_leDT [restriction_adm_processptick_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<sqsubseteq>DT g x)\<close> 
+  by (simp add: trace_divergence_refine_def restriction_adm_processptick_simpset)
+
+
+
+
+subsubsection \<open>Transitions\<close>
+
+lemma (in After) restriction_cont_After [restriction_adm_simpset] :
+  \<open>cont\<down> (\<lambda>x. f x after a)\<close> if \<open>cont\<down> f\<close> and \<open>cont\<down> \<Psi>\<close>
+  \<comment> \<open>We could imagine more precise assumptions, but is it useful?\<close>
+proof (rule restriction_cont_comp[OF _ \<open>cont\<down> f\<close>])
+  show \<open>cont\<down> (\<lambda>P. P after a)\<close>
+  proof (rule restriction_contI)
+    show \<open>(\<lambda>n. \<sigma> n after a) \<midarrow>\<down>\<rightarrow> \<Sigma> after a\<close> if \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> for \<sigma> \<Sigma>
+    proof (rule restriction_tendstoI)
+      fix n
+      from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> obtain n0
+        where * : \<open>\<forall>k\<ge>n0. \<Sigma> \<down> Suc n = \<sigma> k \<down> Suc n\<close>
+        by (blast dest: restriction_tendstoD)
+      consider \<open>ev a \<in> \<Sigma>0\<close> \<open>\<forall>n\<ge>Suc n0. ev a \<in> (\<sigma> n)0\<close>
+        | \<open>ev a \<notin> \<Sigma>0\<close> \<open>\<forall>n\<ge>Suc n0. ev a \<notin> (\<sigma> n)0\<close>
+        by (metis "*" Suc_leD initials_restriction_processptick nat.distinct(1))
+      thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> after a \<down> n = \<sigma> k after a \<down> n\<close>
+      proof cases
+        assume \<open>ev a \<in> \<Sigma>0\<close> \<open>\<forall>n\<ge>Suc n0. ev a \<in> (\<sigma> n)0\<close>
+        hence \<open>\<forall>k\<ge>Suc n0. \<Sigma> after a \<down> n = \<sigma> k after a \<down> n\<close>
+          by (metis "*" Suc_leD restriction_processptick_After)
+        thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> after a \<down> n = \<sigma> k after a \<down> n\<close> ..
+      next
+        assume \<open>ev a \<notin> \<Sigma>0\<close> \<open>\<forall>n\<ge>Suc n0. ev a \<notin> (\<sigma> n)0\<close>
+        hence \<open>\<Sigma> after a = \<Psi> \<Sigma> a\<close> \<open>\<forall>k\<ge>Suc n0. \<sigma> k after a = \<Psi> (\<sigma> k) a\<close>
+          by (simp_all add: not_initial_After)
+        moreover from \<open>cont\<down> \<Psi>\<close>[THEN restriction_contD]
+        obtain n1 where \<open>\<forall>k\<ge>n1. \<Psi> \<Sigma> \<down> n = \<Psi> (\<sigma> k) \<down> n\<close>
+          by (blast intro: \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> dest: restriction_tendstoD)
+        ultimately have \<open>\<forall>k\<ge>max n1 (Suc n0). \<Sigma> after a \<down> n = \<sigma> k after a \<down> n\<close>
+          by simp (metis restriction_fun_def)
+        thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> after a \<down> n = \<sigma> k after a \<down> n\<close> ..
+      qed
+    qed
+  qed
+qed
+
+lemma (in AfterExt) restriction_cont_Aftertick [restriction_adm_simpset] :
+  \<open>cont\<down> (\<lambda>x. f x after\<checkmark> e)\<close> if \<open>cont\<down> f\<close> and \<open>cont\<down> \<Psi>\<close> and \<open>cont\<down> \<Omega>\<close>
+  \<comment> \<open>We could imagine more precise assumptions, but is it useful?\<close>
+proof (cases e)
+  show \<open>e = ev a \<Longrightarrow> cont\<down> (\<lambda>x. f x after\<checkmark> e)\<close> for a
+    by (simp add: Aftertick_def restriction_cont_After \<open>cont\<down> f\<close> \<open>cont\<down> \<Psi>\<close>)
+next
+  fix r assume \<open>e = \<checkmark>(r)\<close>
+  hence \<open>(\<lambda>x. f x after\<checkmark> e) = (\<lambda>x. \<Omega> (f x) r)\<close> by (simp add: Aftertick_def)
+  thus \<open>cont\<down> (\<lambda>x. f x after\<checkmark> e)\<close>
+    by (metis restriction_cont_comp restriction_cont_fun_imp that(1,3))
+qed
+
+lemma (in AfterExt) restriction_cont_Aftertrace [restriction_adm_simpset] :
+  \<open>cont\<down> (\<lambda>x. f x after\<T> t)\<close> if \<open>cont\<down> f\<close> and \<open>cont\<down> \<Psi>\<close> and \<open>cont\<down> \<Omega>\<close>
+  \<comment> \<open>We could imagine more precise assumptions, but is it useful?\<close>
+proof (rule restriction_cont_comp[OF _ \<open>cont\<down> f\<close>])
+  show \<open>cont\<down> (\<lambda>P. P after\<T> t)\<close>
+  proof (induct t)
+    show \<open>cont\<down> (\<lambda>P. P after\<T> [])\<close> by simp
+  next
+    fix e t assume \<open>cont\<down> (\<lambda>P. P after\<T> t)\<close>
+    show \<open>cont\<down> (\<lambda>P. P after\<T> (e # t))\<close>
+      by (simp, rule restriction_cont_comp[OF \<open>cont\<down> (\<lambda>P. P after\<T> t)\<close>])
+        (simp add: restriction_cont_Aftertick \<open>cont\<down> \<Psi>\<close> \<open>cont\<down> \<Omega>\<close>)
+  qed
+qed
+
+
+lemma (in OpSemTransitions) restriction_adm_weak_ev_trans [restriction_adm_processptick_simpset]:
+  \<comment> \<open>Could be weakened to a continuity assumption on term\<open>\<Psi>\<close>.\<close>
+  fixes f g :: \<open>'b :: restriction \<Rightarrow> ('a, 'r) processptick\<close>
+  assumes \<tau>_trans_restriction_adm:
+    \<open>\<And>f g :: 'b \<Rightarrow> ('a, 'r) processptick. cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<leadsto>\<tau> g x)\<close>
+    and \<open>cont\<down> f\<close> and \<open>cont\<down> g\<close> and \<open>cont\<down> \<Psi>\<close> and \<open>cont\<down> \<Omega>\<close>
+  shows \<open>adm\<down> (\<lambda>x. f x \<leadsto>e g x)\<close>
+proof (intro restriction_adm_conj)
+  show \<open>adm\<down> (\<lambda>x. ev e \<in> (f x)0)\<close>
+    by (simp add: \<open>cont\<down> f\<close> restriction_adm_in_initials)
+next
+  show \<open>adm\<down> (\<lambda>x. f x after\<checkmark> ev e \<leadsto>\<tau> g x)\<close>
+  proof (rule \<tau>_trans_restriction_adm[OF _ \<open>cont\<down> g\<close>],
+      rule restriction_cont_comp[OF _ \<open>cont\<down> f\<close>])
+    show \<open>cont\<down> (\<lambda>x. x after\<checkmark> ev e)\<close>
+      by (simp add: \<open>cont\<down> \<Psi>\<close> \<open>cont\<down> \<Omega>\<close> restriction_cont_Aftertick)
+  qed
+qed
+
+lemma (in OpSemTransitions) restriction_adm_weak_tick_trans [restriction_adm_processptick_simpset]:
+  fixes f g :: \<open>'b :: restriction \<Rightarrow> ('a, 'r) processptick\<close>
+  assumes \<tau>_trans_restriction_adm:
+    \<open>\<And>f g :: 'b \<Rightarrow> ('a, 'r) processptick. cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<leadsto>\<tau> g x)\<close>
+    and \<open>cont\<down> f\<close> and \<open>cont\<down> g\<close> and \<open>cont\<down> \<Psi>\<close> and \<open>cont\<down> \<Omega>\<close>
+  shows \<open>adm\<down> (\<lambda>x. f x \<leadsto>\<checkmark>r (g x))\<close>
+proof (intro restriction_adm_conj)
+  show \<open>adm\<down> (\<lambda>x. \<checkmark>(r) \<in> (f x)0)\<close>
+    by (simp add: \<open>cont\<down> f\<close> restriction_adm_in_initials)
+next
+  show \<open>adm\<down> (\<lambda>x. f x after\<checkmark> \<checkmark>(r) \<leadsto>\<tau> g x)\<close>
+  proof (rule \<tau>_trans_restriction_adm[OF _ \<open>cont\<down> g\<close>],
+      rule restriction_cont_comp[OF _ \<open>cont\<down> f\<close>])
+    show \<open>cont\<down> (\<lambda>x. x after\<checkmark> \<checkmark>(r))\<close>
+      by (simp add: \<open>cont\<down> \<Psi>\<close> \<open>cont\<down> \<Omega>\<close> restriction_cont_Aftertick)
+  qed
+qed
+
+lemma (in OpSemTransitions) restriction_adm_weak_trace_trans [restriction_adm_processptick_simpset]:
+  fixes f g :: \<open>'b :: restriction \<Rightarrow> ('a, 'r) processptick\<close>
+  assumes \<tau>_trans_restriction_adm:
+    \<open>\<And>f g :: 'b \<Rightarrow> ('a, 'r) processptick. cont\<down> f \<Longrightarrow> cont\<down> g \<Longrightarrow> adm\<down> (\<lambda>x. f x \<leadsto>\<tau> g x)\<close>
+    and \<open>cont\<down> f\<close> and \<open>cont\<down> g\<close> and \<open>cont\<down> \<Psi>\<close> and \<open>cont\<down> \<Omega>\<close>
+  shows \<open>adm\<down> (\<lambda>x. f x \<leadsto>* t (g x))\<close>
+proof (subst trace_trans_iff_T_and_Aftertrace_\<tau>_trans, intro restriction_adm_conj)
+  show \<open>adm\<down> (\<lambda>x. t \<in> \<T> (f x))\<close> by (simp add: \<open>cont\<down> f\<close> restriction_adm_in_T)
+next
+  show \<open>adm\<down> (\<lambda>x. f x after\<T> t \<leadsto>\<tau> g x)\<close>
+  proof (rule \<tau>_trans_restriction_adm[OF _ \<open>cont\<down> g\<close>])
+    show \<open>cont\<down> (\<lambda>x. f x after\<T> t)\<close>
+      by (simp add: \<open>cont\<down> f\<close> \<open>cont\<down> \<Psi>\<close> \<open>cont\<down> \<Omega>\<close> restriction_cont_Aftertrace)
+  qed
+qed
+
+
+
+declare restriction_adm_processptick_simpset [simp]
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Choices_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Choices_Non_Destructive.thy
@@ -0,0 +1,125 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Non Destructiveness of Choices\<close>
+
+(*<*)
+theory Choices_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSPM.CSPM_Laws"
+begin
+  (*>*)
+
+
+subsection \<open>Equality\<close>
+
+lemma restriction_processptick_Ndet : \<open>P \<sqinter> Q \<down> n = (P \<down> n) \<sqinter> (Q \<down> n)\<close>
+  by (auto simp add: Process_eq_spec Ndet_projs restriction_processptick_projs)
+
+lemma restriction_processptick_GlobalNdet :
+  \<open>(\<sqinter>a \<in> A. P a) \<down> n = (if n = 0 then \<bottom> else \<sqinter>a \<in> A. (P a \<down> n))\<close>
+  by simp (auto simp add: Process_eq_spec GlobalNdet_projs restriction_processptick_projs)
+
+
+lemma restriction_processptick_GlobalDet :
+  \<open>(\<box>a \<in> A. P a) \<down> n = (if n = 0 then \<bottom> else \<box>a \<in> A. (P a \<down> n))\<close>
+  (is \<open>?lhs = (if n = 0 then \<bottom> else ?rhs)\<close>)
+proof (split if_split, intro conjI impI)
+  show \<open>n = 0 \<Longrightarrow> ?lhs = \<bottom>\<close> by simp
+next
+  show \<open>?lhs = ?rhs\<close> if \<open>n \<noteq> 0\<close>
+  proof (rule Process_eq_optimized_bisI)
+    show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+      by (auto simp add: GlobalDet_projs D_restriction_processptick \<open>n \<noteq> 0\<close> split: if_split_asm)
+  next
+    show \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+      by (auto simp add: GlobalDet_projs D_restriction_processptick)
+  next
+    show \<open>([], X) \<in> \<F> ?lhs \<Longrightarrow> ([], X) \<in> \<F> ?rhs\<close> for X
+      by (auto simp add: restriction_processptick_projs F_GlobalDet)
+  next
+    show \<open>([], X) \<in> \<F> ?rhs \<Longrightarrow> ([], X) \<in> \<F> ?lhs\<close> for X
+      by (auto simp add: restriction_processptick_projs F_GlobalDet)
+        (metis append_eq_Cons_conv eventptick.disc(2) tickFree_Cons_iff)
+  next
+    show \<open>(e # t, X) \<in> \<F> ?lhs \<Longrightarrow> (e # t, X) \<in> \<F> ?rhs\<close> for e t X
+      by (auto simp add: \<open>n \<noteq> 0\<close> F_restriction_processptick GlobalDet_projs split: if_split_asm)
+  next
+    show \<open>(e # t, X) \<in> \<F> ?rhs \<Longrightarrow> (e # t, X) \<in> \<F> ?lhs\<close> for e t X
+      by (auto simp add: F_restriction_processptick GlobalDet_projs)
+  qed
+qed
+
+
+lemma restriction_processptick_Det: \<open>P \<box> Q \<down> n = (P \<down> n) \<box> (Q \<down> n)\<close> (is \<open>?lhs = ?rhs\<close>)
+proof -
+  have \<open>P \<box> Q \<down> n = \<box>i \<in> {0, 1 :: nat}. (if i = 0 then P else Q) \<down> n\<close>
+    by (simp add: GlobalDet_distrib_unit_bis)
+  also have \<open>\<dots> = (if n = 0 then \<bottom> else \<box>i \<in> {0, 1 :: nat}. (if i = 0 then P \<down> n else Q \<down> n))\<close>
+    by (simp add: restriction_processptick_GlobalDet if_distrib if_distribR)
+  also have \<open>\<dots> = (P \<down> n) \<box> (Q \<down> n)\<close> by (simp add: GlobalDet_distrib_unit_bis)
+  finally show \<open>P \<box> Q \<down> n = (P \<down> n) \<box> (Q \<down> n)\<close> .
+qed
+
+
+corollary restriction_processptick_Sliding: \<open>P \<rhd> Q \<down> n = (P \<down> n) \<rhd> (Q \<down> n)\<close>
+  by (simp add: restriction_processptick_Det restriction_processptick_Ndet Sliding_def)
+
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma GlobalNdet_non_destructive : \<open>non_destructive (\<lambda>P. \<sqinter>a \<in> A. P a)\<close>
+  by (auto intro: non_destructiveI
+      simp add: restriction_processptick_GlobalNdet restriction_fun_def)
+
+lemma Ndet_non_destructive : \<open>non_destructive (\<lambda>(P, Q). P \<sqinter> Q)\<close>
+  by (auto intro: non_destructiveI
+      simp add: restriction_processptick_Ndet restriction_prod_def)
+
+lemma GlobalDet_non_destructive : \<open>non_destructive (\<lambda>P. \<box>a \<in> A. P a)\<close>
+  by (auto intro: non_destructiveI
+      simp add: restriction_processptick_GlobalDet restriction_fun_def)
+
+lemma Det_non_destructive : \<open>non_destructive (\<lambda>(P, Q). P \<box> Q)\<close>
+  by (auto intro: non_destructiveI
+      simp add: restriction_processptick_Det restriction_prod_def)
+
+corollary Sliding_non_destructive : \<open>non_destructive (\<lambda>(P, Q). P \<rhd> Q)\<close>
+  by (auto intro: non_destructiveI
+      simp add: restriction_processptick_Sliding restriction_prod_def)
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/HOL-CSP_RS.thy
--- /dev/null
+++ thys/HOL-CSP_RS/HOL-CSP_RS.thy
@@ -0,0 +1,285 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Higher-Order Rules\<close>
+
+(*<*)
+theory "HOL-CSP_RS"
+  imports
+    Prefixes_Constructive
+    Choices_Non_Destructive
+    Renaming_Non_Destructive
+    Sequential_Composition_Non_Destructive
+    Synchronization_Product_Non_Destructive
+    Throw_Non_Destructive
+    Interrupt_Non_Destructive
+    Hiding_Destructive
+    CSP_Restriction_Adm
+begin
+  (*>*)
+
+text \<open>This is the main entry point.
+      We configure the simplifier below.\<close>
+
+
+named_theorems restriction_shift_processptick_simpset
+
+subsection \<open>Prefixes\<close>
+
+lemma Mprefix_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>constructive (\<lambda>x. \<box>a \<in> A \<rightarrow> f a x)\<close> if \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a))\<close>
+proof -
+  have * : \<open>\<box>a \<in> A \<rightarrow> f a x = \<box>a \<in> A \<rightarrow> (if a \<in> A then f a x else STOP)\<close> for x
+    by (auto intro: mono_Mprefix_eq)
+  show \<open>constructive (\<lambda>x. \<box>a \<in> A \<rightarrow> f a x)\<close>
+    by (subst "*", rule constructive_comp_non_destructive
+        [OF Mprefix_constructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>])
+      (auto intro: that)
+qed
+
+lemma Mndetprefix_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>constructive (\<lambda>x. \<sqinter>a \<in> A \<rightarrow> f a x)\<close> if \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a))\<close>
+proof -
+  have * : \<open>\<sqinter>a \<in> A \<rightarrow> f a x = \<sqinter>a \<in> A \<rightarrow> (if a \<in> A then f a x else STOP)\<close> for x
+    by (auto intro: mono_Mndetprefix_eq)
+  show \<open>constructive (\<lambda>x. \<sqinter>a \<in> A \<rightarrow> f a x)\<close>
+    by (subst "*", rule constructive_comp_non_destructive
+        [OF Mndetprefix_constructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>])
+      (auto intro: that)
+qed
+
+corollary write0_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> constructive (\<lambda>x. a \<rightarrow> f x)\<close>
+  by (simp add: write0_def Mprefix_restriction_shift_processptick)
+
+corollary write_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> constructive (\<lambda>x. c!a \<rightarrow> f x)\<close>
+  by (simp add: write_def Mprefix_restriction_shift_processptick)
+
+corollary read_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> constructive (\<lambda>x. c?a \<in> A \<rightarrow> f a x)\<close>
+  by (simp add: read_def Mprefix_restriction_shift_processptick inv_into_into)
+
+corollary ndet_write_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> constructive (\<lambda>x. c!!a \<in> A \<rightarrow> f a x)\<close>
+  by (simp add: ndet_write_def Mndetprefix_restriction_shift_processptick inv_into_into)
+
+
+subsection \<open>Choices\<close>
+
+lemma GlobalNdet_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> non_destructive (\<lambda>x. \<sqinter>a \<in> A. f a x)\<close>
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> constructive (f a)) \<Longrightarrow> constructive (\<lambda>x. \<sqinter>a \<in> A. f a x)\<close>
+proof -
+  have * : \<open>\<sqinter>a \<in> A. f a x = \<sqinter>a \<in> A. (if a \<in> A then f a x else STOP)\<close> for x
+    by (auto intro: mono_GlobalNdet_eq)
+
+  show \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> non_destructive (\<lambda>x. \<sqinter>a \<in> A. f a x)\<close>
+    by (subst "*", rule non_destructive_comp_non_destructive
+        [OF GlobalNdet_non_destructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>]) auto
+
+  show \<open>(\<And>a. a \<in> A \<Longrightarrow> constructive (f a)) \<Longrightarrow> constructive (\<lambda>x. \<sqinter>a \<in> A. f a x)\<close>
+    by (subst "*", rule non_destructive_comp_constructive
+        [OF GlobalNdet_non_destructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>]) auto
+qed
+
+lemma GlobalDet_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> non_destructive (\<lambda>x. \<box>a \<in> A. f a x)\<close>
+  \<open>(\<And>a. a \<in> A \<Longrightarrow> constructive (f a)) \<Longrightarrow> constructive (\<lambda>x. \<box>a \<in> A. f a x)\<close>
+proof -
+  have * : \<open>\<box>a \<in> A. f a x = \<box>a \<in> A. (if a \<in> A then f a x else STOP)\<close> for x
+    by (auto intro: mono_GlobalDet_eq)
+
+  show \<open>(\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)) \<Longrightarrow> non_destructive (\<lambda>x. \<box>a \<in> A. f a x)\<close>
+    by (subst "*", rule non_destructive_comp_non_destructive
+        [OF GlobalDet_non_destructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>]) auto
+
+  show \<open>(\<And>a. a \<in> A \<Longrightarrow> constructive (f a)) \<Longrightarrow> constructive (\<lambda>x. \<box>a \<in> A. f a x)\<close>
+    by (subst "*", rule non_destructive_comp_constructive
+        [OF GlobalDet_non_destructive, of \<open>\<lambda>x a. if a \<in> A then f a x else STOP\<close>]) auto
+qed
+
+
+lemma Ndet_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x \<sqinter> g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x \<sqinter> g x)\<close>
+  by (auto intro!: non_destructiveI constructiveI
+      simp add: restriction_processptick_Ndet dest!: non_destructiveD constructiveD)
+
+lemma Det_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x \<box> g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x \<box> g x)\<close>
+  by (auto intro!: non_destructiveI constructiveI
+      simp add: restriction_processptick_Det dest!: non_destructiveD constructiveD)
+
+lemma Sliding_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x \<rhd> g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x \<rhd> g x)\<close>
+  by (auto intro!: non_destructiveI constructiveI
+      simp add: restriction_processptick_Sliding dest!: non_destructiveD constructiveD)
+
+
+subsection \<open>Renaming\<close>
+
+lemma Renaming_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive P \<Longrightarrow> non_destructive (\<lambda>x. Renaming (P x) f g)\<close>
+  \<open>constructive P \<Longrightarrow> constructive (\<lambda>x. Renaming (P x) f g)\<close>
+  by (auto intro!: non_destructiveI constructiveI
+      simp add: restriction_processptick_Renaming dest!: non_destructiveD constructiveD)
+
+
+subsection \<open>Sequential Composition\<close>
+
+lemma Seq_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x ; g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x ; g x)\<close>
+  by (fact non_destructive_comp_non_destructive[OF Seq_non_destructive non_destructive_prod_codomain, simplified])
+    (fact non_destructive_comp_constructive[OF Seq_non_destructive constructive_prod_codomain, simplified])
+
+
+lemma MultiSeq_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>l. l \<in> set L \<Longrightarrow> non_destructive (f l)) \<Longrightarrow> non_destructive (\<lambda>x. SEQ l \<in>@ L. f l x)\<close>
+  \<open>(\<And>l. l \<in> set L \<Longrightarrow> constructive (f l)) \<Longrightarrow> constructive (\<lambda>x. SEQ l \<in>@ L. f l x)\<close>
+  by (induct L rule: rev_induct; simp add: Seq_restriction_shift_processptick)+
+
+
+corollary MultiSeq_non_destructive : \<open>non_destructive (\<lambda>P. SEQ l \<in>@ L. P l)\<close>
+  by (simp add: MultiSeq_restriction_shift_processptick(1)[of L \<open>\<lambda>l x. x l\<close>])
+
+
+
+subsection \<open>Synchronization Product\<close>
+
+lemma Sync_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x \<lbrakk>S\<rbrakk> g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x \<lbrakk>S\<rbrakk> g x)\<close>
+  by (fact non_destructive_comp_non_destructive[OF Sync_non_destructive non_destructive_prod_codomain, simplified])
+    (fact non_destructive_comp_constructive[OF Sync_non_destructive constructive_prod_codomain, simplified])
+
+
+lemma MultiSync_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>(\<And>m. m \<in># M \<Longrightarrow> non_destructive (f m)) \<Longrightarrow> non_destructive (\<lambda>x. \<lbrakk>S\<rbrakk> m \<in># M. f m x)\<close>
+  \<open>(\<And>m. m \<in># M \<Longrightarrow> constructive (f m)) \<Longrightarrow> constructive (\<lambda>x. \<lbrakk>S\<rbrakk> m \<in># M. f m x)\<close>
+  by (induct M rule: induct_subset_mset_empty_single;
+      simp add: Sync_restriction_shift_processptick)+
+
+
+corollary MultiSync_non_destructive : \<open>non_destructive (\<lambda>P. \<lbrakk>S\<rbrakk> m \<in># M. P m)\<close>
+  by (rule MultiSync_restriction_shift_processptick(1)[of M \<open>\<lambda>m x. x m\<close>]) simp
+
+
+
+subsection \<open>Throw\<close>
+
+lemma Throw_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> non_destructive (g a)) \<Longrightarrow> non_destructive (\<lambda>x. f x \<Theta> a \<in> A. g a x)\<close>
+  \<open>constructive f \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> constructive (g a)) \<Longrightarrow> constructive (\<lambda>x. f x \<Theta> a \<in> A. g a x)\<close>
+proof -
+  have * : \<open>f x \<Theta> a \<in> A. g a x = f x \<Theta> a \<in> A. (if a \<in> A then g a x else STOP)\<close> for x
+    by (auto intro: mono_Throw_eq)
+
+  show \<open>non_destructive f \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> non_destructive (g a)) \<Longrightarrow> non_destructive (\<lambda>x. f x \<Theta> a \<in> A. g a x)\<close>
+    by (subst "*", erule non_destructive_comp_non_destructive
+        [OF Throw_non_destructive non_destructive_prod_codomain, simplified]) auto
+
+  show \<open>constructive f \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> constructive (g a)) \<Longrightarrow> constructive (\<lambda>x. f x \<Theta> a \<in> A. g a x)\<close>
+    by (subst "*", erule non_destructive_comp_constructive
+        [OF Throw_non_destructive constructive_prod_codomain, simplified]) auto
+qed
+
+
+
+subsection \<open>Interrupt\<close>
+
+lemma Interrupt_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive g \<Longrightarrow> non_destructive (\<lambda>x. f x \<triangle> g x)\<close>
+  \<open>constructive f \<Longrightarrow> constructive g \<Longrightarrow> constructive (\<lambda>x. f x \<triangle> g x)\<close>
+  by (fact non_destructive_comp_non_destructive[OF Interrupt_non_destructive non_destructive_prod_codomain, simplified])
+    (fact non_destructive_comp_constructive[OF Interrupt_non_destructive constructive_prod_codomain, simplified])
+
+
+
+subsection \<open>After\<close>
+
+lemma (in After) After_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>non_too_destructive \<Psi> \<Longrightarrow> constructive f \<Longrightarrow> non_destructive (\<lambda>x. f x  after a)\<close>
+  \<open>non_too_destructive \<Psi> \<Longrightarrow> non_destructive f \<Longrightarrow> non_too_destructive (\<lambda>x. f x  after a)\<close>
+  by (auto intro!: non_too_destructive_comp_constructive[OF non_too_destructive_After]
+      non_too_destructive_comp_non_destructive[OF non_too_destructive_After])
+
+lemma (in AfterExt) Aftertick_restriction_shift_processptick [restriction_shift_processptick_simpset] :
+  \<open>\<lbrakk>non_too_destructive \<Psi>; non_too_destructive \<Omega>; constructive f\<rbrakk>
+   \<Longrightarrow> non_destructive (\<lambda>x. f x  after\<checkmark> e)\<close>
+  \<open>\<lbrakk>non_too_destructive \<Psi>; non_too_destructive \<Omega>; non_destructive f\<rbrakk>
+   \<Longrightarrow> non_too_destructive (\<lambda>x. f x  after\<checkmark> e)\<close>
+  by (auto intro!: non_too_destructive_comp_constructive[OF non_too_destructive_Aftertick]
+      non_too_destructive_comp_non_destructive[OF non_too_destructive_Aftertick]
+      simp add: non_too_destructive_fun_iff)
+
+
+
+
+subsection \<open>Illustration\<close>
+
+declare restriction_shift_processptick_simpset [simp]
+
+
+notepad begin
+  fix e f g :: 'a fix r s :: 'r
+  fix A B C :: \<open>'a set\<close>
+  fix S :: \<open>'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a set\<close>
+  define P where \<open>P \<equiv> \<upsilon> X. ((\<box>a \<in> A \<rightarrow> X \<sqinter> SKIP r) \<triangle> (f \<rightarrow> STOP))
+                            \<box> (g \<rightarrow> X)
+                            \<box> ((f \<rightarrow> e \<rightarrow> (\<bottom> \<sqinter> (e \<rightarrow> X))) \<Theta> b \<in> insert e B. (e \<rightarrow> SKIP s))\<close>  
+    (is \<open>P \<equiv> \<upsilon> X. ?f X\<close>)
+  have \<open>constructive ?f\<close> by simp
+  have \<open>cont ?f\<close> by simp
+  have \<open>P = ?f P\<close>
+    unfolding P_def by (subst restriction_fix_eq) simp_all
+
+
+  define Q where \<open>Q \<equiv> \<upsilon> X. (\<lambda>\<sigma> \<sigma>' \<sigma>'' . e \<rightarrow> \<sqinter> b \<in> S \<sigma> \<sigma>' \<sigma>''. X b b b \<sqinter> SKIP r)\<close> (is \<open>Q \<equiv> \<upsilon> X. ?g X\<close>)
+  have \<open>constructive ?g\<close> by simp
+
+
+  define R where \<open>R \<equiv> \<upsilon> (x, y). (e \<rightarrow> y \<sqinter> SKIP r, \<box>a \<in> A \<rightarrow> x)\<close>
+    (is \<open>R \<equiv> \<upsilon> (x, y). (?h y, ?i x)\<close>)
+  have \<open>snd R = \<box>a \<in> A \<rightarrow> fst R\<close>
+    by (unfold R_def, subst restriction_fix_eq)
+      (simp_all add: case_prod_beta')
+
+end
+
+
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/HOL-CSP_RS.zip
Binary file thys/HOL-CSP_RS/HOL-CSP_RS.zip has changed
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Hiding_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Hiding_Destructive.thy
@@ -0,0 +1,120 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Destructiveness of Hiding\<close>
+
+(*>*)
+theory Hiding_Destructive
+  imports "HOL-CSPM.CSPM_Laws" Prefixes_Constructive
+begin
+  (*>*)
+
+
+subsection \<open>Refinement\<close>
+
+lemma Hiding_restriction_processptick_FD : \<open>(P \<down> n) \ S \<sqsubseteq>FD P \ S \<down> n\<close>
+proof (unfold refine_defs, safe)
+  show * : \<open>t \<in> \<D> (P \ S \<down> n) \<Longrightarrow> t \<in> \<D> ((P \<down> n) \ S)\<close> for t
+  proof (elim D_restriction_processptickE)
+    show \<open>t \<in> \<D> (P \ S) \<Longrightarrow> t \<in> \<D> ((P \<down> n) \ S)\<close>
+      by (simp add: D_Hiding restriction_processptick_projs) blast
+  next
+    fix u v
+    assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (P \ S)\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    from \<open>u \<in> \<T> (P \ S)\<close>[simplified T_Hiding, simplified]
+    consider \<open>u \<in> \<D> (P \ S)\<close> | u' where \<open>u = trace_hide u' (ev ` S)\<close> \<open>(u', ev ` S) \<in> \<F> P\<close>
+      by (simp add: F_Hiding D_Hiding) blast
+    thus \<open>t \<in> \<D> ((P \<down> n) \ S)\<close>
+    proof cases
+      assume \<open>u \<in> \<D> (P \ S)\<close>
+      hence \<open>t \<in> \<D> (P \ S)\<close> by (simp add: \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> is_processT7)
+      with restriction_processptick_approx_self le_approx1 mono_Hiding
+      show \<open>t \<in> \<D> ((P \<down> n) \ S)\<close> by blast
+    next
+      fix u' assume \<open>u = trace_hide u' (ev ` S)\<close> \<open>(u', ev ` S) \<in> \<F> P\<close>
+      with \<open>length u = n\<close> \<open>tF u\<close> Hiding_tickFree length_filter_le F_T
+      have \<open>n \<le> length u'\<close> \<open>tickFree u'\<close> \<open>u' \<in> \<T> P\<close> by blast+
+      with \<open>u = trace_hide u' (ev ` S)\<close>
+      have \<open>u' = (take n u') @ (drop n u') \<and> take n u' \<in> \<T> P \<and>
+            length (take n u') = n \<and> tF (take n u') \<and> ftF (drop n u')\<close>
+        by (simp add: min_def) (metis append_take_drop_id is_processT3_TR_append
+            tickFree_append_iff tickFree_imp_front_tickFree)
+      with D_restriction_processptick have \<open>u' \<in> \<D> (P \<down> n)\<close> by blast
+      with Hiding_tickFree \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>u = trace_hide u' (ev ` S)\<close> \<open>tF u\<close>
+      show \<open>t \<in> \<D> ((P \<down> n) \ S)\<close> by (simp add: D_Hiding) blast
+    qed
+  qed
+
+  fix s X
+  assume \<open>(s, X) \<in> \<F> ((P \ S) \<down> n)\<close>
+  then consider \<open>s \<in> \<D> ((P \ S) \<down> n)\<close> | \<open>(s, X) \<in> \<F> (P \ S)\<close>
+    unfolding restriction_processptick_projs by blast
+  thus \<open>(s, X) \<in> \<F> ((P \<down> n) \ S)\<close>
+  proof cases
+    from "*" D_F show \<open>s \<in> \<D> ((P \ S) \<down> n) \<Longrightarrow> (s, X) \<in> \<F> ((P \<down> n) \ S)\<close> by blast
+  next
+    show \<open>(s, X) \<in> \<F> (P \ S) \<Longrightarrow> (s, X) \<in> \<F> ((P \<down> n) \ S)\<close>
+      using restriction_processptick_approx_self D_F mono_Hiding proc_ord2a by blast
+  qed
+qed
+
+
+subsection \<open>Destructiveness\<close>
+
+lemma Hiding_destructive :
+  \<open>\<exists>P Q :: ('a, 'r) processptick. P \<down> n = Q \<down> n \<and> (P \ S) \<down> Suc 0 \<noteq> (Q \ S) \<down> Suc 0\<close> if \<open>S \<noteq> {}\<close>
+proof -
+  from \<open>S \<noteq> {}\<close> obtain e where \<open>e \<in> S\<close> by blast
+  define P :: \<open>('a, 'r) processptick\<close> where \<open>P \<equiv> iterate (Suc n)\<cdot>(\<Lambda> X. write0 e X)\<cdot>(SKIP undefined)\<close>
+  define Q :: \<open>('a, 'r) processptick\<close> where \<open>Q \<equiv> iterate (Suc n)\<cdot>(\<Lambda> X. write0 e X)\<cdot>STOP\<close>
+  have \<open>P \<down> n = Q \<down> n\<close>
+    unfolding P_def Q_def by (induct n) (simp_all add: restriction_processptick_write0) 
+  have \<open>P \ S = SKIP undefined\<close>
+    unfolding P_def by (induct n) (simp_all add: Hiding_write0_non_disjoint \<open>e \<in> S\<close>)
+  hence \<open>(P \ S) \<down> Suc 0 = SKIP undefined\<close> by simp
+  have \<open>Q \ S = STOP\<close>
+    unfolding Q_def by (induct n) (simp_all add: Hiding_write0_non_disjoint \<open>e \<in> S\<close>)
+  hence \<open>(Q \ S) \<down> Suc 0 = STOP\<close> by simp
+
+  have \<open>P \<down> n = Q \<down> n \<and> (P \ S) \<down> Suc 0 \<noteq> (Q \ S) \<down> Suc 0\<close>
+    by (simp add: \<open>P \<down> n = Q \<down> n\<close> \<open>(P \ S) \<down> Suc 0 = SKIP undefined\<close>
+        \<open>(Q \ S) \<down> Suc 0 = STOP\<close> SKIP_Neq_STOP)
+  thus ?thesis by blast
+qed
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Interrupt_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Interrupt_Non_Destructive.thy
@@ -0,0 +1,344 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Non Destructiveness of Interrupt\<close>
+
+(*<*)
+theory Interrupt_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSPM.CSPM_Laws"
+begin
+  (*>*)
+
+subsection \<open>Refinement\<close>
+
+lemma restriction_processptick_Interrupt_FD :
+  \<open>P \<triangle> Q \<down> n \<sqsubseteq>FD (P \<down> n) \<triangle> (Q \<down> n)\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+proof (unfold refine_defs, safe)  
+  show * : \<open>t \<in> \<D> ?lhs\<close> if \<open>t \<in> \<D> ?rhs\<close> for t
+  proof -
+    from \<open>t \<in> \<D> ?rhs\<close> consider \<open>t \<in> \<D> (P \<down> n)\<close>
+      | u v where \<open>t = u @ v\<close> \<open>u \<in> \<T> (P \<down> n)\<close> \<open>tF u\<close> \<open>v \<in> \<D> (Q \<down> n)\<close>
+      by (simp add: D_Interrupt) blast
+    thus \<open>t \<in> \<D> ?lhs\<close>
+    proof cases
+      show \<open>t \<in> \<D> (P \<down> n) \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+        by (elim D_restriction_processptickE) (auto simp add: D_restriction_processptick Interrupt_projs)
+    next
+      fix u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (P \<down> n)\<close> \<open>tF u\<close> \<open>v \<in> \<D> (Q \<down> n)\<close>
+      from \<open>v \<in> \<D> (Q \<down> n)\<close> show \<open>t \<in> \<D> ?lhs\<close>
+      proof (elim D_restriction_processptickE)
+        from \<open>t = u @ v\<close> \<open>u \<in> \<T> (P \<down> n)\<close> \<open>tF u\<close> show \<open>v \<in> \<D> Q \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+          by (auto simp add: restriction_processptick_projs Interrupt_projs
+              D_imp_front_tickFree front_tickFree_append)
+      next
+        fix w x assume \<open>v = w @ x\<close> \<open>w \<in> \<T> Q\<close> \<open>length w = n\<close> \<open>tF w\<close> \<open>ftF x\<close>
+        from \<open>u \<in> \<T> (P \<down> n)\<close> consider \<open>u \<in> \<D> (P \<down> n)\<close> | \<open>u \<in> \<T> P\<close> \<open>length u \<le> n\<close>
+          by (elim T_restriction_processptickE) (auto simp add: D_restriction_processptick)
+        thus \<open>t \<in> \<D> ?lhs\<close>
+        proof cases
+          assume \<open>u \<in> \<D> (P \<down> n)\<close>
+          with D_imp_front_tickFree \<open>t = u @ v\<close> \<open>tF u\<close> \<open>v \<in> \<D> (Q \<down> n)\<close> is_processT7
+          have \<open>t \<in> \<D> (P \<down> n)\<close> by blast
+          thus \<open>t \<in> \<D> ?lhs\<close> by (elim D_restriction_processptickE)
+              (auto simp add: D_restriction_processptick Interrupt_projs)
+        next
+          assume \<open>u \<in> \<T> P\<close> \<open>length u \<le> n\<close>
+          hence \<open>t = take n (u @ w) @ drop (n - length u) w @ x \<and>
+                 take n (u @ w) \<in> \<T> (P \<triangle> Q) \<and> length (take n (u @ w)) = n \<and>
+                 tF (take n (u @ w)) \<and> ftF (drop (n - length u) w @ x)\<close>
+            by (simp add: \<open>t = u @ v\<close> \<open>v = w @ x\<close> \<open>length w = n\<close> \<open>tF u\<close> T_Interrupt)
+              (metis \<open>ftF x\<close> \<open>tF u\<close> \<open>tF w\<close> \<open>w \<in> \<T> Q\<close> append_take_drop_id
+                front_tickFree_append is_processT3_TR_append tickFree_append_iff)
+          with D_restriction_processptick show \<open>t \<in> \<D> ?lhs\<close> by blast
+        qed
+      qed
+    qed
+  qed
+
+  show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+    by (meson "*" is_processT8 mono_Interrupt proc_ord2a restriction_processptick_approx_self)
+qed
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma Interrupt_non_destructive :
+  \<open>non_destructive (\<lambda>(P :: ('a, 'r) processptick, Q). P \<triangle> Q)\<close>
+proof (rule order_non_destructiveI, clarify)
+  fix P P' Q Q' :: \<open>('a, 'r) processptick\<close> and n
+  assume \<open>(P, Q) \<down> n = (P', Q') \<down> n\<close> \<open>0 < n\<close>
+  hence \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+    by (simp_all add: restriction_prod_def)
+
+  { let ?lhs = \<open>P \<triangle> Q \<down> n\<close>
+    fix t u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (P' \<triangle> Q')\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    from this(2) \<open>tF u\<close> obtain u1 u2
+      where \<open>u = u1 @ u2\<close> \<open>u1 \<in> \<T> P'\<close> \<open>tF u1\<close> \<open>u2 \<in> \<T> Q'\<close> \<open>tF u2\<close>
+      by (simp add: T_Interrupt)
+        (metis append_Nil2 is_processT1_TR tickFree_append_iff)
+
+    from \<open>length u = n\<close> \<open>u = u1 @ u2\<close>
+    have \<open>length u1 \<le> n\<close> \<open>length u2 \<le> n\<close> by simp_all
+    with \<open>u1 \<in> \<T> P'\<close> \<open>P \<down> n = P' \<down> n\<close> \<open>u2 \<in> \<T> Q'\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+    have \<open>u1 \<in> \<T> P\<close> \<open>u2 \<in> \<T> Q\<close>
+      by (metis T_restriction_processptickI
+          length_le_in_T_restriction_processptick)+
+    with \<open>tF u1\<close> have \<open>u \<in> \<T> (P \<triangle> Q)\<close>
+      by (auto simp add: \<open>u = u1 @ u2\<close> T_Interrupt)
+    with \<open>length u = n\<close> \<open>tF u\<close> have \<open>u \<in> \<D> ?lhs\<close>
+      by (simp add: D_restriction_processptickI)
+    hence \<open>t \<in> \<D> ?lhs\<close> by (simp add: \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> is_processT7)
+  } note * = this
+
+  show \<open>P \<triangle> Q \<down> n \<sqsubseteq>FD P' \<triangle> Q' \<down> n\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+  proof (unfold refine_defs, safe)
+    show div : \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    proof (elim D_restriction_processptickE)
+      assume \<open>t \<in> \<D> (P' \<triangle> Q')\<close> \<open>length t \<le> n\<close>
+      from this(1) consider \<open>t \<in> \<D> P'\<close>
+        | (divR) t1 t2 where \<open>t = t1 @ t2\<close> \<open>t1 \<in> \<T> P'\<close> \<open>tF t1\<close> \<open>t2 \<in> \<D> Q'\<close>
+        unfolding D_Interrupt by blast
+      thus \<open>t \<in> \<D> ?lhs\<close>
+      proof cases
+        assume \<open>t \<in> \<D> P'\<close>
+        hence \<open>ftF t\<close> by (simp add: D_imp_front_tickFree)
+        thus \<open>t \<in> \<D> ?lhs\<close>
+        proof (elim front_tickFreeE)
+          show \<open>t \<in> \<D> ?lhs\<close> if \<open>tF t\<close>
+          proof (cases \<open>length t = n\<close>)
+            assume \<open>length t = n\<close>
+            from \<open>P \<down> n = P' \<down> n\<close> \<open>t \<in> \<D> P'\<close> \<open>length t \<le> n\<close> have \<open>t \<in> \<T> P\<close>
+              by (metis D_T T_restriction_processptickI
+                  length_le_in_T_restriction_processptick)
+            with \<open>tF t\<close> \<open>length t = n\<close> front_tickFree_Nil show \<open>t \<in> \<D> ?lhs\<close>
+              by (simp (no_asm) add: D_restriction_processptick T_Interrupt) blast
+          next
+            assume \<open>length t \<noteq> n\<close>
+            with \<open>length t \<le> n\<close> have \<open>length t < n\<close> by simp
+            with \<open>P \<down> n = P' \<down> n\<close> \<open>t \<in> \<D> P'\<close> have \<open>t \<in> \<D> P\<close>
+              by (metis D_restriction_processptickI
+                  length_less_in_D_restriction_processptick)
+            thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI D_Interrupt)
+          qed
+        next
+          fix t' r assume \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>tF t'\<close>
+          with \<open>t \<in> \<D> P'\<close> \<open>length t \<le> n\<close>
+          have \<open>t' \<in> \<D> P'\<close> \<open>length t' < n\<close> by (auto intro: is_processT9)
+          with \<open>P \<down> n = P' \<down> n\<close> have \<open>t' \<in> \<D> P\<close>
+            by (metis D_restriction_processptickI
+                length_less_in_D_restriction_processptick)
+          with \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>tF t'\<close> have \<open>t \<in> \<D> P\<close> by (simp add: is_processT7)
+          thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI D_Interrupt)
+        qed
+      next
+        fix u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> P'\<close> \<open>tF u\<close> \<open>v \<in> \<D> Q'\<close>
+        from \<open>t = u @ v\<close> \<open>length t \<le> n\<close> have \<open>length u \<le> n\<close> by simp
+        with \<open>P \<down> n = P' \<down> n\<close> \<open>u \<in> \<T> P'\<close> have \<open>u \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+        show \<open>t \<in> \<D> ?lhs\<close>
+        proof (cases \<open>length v = n\<close>)
+          assume \<open>length v = n\<close>
+          with \<open>t = u @ v\<close> \<open>length t \<le> n\<close> have \<open>u = []\<close> by simp
+          from D_imp_front_tickFree \<open>v \<in> \<D> Q'\<close> have \<open>ftF v\<close> by blast
+          thus \<open>t \<in> \<D> ?lhs\<close>
+          proof (elim front_tickFreeE)
+            assume \<open>tF v\<close>
+            from \<open>length v = n\<close> \<open>Q \<down> n = Q' \<down> n\<close> \<open>v \<in> \<D> Q'\<close> have \<open>v \<in> \<T> Q\<close>
+              by (metis D_T D_restriction_processptickI dual_order.refl
+                  length_le_in_T_restriction_processptick)
+            with \<open>tF u\<close> \<open>u \<in> \<T> P\<close> have \<open>t \<in> \<T> (P \<triangle> Q)\<close>
+              by (auto simp add: \<open>t = u @ v\<close> T_Interrupt)
+            with \<open>length v = n\<close> \<open>t = u @ v\<close> \<open>tF v\<close> \<open>u = []\<close> show \<open>t \<in> \<D> ?lhs\<close>
+              by (simp add: D_restriction_processptickI)
+          next
+            fix v' r assume \<open>v = v' @ [\<checkmark>(r)]\<close> \<open>tF v'\<close>
+            with \<open>v \<in> \<D> Q'\<close> \<open>length t \<le> n\<close> \<open>t = u @ v\<close>
+            have \<open>v' \<in> \<D> Q'\<close> \<open>length v' < n\<close> by (auto intro: is_processT9)
+            with \<open>Q \<down> n = Q' \<down> n\<close> have \<open>v' \<in> \<D> Q\<close>
+              by (metis D_restriction_processptickI
+                  length_less_in_D_restriction_processptick)
+            with \<open>v = v' @ [\<checkmark>(r)]\<close> \<open>tF v'\<close> have \<open>v \<in> \<D> Q\<close> by (simp add: is_processT7)
+            with \<open>tF u\<close> \<open>u \<in> \<T> P\<close> have \<open>t \<in> \<D> (P \<triangle> Q)\<close>
+              by (auto simp add: \<open>t = u @ v\<close> D_Interrupt)
+            thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+          qed
+        next
+          assume \<open>length v \<noteq> n\<close>
+          with \<open>length t \<le> n\<close> \<open>t = u @ v\<close> have \<open>length v < n\<close> by simp
+          with \<open>Q \<down> n = Q' \<down> n\<close> \<open>v \<in> \<D> Q'\<close> have \<open>v \<in> \<D> Q\<close>
+            by (metis D_restriction_processptickI
+                length_less_in_D_restriction_processptick)
+          with \<open>t = u @ v\<close> \<open>tF u\<close> \<open>u \<in> \<T> P\<close> have \<open>t \<in> \<D> (P \<triangle> Q)\<close>
+            by (auto simp add: D_Interrupt)
+          thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+        qed
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (P' \<triangle> Q') \<Longrightarrow> length u = n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> t \<in> \<D> ?lhs\<close> for u v by (fact "*")
+    qed
+
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+    proof (elim F_restriction_processptickE)
+      assume \<open>(t, X) \<in> \<F> (P' \<triangle> Q')\<close> \<open>length t \<le> n\<close>
+      from this(1) consider \<open>t \<in> \<D> (P' \<triangle> Q')\<close>
+        | u r where \<open>t = u @ [\<checkmark>(r)]\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close>
+        | r where \<open>\<checkmark>(r) \<notin> X\<close> \<open>t @ [\<checkmark>(r)] \<in> \<T> P'\<close>
+        | \<open>(t, X) \<in> \<F> P'\<close> \<open>tF t\<close> \<open>([], X) \<in> \<F> Q'\<close>
+        | u v where \<open>t = u @ v\<close> \<open>u \<in> \<T> P'\<close> \<open>tF u\<close> \<open>(v, X) \<in> \<F> Q'\<close> \<open>v \<noteq> []\<close>
+        | r where \<open>\<checkmark>(r) \<notin> X\<close> \<open>t \<in> \<T> P'\<close> \<open>tF t\<close> \<open>[\<checkmark>(r)] \<in> \<T> Q'\<close>
+        unfolding Interrupt_projs by blast
+      thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+      proof cases
+        assume \<open>t \<in> \<D> (P' \<triangle> Q')\<close>
+        hence \<open>t \<in> \<D> ?rhs\<close> by (simp add: D_restriction_processptickI)
+        with div D_F show \<open>t \<in> \<D> (P' \<triangle> Q') \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> by blast
+      next
+        fix u r assume \<open>t = u @ [\<checkmark>(r)]\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close>
+        with \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close> have \<open>u @ [\<checkmark>(r)] \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+        hence \<open>(t, X) \<in> \<F> (P \<triangle> Q)\<close> by (auto simp add: \<open>t = u @ [\<checkmark>(r)]\<close> F_Interrupt)
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+      next
+        fix r assume \<open>\<checkmark>(r) \<notin> X\<close> \<open>t @ [\<checkmark>(r)] \<in> \<T> P'\<close>
+        show \<open>(t, X) \<in> \<F> ?lhs\<close>
+        proof (cases \<open>length t = n\<close>)
+          assume \<open>length t = n\<close>
+          with \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close> \<open>t @ [\<checkmark>(r)] \<in> \<T> P'\<close> have \<open>t \<in> \<T> P\<close>
+            by (metis T_restriction_processptickI is_processT3_TR_append
+                length_le_in_T_restriction_processptick)
+          moreover from \<open>t @ [\<checkmark>(r)] \<in> \<T> P'\<close> append_T_imp_tickFree have \<open>tF t\<close> by blast
+          ultimately have \<open>t \<in> \<T> (P \<triangle> Q)\<close> by (simp add: T_Interrupt)
+          with \<open>length t = n\<close> \<open>tF t\<close> show \<open>(t, X) \<in> \<F> ?lhs\<close>
+            by (simp add: F_restriction_processptickI)
+        next
+          assume \<open>length t \<noteq> n\<close>
+          with \<open>length t \<le> n\<close> have \<open>length (t @ [\<checkmark>(r)]) \<le> n\<close> by simp
+          with \<open>P \<down> n = P' \<down> n\<close> \<open>t @ [\<checkmark>(r)] \<in> \<T> P'\<close> have \<open>t @ [\<checkmark>(r)] \<in> \<T> P\<close>
+            by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+          with \<open>\<checkmark>(r) \<notin> X\<close> have \<open>(t, X) \<in> \<F> (P \<triangle> Q)\<close>
+            by (simp add: F_Interrupt) (metis Diff_insert_absorb)
+          thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+        qed
+      next
+        assume \<open>(t, X) \<in> \<F> P'\<close> \<open>tF t\<close> \<open>([], X) \<in> \<F> Q'\<close>
+        show \<open>(t, X) \<in> \<F> ?lhs\<close>
+        proof (cases \<open>length t = n\<close>)
+          assume \<open>length t = n\<close>
+          from \<open>(t, X) \<in> \<F> P'\<close> \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close> have \<open>t \<in> \<T> P\<close>
+            by (simp add: F_T T_restriction_processptick
+                length_le_in_T_restriction_processptick)
+          hence \<open>t \<in> \<T> (P \<triangle> Q)\<close> by (simp add: T_Interrupt)
+          with \<open>length t = n\<close> \<open>tF t\<close> show \<open>(t, X) \<in> \<F> ?lhs\<close>
+            by (simp add: F_restriction_processptickI)
+        next
+          assume \<open>length t \<noteq> n\<close>
+          with \<open>length t \<le> n\<close> have \<open>length t < n\<close> by simp
+          with \<open>(t, X) \<in> \<F> P'\<close> \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<noteq> n\<close> have \<open>(t, X) \<in> \<F> P\<close>
+            by (metis F_restriction_processptickI length_less_in_F_restriction_processptick)
+          moreover from \<open>([], X) \<in> \<F> Q'\<close> have \<open>([], X) \<in> \<F> Q\<close>
+            by (metis F_restriction_processptickI \<open>0 < n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+                length_less_in_F_restriction_processptick list.size(3))
+          ultimately have \<open>(t, X) \<in> \<F> (P \<triangle> Q)\<close>
+            using \<open>tF t\<close> by (simp add: F_Interrupt)
+          thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+        qed
+      next
+        fix u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> P'\<close> \<open>tF u\<close> \<open>(v, X) \<in> \<F> Q'\<close> \<open>v \<noteq> []\<close>
+        from \<open>t = u @ v\<close> \<open>length t \<le> n\<close> have \<open>length u \<le> n\<close> by simp
+        with \<open>P \<down> n = P' \<down> n\<close> \<open>u \<in> \<T> P'\<close> have \<open>u \<in> \<T> P\<close> 
+          by (simp add: T_restriction_processptick length_le_in_T_restriction_processptick)
+        show \<open>(t, X) \<in> \<F> ?lhs\<close>
+        proof (cases \<open>length v = n\<close>)
+          assume \<open>length v = n\<close>
+          with \<open>t = u @ v\<close> \<open>length t \<le> n\<close> have \<open>u = []\<close> by simp
+          from F_imp_front_tickFree \<open>(v, X) \<in> \<F> Q'\<close> have \<open>ftF v\<close> by blast
+          thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+          proof (elim front_tickFreeE)
+            assume \<open>tF v\<close>
+            from \<open>length v = n\<close> \<open>Q \<down> n = Q' \<down> n\<close> \<open>(v, X) \<in> \<F> Q'\<close> have \<open>v \<in> \<T> Q\<close>
+              by (metis F_T F_restriction_processptickI dual_order.refl
+                  length_le_in_T_restriction_processptick)
+            with \<open>tF u\<close> \<open>u \<in> \<T> P\<close> have \<open>t \<in> \<T> (P \<triangle> Q)\<close>
+              by (auto simp add: \<open>t = u @ v\<close> T_Interrupt)
+            with \<open>length v = n\<close> \<open>t = u @ v\<close> \<open>tF v\<close> \<open>u = []\<close> show \<open>(t, X) \<in> \<F> ?lhs\<close>
+              by (simp add: F_restriction_processptickI)
+          next
+            fix v' r assume \<open>v = v' @ [\<checkmark>(r)]\<close> \<open>tF v'\<close>
+            with \<open>(v, X) \<in> \<F> Q'\<close> \<open>length t \<le> n\<close> \<open>t = u @ v\<close>
+              \<open>Q \<down> n = Q' \<down> n\<close> \<open>u = []\<close> have \<open>v' @ [\<checkmark>(r)] \<in> \<T> Q\<close>
+              by (metis F_T T_restriction_processptickI append_self_conv2
+                  length_le_in_T_restriction_processptick)
+            with \<open>t = u @ v\<close> \<open>tF u\<close> \<open>u \<in> \<T> P\<close> \<open>v = v' @ [\<checkmark>(r)]\<close> have \<open>t \<in> \<T> (P \<triangle> Q)\<close>
+              by (auto simp add: T_Interrupt)
+            thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+              by (simp add: \<open>t = u @ v\<close> \<open>u = []\<close> \<open>v = v' @ [\<checkmark>(r)]\<close>
+                  restriction_processptick_projs(3) tick_T_F)
+          qed
+        next
+          assume \<open>length v \<noteq> n\<close>
+          with \<open>length t \<le> n\<close> \<open>t = u @ v\<close> have \<open>length v < n\<close> by simp
+          with \<open>Q \<down> n = Q' \<down> n\<close> \<open>(v, X) \<in> \<F> Q'\<close> have \<open>(v, X) \<in> \<F> Q\<close>
+            by (metis F_restriction_processptickI
+                length_less_in_F_restriction_processptick)
+          with \<open>t = u @ v\<close> \<open>tF u\<close> \<open>u \<in> \<T> P\<close> \<open>v \<noteq> []\<close> have \<open>(t, X) \<in> \<F> (P \<triangle> Q)\<close>
+            by (simp add: F_Interrupt) blast
+          thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+        qed
+      next
+        fix r assume \<open>\<checkmark>(r) \<notin> X\<close> \<open>t \<in> \<T> P'\<close> \<open>tF t\<close> \<open>[\<checkmark>(r)] \<in> \<T> Q'\<close>
+        have \<open>t \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close>
+              \<open>t \<in> \<T> P'\<close> length_le_in_T_restriction_processptick)
+        moreover have \<open>[\<checkmark>(r)] \<in> \<T> Q\<close>
+          by (metis (no_types, lifting) Suc_leI T_restriction_processptick
+              Un_iff \<open>0 < n\<close> \<open>Q \<down> n = Q' \<down> n\<close> \<open>[\<checkmark>(r)] \<in> \<T> Q'\<close> length_Cons
+              length_le_in_T_restriction_processptick list.size(3))
+        ultimately have \<open>(t, X) \<in> \<F> (P \<triangle> Q)\<close>
+          using \<open>\<checkmark>(r) \<notin> X\<close> \<open>tF t\<close> by (simp add: F_Interrupt) blast
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (P' \<triangle> Q') \<Longrightarrow> length u = n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for u v
+        by (simp add: "*" is_processT8)
+    qed
+  qed
+qed
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Prefixes_Constructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Prefixes_Constructive.thy
@@ -0,0 +1,151 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Constructiveness of Prefixes\<close>
+
+(*<*)
+theory Prefixes_Constructive
+  imports Process_Restriction_Space
+begin
+  (*>*)
+
+subsection \<open>Equality\<close>
+
+lemma restriction_processptick_Mprefix :
+  \<open>\<box>a \<in> A \<rightarrow> P a \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> \<box>a\<in>A \<rightarrow> (P a \<down> m))\<close> (is \<open>?lhs = ?rhs\<close>)
+proof (cases n)
+  show \<open>n = 0 \<Longrightarrow> ?lhs = ?rhs\<close> by simp
+next
+  fix m assume \<open>n = Suc m\<close>
+  show \<open>?lhs = ?rhs\<close>
+  proof (rule Process_eq_optimizedI)
+    show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+      by (auto simp add: \<open>n = Suc m\<close> Mprefix_projs D_restriction_processptick) blast
+  next
+    fix t assume \<open>t \<in> \<D> ?rhs\<close>
+    with D_imp_front_tickFree obtain a t'
+      where \<open>a \<in> A\<close> \<open>t = ev a # t'\<close> \<open>ftF t'\<close> \<open>t' \<in> \<D> (P a \<down> m)\<close>
+      by (auto simp add: \<open>n = Suc m\<close> D_Mprefix)
+    thus \<open>t \<in> \<D> ?lhs\<close>
+      by (simp add: \<open>n = Suc m\<close> D_restriction_processptick Mprefix_projs)
+        (metis append_Cons eventptick.disc(1) length_Cons tickFree_Cons_iff)
+  next
+    show \<open>(t, X) \<in> \<F> ?lhs \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for t X
+      by (auto simp add: \<open>n = Suc m\<close> restriction_processptick_projs Mprefix_projs)
+  next
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> t \<notin> \<D> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+      by (auto simp add: \<open>n = Suc m\<close> restriction_processptick_projs Mprefix_projs)
+  qed
+qed
+
+
+lemma restriction_processptick_Mndetprefix :
+  \<open>\<sqinter>a \<in> A \<rightarrow> P a \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> \<sqinter>a\<in>A \<rightarrow> (P a \<down> m))\<close> (is \<open>?lhs = ?rhs\<close>)
+proof (cases n)
+  show \<open>n = 0 \<Longrightarrow> ?lhs = ?rhs\<close> by simp
+next
+  fix m assume \<open>n = Suc m\<close>
+  show \<open>?lhs = ?rhs\<close>
+  proof (rule Process_eq_optimizedI)
+    show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+      by (auto simp add: \<open>n = Suc m\<close> Mndetprefix_projs D_restriction_processptick) blast
+  next
+    fix t assume \<open>t \<in> \<D> ?rhs\<close>
+    with D_imp_front_tickFree obtain a t'
+      where \<open>a \<in> A\<close> \<open>t = ev a # t'\<close> \<open>ftF t'\<close> \<open>t' \<in> \<D> (P a \<down> m)\<close>
+      by (auto simp add: \<open>n = Suc m\<close> D_Mndetprefix')
+    thus \<open>t \<in> \<D> ?lhs\<close>
+      by (simp add: \<open>n = Suc m\<close> D_restriction_processptick Mndetprefix_projs)
+        (metis append_Cons eventptick.disc(1) length_Cons tickFree_Cons_iff)
+  next
+    show \<open>(t, X) \<in> \<F> ?lhs \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for t X
+      by (auto simp add: \<open>n = Suc m\<close> restriction_processptick_projs
+          Mndetprefix_projs split: if_split_asm)
+  next
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> t \<notin> \<D> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+      by (auto simp add: \<open>n = Suc m\<close> restriction_processptick_projs
+          Mndetprefix_projs split: if_split_asm)
+  qed
+qed
+
+
+corollary restriction_processptick_write0 :
+  \<open>a \<rightarrow> P \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> a \<rightarrow> (P \<down> m))\<close>
+  unfolding write0_def by (simp add: restriction_processptick_Mprefix)
+
+
+corollary restriction_processptick_write :
+  \<open>c!a \<rightarrow> P \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> c!a \<rightarrow> (P \<down> m))\<close>
+  unfolding write_def by (simp add: restriction_processptick_Mprefix)
+
+
+corollary restriction_processptick_read :
+  \<open>c?a\<in>A \<rightarrow> P a \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> c?a\<in>A \<rightarrow> (P a \<down> m))\<close>
+  unfolding read_def comp_def by (simp add: restriction_processptick_Mprefix)
+
+
+corollary restriction_processptick_ndet_write :
+  \<open>c!!a\<in>A \<rightarrow> P a \<down> n = (case n of 0 \<Rightarrow> \<bottom> | Suc m \<Rightarrow> c!!a\<in>A \<rightarrow> (P a \<down> m))\<close>
+  unfolding ndet_write_def comp_def by (simp add: restriction_processptick_Mndetprefix)
+
+
+
+subsection \<open>Constructiveness\<close>
+
+lemma Mprefix_constructive : \<open>constructive (\<lambda>P. \<box>a \<in> A \<rightarrow> P a)\<close>
+  by (auto intro: constructiveI
+      simp add: restriction_processptick_Mprefix restriction_fun_def)
+
+lemma Mndetprefix_constructive : \<open>constructive (\<lambda>P. \<sqinter>a \<in> A \<rightarrow> P a)\<close>
+  by (auto intro: constructiveI
+      simp add: restriction_processptick_Mndetprefix restriction_fun_def)
+
+lemma write0_constructive : \<open>constructive (\<lambda>P. a \<rightarrow> P)\<close>
+  by (auto intro: constructiveI simp add: restriction_processptick_write0)
+
+lemma write_constructive : \<open>constructive (\<lambda>P. c!a \<rightarrow> P)\<close>
+  by (auto intro: constructiveI simp add: restriction_processptick_write)
+
+lemma read_constructive : \<open>constructive (\<lambda>P. c?a \<in> A \<rightarrow> P a)\<close>
+  by (auto intro: constructiveI
+      simp add: restriction_processptick_read restriction_fun_def)
+
+lemma ndet_write_constructive : \<open>constructive (\<lambda>P. c!!a \<in> A \<rightarrow> P a)\<close>
+  by (auto intro: constructiveI
+      simp add: restriction_processptick_ndet_write restriction_fun_def)
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Process_Restriction_Space.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Process_Restriction_Space.thy
@@ -0,0 +1,830 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Depth Operator\<close>
+
+(*<*)
+theory Process_Restriction_Space
+  imports "Restriction_Spaces-HOLCF" "HOL-CSP.Basic_CSP_Laws"
+begin
+  (*>*)
+
+subsection \<open>Definition\<close>
+
+instantiation processptick :: (type, type) order_restriction_space
+begin
+
+lift_definition restriction_processptick ::
+  \<open>[('a, 'r) processptick, nat] \<Rightarrow> ('a, 'r) processptick\<close>
+  is \<open>\<lambda>P n. (\<F> P \<union> {(t @ u, X) |t u X. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u},
+             \<D> P \<union> { t @ u     |t u.   t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u})\<close>
+proof -
+  show \<open>?thesis P n\<close> (is \<open>is_process (?f, ?d)\<close>) for P and n
+  proof (unfold is_process_def FAILURES_def fst_conv DIVERGENCES_def snd_conv, intro conjI impI allI)
+    show \<open>([], {}) \<in> ?f\<close> by (simp add: process_charn)
+  next
+    show \<open>(t, X) \<in> ?f \<Longrightarrow> ftF t\<close> for t X
+      by simp (meson front_tickFree_append is_processT)
+  next
+    fix t u
+    assume \<open>(t @ u, {}) \<in> ?f\<close>
+    then consider \<open>(t @ u, {}) \<in> \<F> P\<close>
+      | t' u' where \<open>t @ u = t' @ u'\<close> \<open>t' \<in> \<T> P\<close> \<open>length t' = n\<close> \<open>tF t'\<close> \<open>ftF u'\<close> by blast
+    thus \<open>(t, {}) \<in> ?f\<close>
+    proof cases
+      assume \<open>(t @ u, {}) \<in> \<F> P\<close>
+      with is_processT3 have \<open>(t, {}) \<in> \<F> P\<close> by auto
+      thus \<open>(t, {}) \<in> ?f\<close> by fast
+    next
+      fix t' u' assume * : \<open>t @ u = t' @ u'\<close> \<open>t' \<in> \<T> P\<close> \<open>length t' = n\<close> \<open>tF t'\<close> \<open>ftF u'\<close>
+      show \<open>(t, {}) \<in> ?f\<close>
+      proof (cases \<open>t \<le> t'\<close>)
+        assume \<open>t \<le> t'\<close>
+        with "*"(2) is_processT3_TR have \<open>t \<in> \<T> P\<close> by auto
+        thus \<open>(t, {}) \<in> ?f\<close> by (simp add: T_F)
+      next
+        assume \<open>\<not> t \<le> t'\<close>
+        with "*"(1) have \<open>t = t' @ take (length t - length t') u'\<close>
+          by (metis (no_types, lifting) Prefix_Order.prefixI append_Nil2
+              diff_is_0_eq nle_le take_all take_append take_eq_Nil)
+        with "*"(2, 3, 4, 5) show \<open>(t, {}) \<in> ?f\<close>
+          by simp (metis append_take_drop_id front_tickFree_dw_closed)
+      qed
+    qed
+  next
+    show \<open>(t, Y) \<in> ?f \<and> X \<subseteq> Y \<Longrightarrow> (t, X) \<in> ?f\<close> for t X Y
+      by simp (meson is_processT4)
+  next
+    show \<open>(t, X) \<in> ?f \<and> (\<forall>c. c \<in> Y \<longrightarrow> (t @ [c], {}) \<notin> ?f) \<Longrightarrow> (t, X \<union> Y) \<in> ?f\<close> for t X Y
+      by (auto simp add: is_processT5)
+  next
+    show \<open>(t @ [\<checkmark>(r)], {}) \<in> ?f \<Longrightarrow> (t, X - {\<checkmark>(r)}) \<in> ?f\<close> for t r X
+      by (simp, elim disjE exE, solves \<open>simp add: is_processT6\<close>)
+        (metis append_assoc butlast_snoc front_tickFree_dw_closed
+          nonTickFree_n_frontTickFree non_tickFree_tick tickFree_append_iff)
+  next
+    from front_tickFree_append is_processT7 tickFree_append_iff
+    show \<open>t \<in> ?d \<and> tF t \<and> ftF u \<Longrightarrow> t @ u \<in> ?d\<close> for t u by fastforce
+  next
+    from D_F show \<open>t \<in> ?d \<Longrightarrow> (t, X) \<in> ?f\<close> for t X by blast
+  next
+    show \<open>t @ [\<checkmark>(r)] \<in> ?d \<Longrightarrow> t \<in> ?d\<close> for t r
+      by simp (metis butlast_append butlast_snoc front_tickFree_iff_tickFree_butlast is_processT9
+          non_tickFree_tick tickFree_Nil tickFree_append_iff tickFree_imp_front_tickFree)
+  qed
+qed
+
+
+
+subsection \<open>Projections\<close>
+
+context fixes P :: \<open>('a, 'r) processptick\<close> begin
+
+lemma F_restriction_processptick :
+  \<open>\<F> (P \<down> n) = \<F> P \<union> {(t @ u, X) |t u X. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u}\<close>
+  by (simp add: Failures_def FAILURES_def restriction_processptick.rep_eq)     
+
+lemma D_restriction_processptick :
+  \<open>\<D> (P \<down> n) = \<D> P \<union> {t @ u |t u. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u}\<close>
+  by (simp add: Divergences_def DIVERGENCES_def restriction_processptick.rep_eq)
+
+lemma T_restriction_processptick :
+  \<open>\<T> (P \<down> n) = \<T> P \<union> {t @ u |t u. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u}\<close>
+  using F_restriction_processptick by (auto simp add: Failures_def Traces_def TRACES_def)
+
+lemmas restriction_processptick_projs = F_restriction_processptick D_restriction_processptick T_restriction_processptick
+
+
+lemma D_restriction_processptickE:
+  assumes \<open>t \<in> \<D> (P \<down> n)\<close>
+  obtains \<open>t \<in> \<D> P\<close> and \<open>length t \<le> n\<close>
+  | u v where \<open>t = u @ v\<close> \<open>u \<in> \<T> P\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+proof -
+  note assms = that
+  from \<open>t \<in> \<D> (P \<down> n)\<close> have \<open>ftF t\<close> by (simp add: D_imp_front_tickFree)
+  from \<open>t \<in> \<D> (P \<down> n)\<close> consider \<open>t \<in> \<D> P\<close>
+    | u v where \<open>t = u @ v\<close> \<open>u \<in> \<T> P\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    by (simp add: D_restriction_processptick) blast
+  thus thesis
+  proof cases
+    show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> P \<Longrightarrow> length u = n \<Longrightarrow> tF u \<Longrightarrow> ftF v \<Longrightarrow> thesis\<close> for u v
+      by (fact assms(2))
+  next
+    show thesis if \<open>t \<in> \<D> P\<close>
+    proof (cases \<open>length t \<le> n\<close>)
+      from \<open>t \<in> \<D> P\<close> show \<open>length t \<le> n \<Longrightarrow> thesis\<close> by (rule assms(1))
+    next
+      show thesis if \<open>\<not> length t \<le> n\<close>
+      proof (intro assms(2) exI)
+        show \<open>t = take n t @ drop n t\<close> by simp
+      next
+        show \<open>take n t \<in> \<T> P\<close> by (metis D_T \<open>t \<in> \<D> P\<close> append_take_drop_id is_processT3_TR_append)
+      next
+        show \<open>length (take n t) = n\<close> by (simp add: min_def \<open>\<not> length t \<le> n\<close>)
+      next
+        show \<open>tF (take n t)\<close> by (metis \<open>ftF t\<close> append_take_drop_id drop_eq_Nil2
+              front_tickFree_append_iff \<open>\<not> length t \<le> n\<close>)
+      next
+        show \<open>ftF (drop n t)\<close> by (metis \<open>ftF t\<close> append_take_drop_id drop_eq_Nil
+              front_tickFree_append_iff that)
+      qed
+    qed
+  qed
+qed
+
+
+lemma F_restriction_processptickE :
+  assumes \<open>(t, X) \<in> \<F> (P \<down> n)\<close>
+  obtains \<open>(t, X) \<in> \<F> P\<close> and \<open>length t \<le> n\<close>
+  | u v where \<open>t = u @ v\<close> \<open>u \<in> \<T> P\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+proof -
+  from \<open>(t, X) \<in> \<F> (P \<down> n)\<close> consider \<open>(t, X) \<in> \<F> P\<close> | \<open>t \<in> \<D> (P \<down> n)\<close>
+    unfolding restriction_processptick_projs by blast
+  thus thesis
+  proof cases
+    show \<open>(t, X) \<in> \<F> P \<Longrightarrow> thesis\<close>
+      by (metis F_T F_imp_front_tickFree append_take_drop_id
+          drop_eq_Nil front_tickFree_nonempty_append_imp
+          is_processT3_TR_append length_take min_def that)
+  next
+    show \<open>t \<in> \<D> (P \<down> n) \<Longrightarrow> thesis\<close> by (meson D_restriction_processptickE is_processT8 that)
+  qed
+qed
+
+
+lemma T_restriction_processptickE :
+  \<open>\<lbrakk>t \<in> \<T> (P \<down> n); t \<in> \<T> P \<Longrightarrow> length t \<le> n \<Longrightarrow> thesis;
+    \<And>u v. t = u @ v \<Longrightarrow> u \<in> \<T> P \<Longrightarrow> length u = n \<Longrightarrow> tF u \<Longrightarrow> ftF v \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis\<close>
+  by (fold T_F_spec, elim F_restriction_processptickE) (simp_all add: T_F)
+
+
+lemmas restriction_processptick_elims =
+  F_restriction_processptickE D_restriction_processptickE T_restriction_processptickE
+
+
+lemma D_restriction_processptickI :
+  \<open>t \<in> \<D> P \<or> t \<in> \<T> P \<and> (length t = n \<and> tF t \<or> n < length t) \<Longrightarrow> t \<in> \<D> (P \<down> n)\<close>
+  by (simp add: D_restriction_processptick, elim disjE conjE)
+    (solves simp, use front_tickFree_Nil in blast, 
+      metis (no_types) T_imp_front_tickFree append_self_conv front_tickFree_nonempty_append_imp
+      id_take_nth_drop is_processT3_TR_append leD length_take min.absorb4 take_all_iff)
+
+lemma F_restriction_processptickI :
+  \<open>(t, X) \<in> \<F> P \<or> t \<in> \<T> P \<and> (length t = n \<and> tF t \<or> n < length t) \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n)\<close>
+  by (metis (mono_tags, lifting) D_restriction_processptickI F_restriction_processptick Un_iff is_processT8)
+
+lemma T_restriction_processptickI :
+  \<open>t \<in> \<T> P \<or> t \<in> \<T> P \<and> (length t = n \<and> tF t \<or> n < length t) \<Longrightarrow> t \<in> \<T> (P \<down> n)\<close>
+  using F_restriction_processptickI T_F_spec by blast
+
+
+
+lemma F_restriction_processptick_Suc_length_iff_F :
+  \<open>(t, X) \<in> \<F> (P \<down> Suc (length t)) \<longleftrightarrow> (t, X) \<in> \<F> P\<close>
+  and D_restriction_processptick_Suc_length_iff_D :
+  \<open>t \<in> \<D> (P \<down> Suc (length t)) \<longleftrightarrow> t \<in> \<D> P\<close>
+  and T_restriction_processptick_Suc_length_iff_T :
+  \<open>t \<in> \<T> (P \<down> Suc (length t)) \<longleftrightarrow> t \<in> \<T> P\<close>
+  by (auto simp add: restriction_processptick_projs)
+
+
+lemma length_less_in_F_restriction_processptick :
+  \<open>length t < n \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> P\<close>
+  by (auto elim: F_restriction_processptickE)
+
+lemma length_le_in_T_restriction_processptick :
+  \<open>length t \<le> n \<Longrightarrow> t \<in> \<T> (P \<down> n) \<Longrightarrow> t \<in> \<T> P\<close>
+  by (auto elim: T_restriction_processptickE)
+
+lemma length_less_in_D_restriction_processptick :
+  \<open>length t < n \<Longrightarrow> t \<in> \<D> (P \<down> n) \<Longrightarrow> t \<in> \<D> P\<close>
+  by (auto elim: D_restriction_processptickE)
+
+lemma not_tickFree_in_F_restriction_processptick_iff :
+  \<open>length t \<le> n \<Longrightarrow> \<not> tF t \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n) \<longleftrightarrow> (t, X) \<in> \<F> P\<close>
+  by (auto simp add: F_restriction_processptick)
+
+lemma not_tickFree_in_D_restriction_processptick_iff :
+  \<open>length t \<le> n \<Longrightarrow> \<not> tF t \<Longrightarrow> t \<in> \<D> (P \<down> n) \<longleftrightarrow> t \<in> \<D> P\<close>
+  by (auto simp add: D_restriction_processptick)
+
+end
+
+
+(* TODO: move this in HOL-CSP ? *)
+lemma front_tickFreeE :
+  \<open>\<lbrakk>ftF t; tF t \<Longrightarrow> thesis; \<And>t' r. t = t' @ [\<checkmark>(r)] \<Longrightarrow> tF t' \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis\<close>
+  by (metis front_tickFree_append_iff nonTickFree_n_frontTickFree not_Cons_self2)
+
+
+subsection \<open>Proof obligation\<close>
+
+instance
+proof intro_classes
+  fix P Q :: \<open>('a, 'r) processptick\<close>
+  have \<open>P \<down> 0 = \<bottom>\<close> by (simp add: BOT_iff_Nil_D D_restriction_processptick)
+  thus \<open>P \<down> 0 \<sqsubseteq>FD Q \<down> 0\<close> by simp
+next
+  show \<open>P \<down> n \<down> m = P \<down> min n m\<close> for P :: \<open>('a, 'r) processptick\<close> and n m
+  proof (rule Process_eq_optimizedI)
+    show \<open>t \<in> \<D> (P \<down> n \<down> m) \<Longrightarrow> t \<in> \<D> (P \<down> min n m)\<close> for t
+      by (elim restriction_processptick_elims)
+        (auto simp add: D_restriction_processptick intro: front_tickFree_append)
+  next
+    show \<open>t \<in> \<D> (P \<down> min n m) \<Longrightarrow> t \<in> \<D> (P \<down> n \<down> m)\<close> for t
+      by (elim restriction_processptick_elims)
+        (auto simp add: restriction_processptick_projs min_def split: if_split_asm)
+  next
+    fix t X assume \<open>(t, X) \<in> \<F> (P \<down> n \<down> m)\<close> \<open>t \<notin> \<D> (P \<down> n \<down> m)\<close>
+    hence \<open>(t, X) \<in> \<F> P \<and> length t \<le> min n m\<close>
+      by (elim F_restriction_processptickE) (auto simp add: restriction_processptick_projs)
+    thus \<open>(t, X) \<in> \<F> (P \<down> min n m)\<close> unfolding F_restriction_processptick by blast
+  next
+    fix t X assume \<open>(t, X) \<in> \<F> (P \<down> min n m)\<close> \<open>t \<notin> \<D> (P \<down> min n m)\<close>
+    hence \<open>(t, X) \<in> \<F> P \<and> length t \<le> min n m\<close>
+      by (elim F_restriction_processptickE) (auto simp add: restriction_processptick_projs)
+    thus \<open>(t, X) \<in> \<F> (P \<down> n \<down> m)\<close> unfolding F_restriction_processptick by blast
+  qed
+next
+  show \<open>P \<sqsubseteq>FD Q \<Longrightarrow> P \<down> n \<sqsubseteq>FD Q \<down> n\<close> for P Q :: \<open>('a, 'r) processptick\<close> and n
+    by (simp add: refine_defs restriction_processptick_projs
+        flip: T_F_spec) blast
+next  
+  fix P Q :: \<open>('a, 'r) processptick\<close> assume \<open>\<not> P \<sqsubseteq>FD Q\<close>
+  then consider t where \<open>t \<in> \<D> Q\<close> \<open>t \<notin> \<D> P\<close>
+    | t X where \<open>(t, X) \<in> \<F> Q\<close> \<open>(t, X) \<notin> \<F> P\<close>
+    unfolding refine_defs by auto
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>FD Q \<down> n\<close>
+  proof cases
+    fix t assume \<open>t \<in> \<D> Q\<close> \<open>t \<notin> \<D> P\<close>
+    with D_restriction_processptick_Suc_length_iff_D
+    have \<open>t \<in> \<D> (Q \<down> Suc (length t)) \<and> t \<notin> \<D> (P \<down> Suc (length t))\<close> by blast
+    hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq>FD Q \<down> Suc (length t)\<close> unfolding refine_defs by blast
+    thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>FD Q \<down> n\<close> ..
+  next
+    fix t X assume \<open>(t, X) \<in> \<F> Q\<close> \<open>(t, X) \<notin> \<F> P\<close>
+    with F_restriction_processptick_Suc_length_iff_F
+    have \<open>(t, X) \<in> \<F> (Q \<down> Suc (length t)) \<and> (t, X) \<notin> \<F> (P \<down> Suc (length t))\<close> by blast
+    hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq>FD Q \<down> Suc (length t)\<close> unfolding refine_defs by blast
+    thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>FD Q \<down> n\<close> ..
+  qed
+qed
+
+
+\<comment>\<open>Of course, we immediately recover the structure of class\<open>restriction_space\<close>.\<close>
+
+corollary \<open>OFCLASS(('a, 'r) processptick, restriction_space_class)\<close>
+  by intro_classes
+
+end
+
+instance processptick :: (type, type) pcpo_restriction_space
+proof intro_classes
+  show \<open>P \<down> 0 \<sqsubseteq> Q \<down> 0\<close> for P Q :: \<open>('a, 'r) processptick\<close>
+    by (metis below_refl restriction_0_related)
+next
+  show \<open>P \<down> n \<sqsubseteq> Q \<down> n\<close> if \<open>P \<sqsubseteq> Q\<close> for P Q :: \<open>('a, 'r) processptick\<close> and n
+  proof (unfold le_approx_def Refusals_after_def, safe)
+    from \<open>P \<sqsubseteq> Q\<close>[THEN le_approx1] \<open>P \<sqsubseteq> Q\<close>[THEN le_approx2T]
+    show \<open>t \<in> \<D> (Q \<down> n) \<Longrightarrow> t \<in> \<D> (P \<down> n)\<close> for t
+      by (simp add: D_restriction_processptick subset_iff) (metis D_T)
+  next
+    from \<open>P \<sqsubseteq> Q\<close>[THEN le_approx2] \<open>P \<sqsubseteq> Q\<close>[THEN le_approx_lemma_T]
+    show \<open>t \<notin> \<D> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (Q \<down> n)\<close>
+      and \<open>t \<notin> \<D> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (Q \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n)\<close> for t X
+      by (auto simp add: restriction_processptick_projs)
+  next
+    from \<open>P \<sqsubseteq> Q\<close>[THEN le_approx3] \<open>P \<sqsubseteq> Q\<close>[THEN le_approx2T]
+    show \<open>t \<in> min_elems (\<D> (P \<down> n)) \<Longrightarrow> t \<in> \<T> (Q \<down> n)\<close> for t
+      by (simp add: min_elems_def restriction_processptick_projs ball_Un subset_iff)
+        (metis is_processT7)
+  qed
+next
+  fix P Q :: \<open>('a, 'r) processptick\<close>
+  assume \<open>\<not> P \<sqsubseteq> Q\<close>
+  then consider t where \<open>t \<in> \<D> Q\<close> \<open>t \<notin> \<D> P\<close>
+    | t X where \<open>t \<notin> \<D> P\<close> \<open>(t, X) \<in> \<F> P \<longleftrightarrow> (t, X) \<notin> \<F> Q\<close>
+    | t where \<open>t \<in> min_elems (\<D> P)\<close> \<open>t \<notin> \<T> Q\<close>
+    unfolding le_approx_def Refusals_after_def by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq> Q \<down> n\<close>
+  proof cases
+    fix t assume \<open>t \<in> \<D> Q\<close> \<open>t \<notin> \<D> P\<close>
+    hence \<open>t \<in> \<D> (Q \<down> Suc (length t))\<close> \<open>t \<notin> \<D> (P \<down> Suc (length t))\<close>
+      by (simp_all add: D_restriction_processptick_Suc_length_iff_D)
+    hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq> Q \<down> Suc (length t)\<close>
+      unfolding le_approx_def by blast
+    thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq> Q \<down> n\<close> ..
+  next
+    fix t X assume \<open>t \<notin> \<D> P\<close> \<open>(t, X) \<in> \<F> P \<longleftrightarrow> (t, X) \<notin> \<F> Q\<close>
+    hence \<open>t \<notin> \<D> (P \<down> Suc (length t))\<close>
+      \<open>(t, X) \<in> \<F> (P \<down> Suc (length t)) \<longleftrightarrow> (t, X) \<notin> \<F> (Q \<down> Suc (length t))\<close>
+      by (simp_all add: D_restriction_processptick_Suc_length_iff_D
+          F_restriction_processptick_Suc_length_iff_F)
+    hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq> Q \<down> Suc (length t)\<close>
+      unfolding le_approx_def Refusals_after_def by blast
+    thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq> Q \<down> n\<close> ..
+  next
+    fix t assume \<open>t \<in> min_elems (\<D> P)\<close> \<open>t \<notin> \<T> Q\<close>
+    hence \<open>t \<in> min_elems (\<D> (P \<down> Suc (length t)))\<close> \<open>t \<notin> \<T> (Q \<down> Suc (length t))\<close>
+      by (simp_all add: min_elems_def D_restriction_processptick_Suc_length_iff_D
+          T_restriction_processptick_Suc_length_iff_T)
+        (meson length_less_in_D_restriction_processptick less_SucI less_length_mono)
+    hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq> Q \<down> Suc (length t)\<close>
+      unfolding le_approx_def by blast
+    thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq> Q \<down> n\<close> ..
+  qed
+next
+  show \<open>chain S \<Longrightarrow> \<exists>P. range S <<| P\<close> for S :: \<open>nat \<Rightarrow> ('a, 'r) processptick\<close>
+    by (simp add: cpo_class.cpo)
+next
+  show \<open>\<exists>P :: ('a, 'r) processptick. \<forall>Q. P \<sqsubseteq> Q\<close> by blast
+qed
+
+
+setup \<open>Sign.add_const_constraint (const_name\<open>restriction\<close>, SOME typ\<open>('a, 'r) processptick \<Rightarrow> nat \<Rightarrow> ('a, 'r) processptick\<close>)\<close>
+  \<comment> \<open>Only allow const\<open>restriction\<close> for typ\<open>('a, 'r) processptick\<close>
+     (otherwise we would often have to specify).\<close>
+
+
+
+subsection \<open>Compatibility with Refinements\<close>
+
+lemma leF_restriction_processptickI: \<open>P \<down> n \<sqsubseteq>F Q \<down> n\<close>
+  if \<open>\<And>s X. (s, X) \<in> \<F> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> (s, X) \<in> \<F> (P \<down> n)\<close>
+proof (unfold failure_refine_def, safe)
+  show \<open>(s, X) \<in> \<F> (Q \<down> n) \<Longrightarrow> (s, X) \<in> \<F> (P \<down> n)\<close> for s X
+  proof (elim F_restriction_processptickE exE conjE)
+    show \<open>(s, X) \<in> \<F> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> (s, X) \<in> \<F> (P \<down> n)\<close>
+      by (simp add: F_restriction_processptick) (meson F_restriction_processptickE that)
+  next
+    fix s' t'
+    assume \<open>s = s' @ t'\<close> \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+    from \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> have \<open>s' \<in> \<T> P\<close>
+      by (metis F_T T_F dual_order.refl length_le_in_T_restriction_processptick that)
+    with \<open>s = s' @ t'\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+    show \<open>(s, X) \<in> \<F> (P \<down> n)\<close> by (simp add: F_restriction_processptick) blast
+  qed
+qed
+
+
+lemma leT_restriction_processptickI: \<open>P \<down> n \<sqsubseteq>T Q \<down> n\<close>
+  if \<open>\<And>s. s \<in> \<T> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> s \<in> \<T> (P \<down> n)\<close>
+proof (unfold trace_refine_def, safe)
+  show \<open>s \<in> \<T> (Q \<down> n) \<Longrightarrow> s \<in> \<T> (P \<down> n)\<close> for s
+  proof (elim T_restriction_processptickE exE conjE)
+    show \<open>s \<in> \<T> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> s \<in> \<T> (P \<down> n)\<close>
+      by (simp add: T_restriction_processptick) (meson T_restriction_processptickE that)
+  next
+    fix s' t'
+    assume \<open>s = s' @ t'\<close> \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+    from \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> have \<open>s' \<in> \<T> P\<close>
+      using length_le_in_T_restriction_processptick that by blast
+    with \<open>s = s' @ t'\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+    show \<open>s \<in> \<T> (P \<down> n)\<close> by (simp add: T_restriction_processptick) blast
+  qed
+qed
+
+
+lemma leDT_restriction_processptickI: \<open>P \<down> n \<sqsubseteq>DT Q \<down> n\<close>
+  if \<open>\<And>s. s \<in> \<T> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> s \<in> \<T> (P \<down> n)\<close>
+    and \<open>\<And>s. length s \<le> n \<Longrightarrow> s \<in> \<D> Q \<Longrightarrow> s \<in> \<D> (P \<down> n)\<close>
+proof (rule leD_leT_imp_leDT)
+  show \<open>P \<down> n \<sqsubseteq>T Q \<down> n\<close> by (simp add: leT_restriction_processptickI that(1))
+next
+  show \<open>P \<down> n \<sqsubseteq>D Q \<down> n\<close>
+  proof (unfold divergence_refine_def, rule subsetI)
+    show \<open>s \<in> \<D> (Q \<down> n) \<Longrightarrow> s \<in> \<D> (P \<down> n)\<close> for s
+    proof (elim D_restriction_processptickE exE conjE)
+      show \<open>s \<in> \<D> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> s \<in> \<D> (P \<down> n)\<close>
+        by (simp add: D_restriction_processptick) (meson D_restriction_processptickE that(2))
+    next
+      fix s' t'
+      assume \<open>s = s' @ t'\<close> \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+      from \<open>s' \<in> \<T> Q\<close> \<open>length s' = n\<close> have \<open>s' \<in> \<T> P\<close>
+        using length_le_in_T_restriction_processptick that(1) by blast
+      with \<open>s = s' @ t'\<close> \<open>length s' = n\<close> \<open>tickFree s'\<close> \<open>front_tickFree t'\<close>
+      show \<open>s \<in> \<D> (P \<down> n)\<close> by (simp add: D_restriction_processptick) blast
+    qed
+  qed
+qed
+
+
+lemma leFD_restriction_processptickI: \<open>P \<down> n \<sqsubseteq>FD Q \<down> n\<close>
+  if \<open>\<And>s X. (s, X) \<in> \<F> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> (s, X) \<in> \<F> (P \<down> n)\<close>
+    and \<open>\<And>s. s \<in> \<D> Q \<Longrightarrow> length s \<le> n \<Longrightarrow> s \<in> \<D> (P \<down> n)\<close>
+proof (rule leF_leD_imp_leFD)
+  show \<open>P \<down> n \<sqsubseteq>F Q \<down> n\<close> by (simp add: leF_restriction_processptickI that(1))
+next
+  show \<open>P \<down> n \<sqsubseteq>D Q \<down> n\<close>  by (meson T_F_spec leDT_imp_leD leDT_restriction_processptickI that)
+qed
+
+
+
+subsection \<open>First Laws\<close>
+
+lemma restriction_processptick_0 [simp] : \<open>P \<down> 0 = \<bottom>\<close>
+  by (simp add: BOT_iff_Nil_D D_restriction_processptick)
+
+lemma restriction_processptick_BOT [simp] : \<open>(\<bottom> :: ('a, 'r) processptick) \<down> n = \<bottom>\<close>
+  by (simp add: BOT_iff_Nil_D D_restriction_processptick D_BOT)
+
+lemma restriction_processptick_is_BOT_iff :
+  \<open>P \<down> n = \<bottom> \<longleftrightarrow> n = 0 \<or> P = \<bottom>\<close>
+  by (auto simp add: BOT_iff_Nil_D D_restriction_processptick)
+
+
+lemma restriction_processptick_STOP [simp] : \<open>STOP \<down> n = (if n = 0 then \<bottom> else STOP)\<close>
+  by (simp add: STOP_iff_T T_restriction_processptick T_STOP)
+
+lemma restriction_processptick_is_STOP_iff : \<open>P \<down> n = STOP \<longleftrightarrow> n \<noteq> 0 \<and> P = STOP\<close>
+  by (simp add: STOP_iff_T T_restriction_processptick set_eq_iff)
+    (metis (no_types, lifting) append_self_conv2 front_tickFree_single gr0I
+      less_numeral_extra(3) list.discI list.size(3) tickFree_Nil)
+
+
+lemma restriction_processptick_SKIP [simp] : \<open>SKIP r \<down> n = (if n = 0 then \<bottom> else SKIP r)\<close>
+  by simp (auto simp add: Process_eq_spec restriction_processptick_projs SKIP_projs)
+
+lemma restriction_processptick_is_SKIP_iff : \<open>P \<down> n = SKIP r \<longleftrightarrow> n \<noteq> 0 \<and> P = SKIP r\<close>
+proof (intro iffI conjI)
+  show \<open>n \<noteq> 0 \<and> P = SKIP r \<Longrightarrow> P \<down> n = SKIP r\<close> by simp
+next
+  show \<open>P \<down> n = SKIP r \<Longrightarrow> n \<noteq> 0\<close> by (metis restriction_processptick_0 SKIP_neq_BOT)
+next
+  show \<open>P \<down> n = SKIP r \<Longrightarrow> P = SKIP r\<close>
+    by (simp add: Process_eq_spec set_eq_iff SKIP_projs
+        restriction_processptick_projs, safe; metis)
+qed
+
+
+lemma restriction_processptick_SKIPS [simp] : \<open>SKIPS R \<down> n = (if n = 0 then \<bottom> else SKIPS R)\<close>
+  by simp (auto simp add: Process_eq_spec restriction_processptick_projs SKIPS_projs)
+
+lemma restriction_processptick_is_SKIPS_iff : \<open>P \<down> n = SKIPS R \<longleftrightarrow> n \<noteq> 0 \<and> P = SKIPS R\<close>
+proof (cases \<open>R = {}\<close>)
+  show \<open>R = {} \<Longrightarrow> P \<down> n = SKIPS R \<longleftrightarrow> n \<noteq> 0 \<and> P = SKIPS R\<close>
+    by (simp add: restriction_processptick_is_STOP_iff)
+next
+  show \<open>P \<down> n = SKIPS R \<longleftrightarrow> n \<noteq> 0 \<and> P = SKIPS R\<close> if \<open>R \<noteq> {}\<close>
+  proof (intro iffI conjI)
+    show \<open>n \<noteq> 0 \<and> P = SKIPS R \<Longrightarrow> P \<down> n = SKIPS R\<close> by simp
+  next
+    show \<open>P \<down> n = SKIPS R \<Longrightarrow> n \<noteq> 0\<close> 
+      by (metis BOT_iff_Nil_D D_SKIPS empty_iff restriction_processptick_0)
+  next
+    show \<open>P \<down> n = SKIPS R \<Longrightarrow> P = SKIPS R\<close>
+      by (simp add: Process_eq_spec \<open>R \<noteq> {}\<close> SKIPS_projs
+          restriction_processptick_projs, safe; blast)
+  qed
+qed
+
+
+
+subsection \<open>Monotony\<close>
+
+subsubsection \<open>term\<open>P \<down> n\<close> is an Approximation of the term\<open>P\<close>\<close>
+
+lemma restriction_processptick_approx_self : \<open>P \<down> n \<sqsubseteq> P\<close>
+proof (unfold le_approx_def Refusals_after_def, safe)
+  show \<open>t \<in> \<D> P \<Longrightarrow> t \<in> \<D> (P \<down> n)\<close> for t by (simp add: D_restriction_processptick)
+next
+  show \<open>t \<notin> \<D> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> P\<close> for t X
+    by (auto simp add: D_restriction_processptick elim: F_restriction_processptickE)
+next
+  show \<open>t \<notin> \<D> (P \<down> n) \<Longrightarrow> (t, X) \<in> \<F> P \<Longrightarrow> (t, X) \<in> \<F> (P \<down> n)\<close> for t X
+    by (auto simp add: restriction_processptick_projs)
+next
+  show \<open>t \<in> min_elems (\<D> (P \<down> n)) \<Longrightarrow> t \<in> \<T> P\<close> for t
+    by (auto simp add: min_elems_def D_restriction_processptick ball_Un D_T)
+      (metis append.right_neutral front_tickFree_charn less_append nil_less2)
+qed
+
+lemma restriction_processptick_FD_self : \<open>P \<down> n \<sqsubseteq>FD P\<close>
+  by (simp add: le_approx_imp_le_ref restriction_processptick_approx_self)
+
+lemma restriction_processptick_F_self : \<open>P \<down> n \<sqsubseteq>F P\<close>
+  by (simp add: restriction_processptick_FD_self leFD_imp_leF)
+
+lemma restriction_processptick_D_self : \<open>P \<down> n \<sqsubseteq>D P\<close>
+  by (simp add: restriction_processptick_FD_self leFD_imp_leD)
+
+lemma restriction_processptick_T_self : \<open>P \<down> n \<sqsubseteq>T P\<close>
+  by (simp add: restriction_processptick_F_self leF_imp_leT) 
+
+lemma restriction_processptick_DT_self : \<open>P \<down> n \<sqsubseteq>DT P\<close>
+  by (simp add: restriction_processptick_D_self restriction_processptick_T_self leD_leT_imp_leDT)
+
+
+
+subsubsection \<open>Monotony of term\<open>(\<down>)\<close>\<close>
+
+lemma Suc_right_mono_restriction_processptick : \<open>P \<down> n \<sqsubseteq> P \<down> Suc n\<close>
+  by (metis restriction_processptick_approx_self restriction_chainD
+      restriction_chain_restrictions)
+
+lemma Suc_right_mono_restriction_processptick_FD : \<open>P \<down> n \<sqsubseteq>FD P \<down> Suc n\<close>
+  by (simp add: Suc_right_mono_restriction_processptick le_approx_imp_le_ref)
+
+lemma Suc_right_mono_restriction_processptick_F  : \<open>P \<down> n \<sqsubseteq>F P \<down> Suc n\<close>
+  by (simp add: Suc_right_mono_restriction_processptick_FD leFD_imp_leF)
+
+lemma Suc_right_mono_restriction_processptick_D  : \<open>P \<down> n \<sqsubseteq>D P \<down> Suc n\<close>
+  by (simp add: Suc_right_mono_restriction_processptick_FD leFD_imp_leD)
+
+lemma Suc_right_mono_restriction_processptick_T  : \<open>P \<down> n \<sqsubseteq>T P \<down> Suc n\<close>
+  by (simp add: Suc_right_mono_restriction_processptick_FD leFD_imp_leF leF_imp_leT)
+
+lemma Suc_right_mono_restriction_processptick_DT : \<open>P \<down> n \<sqsubseteq>DT P \<down> Suc n\<close>
+  by (simp add: Suc_right_mono_restriction_processptick_D
+      Suc_right_mono_restriction_processptick_T leD_leT_imp_leDT)
+
+
+lemma le_right_mono_restriction_processptick : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq> P \<down> m\<close>
+  by (metis restriction_processptick_approx_self restriction_chain_def_ter
+      restriction_chain_restrictions)
+
+
+lemma le_right_mono_restriction_processptick_FD : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq>FD P \<down> m\<close>
+  by (simp add: le_approx_imp_le_ref le_right_mono_restriction_processptick)
+
+lemma le_right_mono_restriction_processptick_F : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq>F P \<down> m\<close>
+  by (simp add: leFD_imp_leF le_right_mono_restriction_processptick_FD)
+
+lemma restriction_processptick_le_right_mono_D : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq>D P \<down> m\<close>
+  by (simp add: leFD_imp_leD le_right_mono_restriction_processptick_FD)
+
+lemma restriction_processptick_le_right_mono_T : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq>T P \<down> m\<close>
+  by (simp add: leF_imp_leT le_right_mono_restriction_processptick_F)
+
+lemma restriction_processptick_le_right_mono_DT : \<open>n \<le> m \<Longrightarrow> P \<down> n \<sqsubseteq>DT P \<down> m\<close>
+  by (simp add: restriction_processptick_le_right_mono_D
+      restriction_processptick_le_right_mono_T leD_leT_imp_leDT)
+
+
+
+subsubsection \<open>Interpretations of Refinements\<close>
+
+lemma ex_not_restriction_leD : \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>D Q \<down> n\<close> if \<open>\<not> P \<sqsubseteq>D Q\<close>
+proof -
+  from \<open>\<not> P \<sqsubseteq>D Q\<close> obtain t where \<open>t \<in> \<D> Q\<close> \<open>t \<notin> \<D> P\<close>
+    unfolding divergence_refine_def by blast
+  hence \<open>t \<in> \<D> (Q \<down> Suc (length t))\<close> \<open>t \<notin> \<D> (P \<down> Suc (length t))\<close>
+    by (simp_all add: D_restriction_processptick_Suc_length_iff_D)
+  hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq>D Q \<down> Suc (length t)\<close>
+    unfolding divergence_refine_def by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>D Q \<down> n\<close> ..
+qed
+
+interpretation PRS_leF : PreorderRestrictionSpace \<open>(\<down>)\<close> \<open>(\<sqsubseteq>F)\<close>
+proof unfold_locales
+  show \<open>P \<down> 0 \<sqsubseteq>F Q \<down> 0\<close> for P Q :: \<open>('a, 'r) processptick\<close> by simp
+next
+  show \<open>P \<sqsubseteq>F Q \<Longrightarrow> P \<down> n \<sqsubseteq>F Q \<down> n\<close> for P Q :: \<open>('a, 'r) processptick\<close> and n
+    by (simp add: failure_refine_def F_restriction_processptick
+        flip: T_F_spec) blast
+next
+  fix P Q :: \<open>('a, 'r) processptick\<close> assume \<open>\<not> P \<sqsubseteq>F Q\<close>
+  then obtain t X where \<open>(t, X) \<in> \<F> Q\<close> \<open>(t, X) \<notin> \<F> P\<close>
+    unfolding failure_refine_def by auto
+  with F_restriction_processptick_Suc_length_iff_F
+  have \<open>(t, X) \<in> \<F> (Q \<down> Suc (length t)) \<and> (t, X) \<notin> \<F> (P \<down> Suc (length t))\<close> by blast
+  hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq>F Q \<down> Suc (length t)\<close> unfolding failure_refine_def by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>F Q \<down> n\<close> ..
+next
+  show \<open>P \<sqsubseteq>F Q \<Longrightarrow> Q \<sqsubseteq>F R \<Longrightarrow> P \<sqsubseteq>F R\<close> for P Q R :: \<open>('a, 'r) processptick\<close>
+    by (fact trans_F)
+qed
+
+
+interpretation PRS_leT : PreorderRestrictionSpace \<open>(\<down>)\<close> \<open>(\<sqsubseteq>T)\<close>
+proof unfold_locales
+  show \<open>P \<down> 0 \<sqsubseteq>T Q \<down> 0\<close> for P Q :: \<open>('a, 'r) processptick\<close> by simp
+next
+  show \<open>P \<sqsubseteq>T Q \<Longrightarrow> P \<down> n \<sqsubseteq>T Q \<down> n\<close> for P Q :: \<open>('a, 'r) processptick\<close> and n
+    by (auto simp add: trace_refine_def T_restriction_processptick)
+next
+  fix P Q :: \<open>('a, 'r) processptick\<close> assume \<open>\<not> P \<sqsubseteq>T Q\<close>
+  then obtain t where \<open>t \<in> \<T> Q\<close> \<open>t \<notin> \<T> P\<close>
+    unfolding trace_refine_def by auto
+  with T_restriction_processptick_Suc_length_iff_T
+  have \<open>t \<in> \<T> (Q \<down> Suc (length t)) \<and> t \<notin> \<T> (P \<down> Suc (length t))\<close> by blast
+  hence \<open>\<not> P \<down> Suc (length t) \<sqsubseteq>T Q \<down> Suc (length t)\<close> unfolding trace_refine_def by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>T Q \<down> n\<close> ..
+next
+  show \<open>P \<sqsubseteq>T Q \<Longrightarrow> Q \<sqsubseteq>T R \<Longrightarrow> P \<sqsubseteq>T R\<close> for P Q R :: \<open>('a, 'r) processptick\<close>
+    by (fact trans_T)
+qed
+
+
+interpretation PRS_leDT : PreorderRestrictionSpace \<open>(\<down>)\<close> \<open>(\<sqsubseteq>DT)\<close>
+proof unfold_locales
+  show \<open>P \<down> 0 \<sqsubseteq>DT Q \<down> 0\<close> for P Q :: \<open>('a, 'r) processptick\<close> by simp
+next
+  show \<open>P \<sqsubseteq>DT Q \<Longrightarrow> P \<down> n \<sqsubseteq>DT Q \<down> n\<close> for P Q :: \<open>('a, 'r) processptick\<close> and n
+    by (auto simp add: refine_defs restriction_processptick_projs)
+next
+  fix P Q :: \<open>('a, 'r) processptick\<close> assume \<open>\<not> P \<sqsubseteq>DT Q\<close>
+  hence \<open>\<not> P \<sqsubseteq>D Q \<or> \<not> P \<sqsubseteq>T Q\<close> unfolding trace_divergence_refine_def by blast
+  with ex_not_restriction_leD PRS_leT.ex_not_restriction_related
+  have \<open>(\<exists>n. \<not> P \<down> n \<sqsubseteq>D Q \<down> n) \<or> (\<exists>n. \<not> P \<down> n \<sqsubseteq>T Q \<down> n)\<close> by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<sqsubseteq>DT Q \<down> n\<close>
+    unfolding trace_divergence_refine_def by blast
+next
+  show \<open>P \<sqsubseteq>DT Q \<Longrightarrow> Q \<sqsubseteq>DT R \<Longrightarrow> P \<sqsubseteq>DT R\<close> for P Q R :: \<open>('a, 'r) processptick\<close>
+    by (fact trans_DT)
+qed
+
+
+
+subsection \<open>Continuity\<close>
+
+context begin
+
+private lemma chain_restriction_processptick : \<open>chain Y \<Longrightarrow> chain (\<lambda>i. Y i \<down> n)\<close>
+  by (simp add: mono_restriction_below po_class.chain_def)
+
+
+private lemma cont_prem_restriction_processptick :
+  \<open>(\<Squnion> i. Y i) \<down> n = (\<Squnion> i. Y i \<down> n)\<close> (is \<open>?lhs = ?rhs\<close>) if \<open>chain Y\<close>
+proof (rule Process_eq_optimizedI)
+  show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+    by (auto simp add: limproc_is_thelub chain_restriction_processptick
+        D_restriction_processptick LUB_projs \<open>chain Y\<close>)
+next
+  show \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    by (simp add: limproc_is_thelub chain_restriction_processptick
+        D_restriction_processptick LUB_projs \<open>chain Y\<close>)
+      (metis D_T append_eq_append_conv is_processT3_TR_append)
+next
+  show \<open>(t, X) \<in> \<F> ?lhs \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for t X
+    by (auto simp add: limproc_is_thelub chain_restriction_processptick
+        F_restriction_processptick LUB_projs \<open>chain Y\<close>)
+next
+  show \<open>(s, X) \<in> \<F> ?rhs \<Longrightarrow> (s, X) \<in> \<F> ?lhs\<close> for s X
+    by (simp add: limproc_is_thelub chain_restriction_processptick
+        F_restriction_processptick LUB_projs \<open>chain Y\<close>)
+      (metis F_T append_eq_append_conv is_processT3_TR_append)
+qed
+
+
+lemma restriction_processptick_cont [simp] : \<open>cont (\<lambda>x. f x \<down> n)\<close> if \<open>cont f\<close>
+proof (rule contI2)
+  show \<open>monofun (\<lambda>x. f x \<down> n)\<close>
+    by (simp add: cont2monofunE mono_restriction_below monofunI \<open>cont f\<close>)
+next
+  show \<open>chain Y \<Longrightarrow> f (\<Squnion>i. Y i) \<down> n \<sqsubseteq> (\<Squnion>i. f (Y i) \<down> n)\<close> for Y
+    by (simp add: ch2ch_cont cont2contlubE cont_prem_restriction_processptick \<open>cont f\<close>)
+qed
+
+end
+
+
+
+subsection \<open>Completeness\<close>
+
+text \<open>Processes are actually an instance of class\<open>complete_restriction_space\<close>.\<close>
+
+lemma chain_restriction_chain :
+  \<open>restriction_chain \<sigma> \<Longrightarrow> chain \<sigma>\<close> for \<sigma> :: \<open>nat \<Rightarrow> ('a, 'r) processptick\<close>
+  by (metis po_class.chainI restriction_processptick_approx_self restriction_chainD)
+
+
+lemma restricted_LUB_restriction_chain_is :
+  \<open>(\<lambda>n. (\<Squnion>n. \<sigma> n) \<down> n) = \<sigma>\<close> if \<open>restriction_chain \<sigma>\<close>
+proof (rule ext)
+  have \<open>chain \<sigma>\<close> by (simp add: chain_restriction_chain \<open>restriction_chain \<sigma>\<close>)
+  moreover have \<open>\<sigma> = (\<lambda>n. \<sigma> n \<down> n)\<close>
+    by (simp add: restricted_restriction_chain_is \<open>restriction_chain \<sigma>\<close>)
+  ultimately have \<open>chain (\<lambda>n. \<sigma> n \<down> n)\<close> by simp
+
+  have \<open>length t < n \<Longrightarrow> t \<in> \<D> (\<sigma> n) \<longleftrightarrow> (\<forall>i. t \<in> \<D> (\<sigma> i))\<close> for t n
+  proof safe
+    show \<open>t \<in> \<D> (\<sigma> i)\<close> if \<open>length t < n\<close> \<open>t \<in> \<D> (\<sigma> n)\<close> for i
+    proof (cases \<open>i \<le> n\<close>)
+      from \<open>t \<in> \<D> (\<sigma> n)\<close> \<open>chain \<sigma>\<close> le_approx_def po_class.chain_mono
+      show \<open>i \<le> n \<Longrightarrow> t \<in> \<D> (\<sigma> i)\<close> by blast
+    next
+      from \<open>length t < n\<close> \<open>t \<in> \<D> (\<sigma> n)\<close> show \<open>\<not> i \<le> n \<Longrightarrow> t \<in> \<D> (\<sigma> i)\<close>
+        by (induct n, simp_all)
+          (metis \<open>restriction_chain \<sigma>\<close> length_less_in_D_restriction_processptick
+            nat_le_linear restriction_chain_def_ter)
+    qed
+  next
+    show \<open>\<forall>i. t \<in> \<D> (\<sigma> i) \<Longrightarrow> t \<in> \<D> (\<sigma> n)\<close> by simp
+  qed
+  hence * : \<open>length t < n \<Longrightarrow> t \<in> \<D> (\<sigma> n) \<longleftrightarrow> t \<in> \<D> (\<Squnion>i. \<sigma> i)\<close> for t n
+    by (simp add: D_LUB \<open>chain \<sigma>\<close> limproc_is_thelub)
+
+  show \<open>(\<Squnion>n. \<sigma> n) \<down> n = \<sigma> n\<close> for n
+  proof (subst (3) \<open>\<sigma> = (\<lambda>n. \<sigma> n \<down> n)\<close>, rule Process_eq_optimizedI)
+    show \<open>t \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n) \<Longrightarrow> t \<in> \<D> (\<sigma> n \<down> n)\<close> for t
+    proof (elim D_restriction_processptickE)
+      show \<open>t \<in> \<D> (\<Squnion>n. \<sigma> n) \<Longrightarrow> length t \<le> n \<Longrightarrow> t \<in> \<D> (\<sigma> n \<down> n)\<close>
+        by (simp add: \<open>chain \<sigma>\<close> limproc_is_thelub D_LUB D_restriction_processptick)
+    next
+      show \<open>\<lbrakk>t = u @ v; u \<in> \<T> (\<Squnion>n. \<sigma> n); length u = n; tF u; ftF v\<rbrakk> \<Longrightarrow> t \<in> \<D> (\<sigma> n \<down> n)\<close> for u v
+        by (auto simp add: \<open>chain \<sigma>\<close> limproc_is_thelub LUB_projs restriction_processptick_projs)
+    qed
+  next
+    fix t assume \<open>t \<in> \<D> (\<sigma> n \<down> n)\<close>
+    hence \<open>ftF t\<close> by (simp add: D_imp_front_tickFree)
+    with \<open>t \<in> \<D> (\<sigma> n \<down> n)\<close> consider \<open>length t < n\<close> \<open>t \<in> \<D> (\<sigma> n)\<close>
+      | t' r where \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>tF t'\<close> \<open>length t' < n\<close> \<open>t' \<in> \<D> (\<sigma> n)\<close>
+      | u v  where \<open>t = u @ v\<close> \<open>u \<in> \<T> (\<sigma> n)\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+      by (auto elim!: D_restriction_processptickE)
+        (metis D_T Suc_le_lessD antisym_conv2 append.right_neutral
+          front_tickFree_Nil front_tickFree_nonempty_append_imp
+          is_processT9 length_append_singleton nonTickFree_n_frontTickFree)
+    thus \<open>t \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n)\<close>
+    proof cases
+      show \<open>length t < n \<Longrightarrow> t \<in> \<D> (\<sigma> n) \<Longrightarrow> t \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n)\<close>
+        by (simp add: D_restriction_processptick "*")
+    next
+      fix t' r assume \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>tF t'\<close> \<open>length t' < n\<close> \<open>t' \<in> \<D> (\<sigma> n)\<close>
+      from \<open>length t' < n\<close> \<open>t' \<in> \<D> (\<sigma> n)\<close> have \<open>t' \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n)\<close>
+        by (simp add: D_restriction_processptick "*")
+      thus \<open>t \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n)\<close>
+        by (simp add: \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>tF t'\<close> is_processT7)
+    next
+      fix u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (\<sigma> n)\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+      from \<open>length u = n\<close> \<open>u \<in> \<T> (\<sigma> n)\<close> have \<open>u \<in> \<T> (\<sigma> (Suc n))\<close>
+        by (metis length_le_in_T_restriction_processptick nat_le_linear
+            restriction_chainD \<open>restriction_chain \<sigma>\<close>)
+      from \<open>chain \<sigma>\<close> \<open>u \<in> \<T> (\<sigma> (Suc n))\<close> D_T le_approx2T po_class.chain_mono
+      have \<open>i \<le> Suc n \<Longrightarrow> u \<in> \<T> (\<sigma> i)\<close> for i by blast
+      moreover have \<open>Suc n < i \<Longrightarrow> u \<in> \<T> (\<sigma> i)\<close> for i
+        by (subst T_restriction_processptick_Suc_length_iff_T[symmetric])
+          (metis \<open>length u = n\<close> \<open>u \<in> \<T> (\<sigma> (Suc n))\<close>
+            restriction_chain_def_bis \<open>restriction_chain \<sigma>\<close>)
+      ultimately have \<open>u \<in> \<T> (\<Squnion>i. \<sigma> i)\<close>
+        by (metis T_LUB_2 \<open>chain \<sigma>\<close> limproc_is_thelub linorder_not_le)
+      with \<open>ftF v\<close> \<open>length u = n\<close> \<open>tF u\<close> show \<open>t \<in> \<D> ((\<Squnion>n. \<sigma> n) \<down> n)\<close>
+        by (auto simp add: \<open>t = u @ v\<close> D_restriction_processptick)
+    qed
+  next
+    show \<open>(t, X) \<in> \<F> ((\<Squnion>n. \<sigma> n) \<down> n) \<Longrightarrow> t \<notin> \<D> ((\<Squnion>n. \<sigma> n) \<down> n) \<Longrightarrow> (t, X) \<in> \<F> (\<sigma> n \<down> n)\<close> for t X
+      by (meson \<open>chain \<sigma>\<close> is_processT8 is_ub_thelub proc_ord2a restriction_processptick_approx_self)
+  next
+    fix t X assume \<open>(t, X) \<in> \<F> (\<sigma> n \<down> n)\<close> \<open>t \<notin> \<D> (\<sigma> n \<down> n)\<close>
+    hence \<open>length t \<le> n\<close> \<open>(t, X) \<in> \<F> (\<sigma> n)\<close> \<open>t \<notin> \<D> (\<sigma> n)\<close>
+      by (auto elim!: F_restriction_processptickE simp add: D_restriction_processptick)
+    thus \<open>(t, X) \<in> \<F> ((\<Squnion>i. \<sigma> i) \<down> n)\<close>
+      by (simp add: F_restriction_processptick)
+        (meson \<open>chain \<sigma>\<close> is_ub_thelub le_approx2)
+  qed
+qed
+
+
+instance processptick :: (type, type) complete_restriction_space
+proof (intro_classes, rule restriction_convergentI)
+  show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> (\<Squnion>i. \<sigma> i)\<close> if \<open>restriction_chain \<sigma>\<close> for \<sigma> :: \<open>nat \<Rightarrow> ('a, 'b) processptick\<close>
+  proof (subst restricted_LUB_restriction_chain_is[symmetric])
+    from \<open>restriction_chain \<sigma>\<close> show \<open>restriction_chain \<sigma>\<close> .
+  next
+    from restriction_tendsto_restrictions
+    show \<open>(\<lambda>n. (\<Squnion>i. \<sigma> i) \<down> n) \<midarrow>\<down>\<rightarrow> (\<Squnion>i. \<sigma> i)\<close> .
+  qed
+qed
+
+
+text \<open>This is a very powerful result. Now we can write fixed-point equations for processes
+      like term\<open>\<upsilon> X. f X\<close>, providing the fact that term\<open>f\<close> is const\<open>constructive\<close>.\<close>
+
+
+
+setup \<open>Sign.add_const_constraint (const_name\<open>restriction\<close>, NONE)\<close>
+  \<comment> \<open>Back to normal.\<close>
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/ROOT
--- /dev/null
+++ thys/HOL-CSP_RS/ROOT
@@ -0,0 +1,57 @@
+chapter AFP
+
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+session "HOL-CSP_RS" = "HOL-CSP_OpSem" +
+  description "Instantiation of HOL-CSP as a restriction Space."
+  options [timeout = 300]
+  sessions
+    "Restriction_Spaces-HOLCF"
+  theories
+    Process_Restriction_Space
+    Prefixes_Constructive
+    Choices_Non_Destructive
+    Renaming_Non_Destructive
+    Sequential_Composition_Non_Destructive
+    Synchronization_Product_Non_Destructive
+    Throw_Non_Destructive
+    Interrupt_Non_Destructive
+    After_Operator_Non_Too_Destructive
+    Hiding_Destructive
+    CSP_Restriction_Adm
+    "HOL-CSP_RS" (global)
+  document_files
+    "root.tex"
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Renaming_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Renaming_Non_Destructive.thy
@@ -0,0 +1,103 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+section \<open>Non Destructiveness of Renaming\<close>
+
+(*<*)
+theory Renaming_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSPM.CSPM_Laws"
+begin
+  (*>*)
+
+
+subsection \<open>Equality\<close>
+
+lemma restriction_processptick_Renaming:
+  \<open>Renaming P f g \<down> n = Renaming (P \<down> n) f g\<close> (is \<open>?lhs = ?rhs\<close>)
+proof (rule Process_eq_optimizedI)
+  show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+    by (auto simp add: Renaming_projs D_restriction_processptick)
+      (metis append.right_neutral front_tickFree_Nil map_eventptick_tickFree,
+        use front_tickFree_append in blast)
+next
+  show \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    by (auto simp add: Renaming_projs D_restriction_processptick
+        front_tickFree_append map_eventptick_tickFree)
+next
+  fix t X assume \<open>(t, X) \<in> \<F> ?lhs\<close> \<open>t \<notin> \<D> ?lhs\<close>
+  then obtain u where \<open>(u, map_eventptick f g -` X) \<in> \<F> P\<close> \<open>t = map (map_eventptick f g) u\<close>
+    by (simp add: Renaming_projs restriction_processptick_projs) blast
+  thus \<open>(t, X) \<in> \<F> ?rhs\<close> by (auto simp add: F_Renaming F_restriction_processptick)
+next
+  fix t X assume \<open>(t, X) \<in> \<F> ?rhs\<close> \<open>t \<notin> \<D> ?rhs\<close>
+  then obtain u where \<open>(u, map_eventptick f g -` X) \<in> \<F> (P \<down> n)\<close> \<open>t = map (map_eventptick f g) u\<close>
+    unfolding Renaming_projs by blast
+  from \<open>(u, map_eventptick f g -` X) \<in> \<F> (P \<down> n)\<close>
+  consider \<open>u \<in> \<D> (P \<down> n)\<close> | \<open>(u, map_eventptick f g -` X) \<in> \<F> P\<close>
+    unfolding restriction_processptick_projs by blast
+  thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+  proof cases
+    assume \<open>u \<in> \<D> (P \<down> n)\<close>
+    hence \<open>t \<in> \<D> ?rhs\<close>
+    proof (elim D_restriction_processptickE)
+      from \<open>t = map (map_eventptick f g) u\<close> show \<open>u \<in> \<D> P \<Longrightarrow> t \<in> \<D> ?rhs\<close>
+        by (cases \<open>tF u\<close>, simp_all add: D_Renaming D_restriction_processptick)
+          (use front_tickFree_Nil in blast,
+            metis D_imp_front_tickFree butlast_snoc div_butlast_when_non_tickFree_iff
+            front_tickFree_iff_tickFree_butlast front_tickFree_single map_append
+            map_eventptick_front_tickFree nonTickFree_n_frontTickFree)
+    next
+      show \<open>\<lbrakk>u = v @ w; v \<in> \<T> P; length v = n; tF v; ftF w\<rbrakk> \<Longrightarrow> t \<in> \<D> ?rhs\<close> for v w
+        by (simp add: D_Renaming D_restriction_processptick \<open>t = map (map_eventptick f g) u\<close>)
+          (use front_tickFree_Nil map_eventptick_front_tickFree in blast)
+    qed
+    with \<open>t \<notin> \<D> ?rhs\<close> have False ..
+    thus \<open>(t, X) \<in> \<F> ?lhs\<close> ..
+  next
+    show \<open>(u, map_eventptick f g -` X) \<in> \<F> P \<Longrightarrow> (t, X) \<in> \<F> (Renaming P f g \<down> n)\<close>
+      by (auto simp add: F_restriction_processptick F_Renaming \<open>t = map (map_eventptick f g) u\<close>)
+  qed
+qed
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma Renaming_non_destructive [simp] :
+  \<open>non_destructive (\<lambda>P. Renaming P f g)\<close>
+  by (auto intro: non_destructiveI simp add: restriction_processptick_Renaming)
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Sequential_Composition_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Sequential_Composition_Non_Destructive.thy
@@ -0,0 +1,293 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Non Destructiveness of Sequential Composition\<close>
+
+(*<*)
+theory Sequential_Composition_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSPM.CSPM"
+begin
+  (*>*)
+
+subsection \<open>Refinement\<close>
+
+lemma restriction_processptick_Seq_FD :
+  \<open>P ; Q \<down> n \<sqsubseteq>FD (P \<down> n) ; (Q \<down> n)\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+proof -
+  have * : \<open>t \<in> \<D> (P \<down> n) \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    by (elim D_restriction_processptickE)
+      (auto simp add: Seq_projs D_restriction_processptick)
+  { fix t u v r w x
+    assume \<open>u @ [\<checkmark>(r)] \<in> \<T> P\<close> \<open>length u < n\<close> \<open>v = w @ x\<close> \<open>w \<in> \<T> Q\<close>
+      \<open>length w = n\<close> \<open>tF w\<close> \<open>ftF x\<close> \<open>t = u @ v\<close>
+    hence \<open>t = (u @ take (n - length u) w) @ drop (n - length u) w @ x \<and>
+           u @ take (n - length u) w \<in> \<T> (P ; Q) \<and>
+           length (u @ take (n - length u) w) = n \<and>
+           tF (u @ take (n - length u) w) \<and> ftF (drop (n - length u) w @ x)\<close>
+      by (simp add: \<open>t = u @ v\<close> T_Seq)
+        (metis append_T_imp_tickFree append_take_drop_id front_tickFree_append
+          is_processT3_TR_append list.distinct(1) tickFree_append_iff)
+    with D_restriction_processptick have \<open>t \<in> \<D> ?lhs\<close> by blast
+  } note ** = this
+
+  show \<open>?lhs \<sqsubseteq>FD ?rhs\<close>
+  proof (unfold refine_defs, safe)
+    show div : \<open>t \<in> \<D> ?lhs\<close> if \<open>t \<in> \<D> ?rhs\<close> for t
+    proof -
+      from \<open>t \<in> \<D> ?rhs\<close> consider \<open>t \<in> \<D> (P \<down> n)\<close>
+        | u v r where \<open>t = u @ v\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> (P \<down> n)\<close> \<open>v \<in> \<D> (Q \<down> n)\<close>
+        unfolding D_Seq by blast
+      thus \<open>t \<in> \<D> ?lhs\<close>
+      proof cases
+        show \<open>t \<in> \<D> (P \<down> n) \<Longrightarrow> t \<in> \<D> ?lhs\<close> by (fact "*")
+      next
+        fix u v r assume \<open>t = u @ v\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> (P \<down> n)\<close> \<open>v \<in> \<D> (Q \<down> n)\<close>
+        from \<open>u @ [\<checkmark>(r)] \<in> \<T> (P \<down> n)\<close> consider \<open>u @ [\<checkmark>(r)] \<in> \<D> (P \<down> n)\<close> | \<open>u @ [\<checkmark>(r)] \<in> \<T> P\<close> \<open>length u < n\<close>
+          by (elim T_restriction_processptickE) (auto simp add: D_restriction_processptick)
+        thus \<open>t \<in> \<D> ?lhs\<close>
+        proof cases
+          show \<open>u @ [\<checkmark>(r)] \<in> \<D> (P \<down> n) \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+            by (metis "*" D_imp_front_tickFree \<open>t = u @ v\<close> \<open>v \<in> \<D> (Q \<down> n)\<close>
+                front_tickFree_append_iff is_processT7 is_processT9 not_Cons_self)
+        next
+          from \<open>v \<in> \<D> (Q \<down> n)\<close> show \<open>u @ [\<checkmark>(r)] \<in> \<T> P \<Longrightarrow> length u < n \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+          proof (elim D_restriction_processptickE exE conjE)
+            show \<open>u @ [\<checkmark>(r)] \<in> \<T> P \<Longrightarrow> v \<in> \<D> Q \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+              by (simp add: \<open>t = u @ v\<close> D_restriction_processptick D_Seq) blast
+          next
+            show \<open>\<lbrakk>u @ [\<checkmark>(r)] \<in> \<T> P; length u < n; v = w @ x; w \<in> \<T> Q;
+                   length w = n; tF w; ftF x\<rbrakk> \<Longrightarrow> t \<in> \<D> ?lhs\<close> for w x
+              using "**" \<open>t = u @ v\<close> by blast
+          qed
+        qed
+      qed
+    qed
+
+    have mono : \<open>(P \<down> n) ; (Q \<down> n) \<sqsubseteq> P ; Q\<close>
+      by (simp add: fun_below_iff mono_Seq restriction_fun_def
+          restriction_processptick_approx_self)
+
+    show \<open>(t, X) \<in> \<F> ?lhs\<close> if \<open>(t, X) \<in> \<F> ?rhs\<close> for t X
+      by (meson F_restriction_processptickI div is_processT8 mono proc_ord2a that)
+  qed
+qed
+
+
+corollary restriction_processptick_MultiSeq_FD :
+  \<open>(SEQ l \<in>@ L. P l) \<down> n \<sqsubseteq>FD SEQ l \<in>@ L. (P l \<down> n)\<close>
+proof (induct L rule: rev_induct)
+  show \<open>(SEQ l \<in>@ []. P l) \<down> n \<sqsubseteq>FD SEQ l \<in>@ []. (P l \<down> n)\<close> by simp
+next
+  fix a L
+  assume hyp: \<open>(SEQ l \<in>@ L. P l) \<down> n \<sqsubseteq>FD SEQ l \<in>@ L. (P l \<down> n)\<close>
+  have \<open>(SEQ l \<in>@ (L @ [a]). P l) \<down> n = (SEQ l \<in>@ L. P l ; P a) \<down> n\<close> by simp
+  also have \<open>\<dots> \<sqsubseteq>FD SEQ l \<in>@ L. (P l \<down> n) ; (P a \<down> n)\<close>
+    by (fact trans_FD[OF restriction_processptick_Seq_FD mono_Seq_FD[OF hyp idem_FD]])
+  also have \<open>\<dots> = SEQ l\<in>@(L @ [a]). (P l \<down> n)\<close> by simp
+  finally show \<open>(SEQ l \<in>@ (L @ [a]). P l) \<down> n \<sqsubseteq>FD \<dots>\<close> .
+qed
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma Seq_non_destructive :
+  \<open>non_destructive (\<lambda>(P :: ('a, 'r) processptick, Q). P ; Q)\<close>
+proof (rule order_non_destructiveI, clarify)
+  fix P P' Q Q' :: \<open>('a, 'r) processptick\<close> and n
+  assume \<open>(P, Q) \<down> n = (P', Q') \<down> n\<close> \<open>0 < n\<close>
+  hence \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+    by (simp_all add: restriction_prod_def)
+  show \<open>P ; Q \<down> n \<sqsubseteq>FD P' ; Q' \<down> n\<close>
+  proof (rule leFD_restriction_processptickI)
+    show div : \<open>t \<in> \<D> (P' ; Q') \<Longrightarrow> t \<in> \<D> (P ; Q \<down> n)\<close> if \<open>length t \<le> n\<close> for t
+    proof (unfold D_Seq, safe)
+      show \<open>t \<in> \<D> P' \<Longrightarrow> t \<in> \<D> (P ; Q \<down> n)\<close>
+        by (simp add: D_restriction_processptick Seq_projs)
+          (metis (no_types, opaque_lifting) D_restriction_processptickE
+            D_restriction_processptickI \<open>P \<down> n = P' \<down> n\<close>)
+    next
+      fix u r v assume \<open>t = u @ v\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>v \<in> \<D> Q'\<close>
+      from \<open>t = u @ v\<close> \<open>length t \<le> n\<close> consider \<open>v = []\<close> \<open>length u = n\<close>
+        | \<open>u = []\<close> \<open>length v = n\<close> | \<open>length u < n\<close> \<open>length v < n\<close>
+        using nless_le by (cases u; cases v, auto)
+      thus \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+      proof cases
+        assume \<open>v = []\<close> \<open>length u = n\<close>
+        from \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> append_T_imp_tickFree is_processT3_TR_append
+        have \<open>tF u\<close> \<open>u \<in> \<T> P'\<close> by auto
+        from \<open>u \<in> \<T> P'\<close> \<open>length u = n\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>u \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI less_or_eq_imp_le
+              length_le_in_T_restriction_processptick)
+        with \<open>tF u\<close> have \<open>u \<in> \<T> (P ; Q)\<close> by (simp add: T_Seq)
+        with \<open>length u = n\<close> show \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+          by (simp add: \<open>v = []\<close> \<open>tF u\<close> D_restriction_processptickI)
+      next
+        assume \<open>u = []\<close> \<open>length v = n\<close>
+        from \<open>0 < n\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>u = []\<close> \<open>P \<down> n = P' \<down> n\<close>
+        have \<open>[\<checkmark>(r)] \<in> \<T> P\<close>
+          by (cases n, simp_all)
+            (metis Suc_leI T_restriction_processptickI length_Cons
+              length_le_in_T_restriction_processptick list.size(3) zero_less_Suc)
+        show \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+        proof (cases \<open>tF v\<close>)
+          assume \<open>tF v\<close>
+          have \<open>v \<in> \<T> Q'\<close> by (simp add: D_T \<open>v \<in> \<D> Q'\<close>)
+          with \<open>length v = n\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>v \<in> \<T> Q\<close>
+            unfolding restriction_fun_def
+            by (metis T_restriction_processptickI less_or_eq_imp_le
+                length_le_in_T_restriction_processptick)
+          with \<open>[\<checkmark>(r)] \<in> \<T> P\<close> have \<open>v \<in> \<T> (P ; Q)\<close>
+            by (simp add: T_Seq) (metis append_Nil)
+          with \<open>length v = n\<close> show \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+            by (simp add: \<open>u = []\<close> \<open>tF v\<close> D_restriction_processptickI)
+        next
+          assume \<open>\<not> tF v\<close>
+          with \<open>u = []\<close> \<open>Q \<down> n = Q' \<down> n\<close> \<open>\<not> tF v\<close> \<open>length t \<le> n\<close>
+            \<open>t = u @ v\<close> \<open>v \<in> \<D> Q'\<close> have \<open>v \<in> \<D> Q\<close>
+            by (metis append_self_conv2 not_tickFree_in_D_restriction_processptick_iff)
+          with \<open>[\<checkmark>(r)] \<in> \<T> P\<close> have \<open>v \<in> \<D> (P ; Q)\<close>
+            by (simp add: D_Seq) (metis append_Nil)
+          thus \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+            by (simp add: D_restriction_processptickI \<open>u = []\<close>)
+        qed
+      next
+        assume \<open>length u < n\<close> \<open>length v < n\<close>
+        from \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>length u < n\<close> \<open>P \<down> n = P' \<down> n\<close>
+        have \<open>u @ [\<checkmark>(r)] \<in> \<T> P\<close>
+          by (metis length_le_in_T_restriction_processptick Suc_le_eq 
+              length_append_singleton T_restriction_processptickI)
+        moreover from \<open>v \<in> \<D> Q'\<close> \<open>length v < n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>v \<in> \<D> Q\<close>
+          by (metis D_restriction_processptickI length_less_in_D_restriction_processptick)
+        ultimately show \<open>u @ v \<in> \<D> (P ; Q \<down> n)\<close>
+          by (auto simp add: D_restriction_processptick D_Seq)
+      qed
+    qed
+
+    fix t X assume \<open>(t, X) \<in> \<F> (P' ; Q')\<close> \<open>length t \<le> n\<close>
+    consider \<open>t \<in> \<D> (P' ; Q')\<close> | \<open>(t, X \<union> range tick) \<in> \<F> P'\<close> \<open>tF t\<close>
+      | u r v where \<open>t = u @ v\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>(v, X) \<in> \<F> Q'\<close>
+      using \<open>(t, X) \<in> \<F> (P' ; Q')\<close> by (auto simp add: Seq_projs)
+    thus \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+    proof cases
+      from div \<open>length t \<le> n\<close> D_F
+      show \<open>t \<in> \<D> (P' ; Q') \<Longrightarrow> (t, X) \<in> \<F> (P ; Q \<down> n)\<close> by blast
+    next
+      show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close> if \<open>(t, X \<union> range tick) \<in> \<F> P'\<close> \<open>tF t\<close>
+      proof (cases \<open>length t = n\<close>)
+        assume \<open>length t = n\<close>
+        from \<open>(t, X \<union> range tick) \<in> \<F> P'\<close> have \<open>t \<in> \<T> P'\<close> by (simp add: F_T)
+        with \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close> have \<open>t \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+        with \<open>tF t\<close> have \<open>t \<in> \<T> (P ; Q)\<close> by (simp add: T_Seq)
+        with \<open>length t = n\<close> \<open>tF t\<close> show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+          by (simp add: F_restriction_processptickI)
+      next
+        assume \<open>length t \<noteq> n\<close>
+        with \<open>length t \<le> n\<close> have \<open>length t < n\<close> by linarith
+        with \<open>P \<down> n = P' \<down> n\<close> \<open>(t, X \<union> range tick) \<in> \<F> P'\<close>
+        have \<open>(t, X \<union> range tick) \<in> \<F> P\<close>
+          by (metis F_restriction_processptickI length_less_in_F_restriction_processptick)
+        with \<open>tF t\<close> show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+          by (simp add: F_restriction_processptickI F_Seq)
+      qed
+    next
+      fix u r v assume \<open>t = u @ v\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>(v, X) \<in> \<F> Q'\<close>
+      from \<open>t = u @ v\<close> \<open>length t \<le> n\<close> consider \<open>v = []\<close> \<open>length u = n\<close>
+        | \<open>u = []\<close> \<open>length v = n\<close> | \<open>length u < n\<close> \<open>length v < n\<close>
+        using nless_le by (cases u; cases v, auto)
+      thus \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+      proof cases
+        assume \<open>v = []\<close> \<open>length u = n\<close>
+        from \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> append_T_imp_tickFree is_processT3_TR_append
+        have \<open>tF u\<close> \<open>u \<in> \<T> P'\<close> by auto
+        from \<open>u \<in> \<T> P'\<close> \<open>length u = n\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>u \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI less_or_eq_imp_le
+              length_le_in_T_restriction_processptick)
+        with \<open>tF u\<close> have \<open>u \<in> \<T> (P ; Q)\<close> by (simp add: T_Seq)
+        with \<open>length u = n\<close> show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+          by (simp add: \<open>v = []\<close> \<open>tF u\<close> F_restriction_processptick)
+            (use \<open>t = u @ v\<close> \<open>tF u\<close> \<open>v = []\<close> front_tickFree_Nil in blast)
+      next
+        assume \<open>u = []\<close> \<open>length v = n\<close>
+        from \<open>0 < n\<close> \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>u = []\<close> \<open>P \<down> n = P' \<down> n\<close>
+        have \<open>[\<checkmark>(r)] \<in> \<T> P\<close>
+          by (cases n, simp_all)
+            (metis Suc_leI T_restriction_processptickI length_Cons
+              length_le_in_T_restriction_processptick list.size(3) zero_less_Suc)
+        show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+        proof (cases \<open>tF v\<close>)
+          assume \<open>tF v\<close>
+          from F_T \<open>(v, X) \<in> \<F> Q'\<close> have \<open>v \<in> \<T> Q'\<close> by blast
+          with \<open>length v = n\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>v \<in> \<T> Q\<close>
+            unfolding restriction_fun_def
+            by (metis T_restriction_processptickI less_or_eq_imp_le
+                length_le_in_T_restriction_processptick)
+          with \<open>[\<checkmark>(r)] \<in> \<T> P\<close> have \<open>v \<in> \<T> (P ; Q)\<close>
+            by (simp add: T_Seq) (metis append_Nil)
+          with \<open>length v = n\<close> have \<open>t \<in> \<D> (P ; Q \<down> n)\<close>
+            by (simp add: \<open>u = []\<close> \<open>tF v\<close> \<open>t = u @ v\<close> D_restriction_processptickI)
+          with D_F show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close> by blast
+        next
+          assume \<open>\<not> tF v\<close>
+          with \<open>u = []\<close> \<open>Q \<down> n = Q' \<down> n\<close> \<open>\<not> tF v\<close> \<open>length t \<le> n\<close>
+            \<open>t = u @ v\<close> \<open>(v, X) \<in> \<F> Q'\<close> have \<open>(v, X) \<in> \<F> Q\<close>
+            by (metis append_self_conv2 not_tickFree_in_F_restriction_processptick_iff)
+          with \<open>[\<checkmark>(r)] \<in> \<T> P\<close> have \<open>(t, X) \<in> \<F> (P ; Q)\<close>
+            by (simp add: \<open>t = u @ v\<close> \<open>u = []\<close> F_Seq) (metis append_Nil)
+          thus \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+            by (simp add: F_restriction_processptickI)
+        qed
+      next
+        assume \<open>length u < n\<close> \<open>length v < n\<close>
+        from \<open>u @ [\<checkmark>(r)] \<in> \<T> P'\<close> \<open>length u < n\<close> \<open>P \<down> n = P' \<down> n\<close>
+        have \<open>u @ [\<checkmark>(r)] \<in> \<T> P\<close>
+          by (metis length_le_in_T_restriction_processptick Suc_le_eq 
+              length_append_singleton T_restriction_processptickI)
+        moreover from \<open>(v, X) \<in> \<F> Q'\<close> \<open>length v < n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>(v, X) \<in> \<F> Q\<close>
+          by (metis F_restriction_processptickI length_less_in_F_restriction_processptick)
+        ultimately show \<open>(t, X) \<in> \<F> (P ; Q \<down> n)\<close>
+          by (auto simp add: \<open>t = u @ v\<close> F_restriction_processptick F_Seq)
+      qed
+    qed
+  qed
+qed
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Synchronization_Product_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Synchronization_Product_Non_Destructive.thy
@@ -0,0 +1,653 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Non Destructiveness of Synchronization Product\<close>
+
+(*<*)
+theory Synchronization_Product_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSPM.CSPM"
+begin
+  (*>*)
+
+subsection \<open>Preliminaries\<close>
+
+lemma D_Sync_optimized :
+  \<open>\<D> (P \<lbrakk>A\<rbrakk> Q) =
+   {v @ w |t u v w. tF v \<and> ftF w \<and>
+                    v setinterleaves ((t, u), range tick \<union> ev ` A) \<and>
+                    (t \<in> \<D> P \<and> u \<in> \<T> Q \<or> t \<in> \<D> Q \<and> u \<in> \<T> P)}\<close>
+  (is \<open>_ = ?rhs\<close>)
+proof (intro subset_antisym subsetI)
+  show \<open>d \<in> ?rhs \<Longrightarrow> d \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close> for d
+    by (auto simp add: D_Sync)
+next
+  fix d assume \<open>d \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+  then obtain t u v w
+    where * : \<open>d = v @ w\<close> \<open>ftF w\<close> \<open>tF v \<or> w = []\<close>
+      \<open>v setinterleaves ((t, u), range tick \<union> ev ` A)\<close>
+      \<open>t \<in> \<D> P \<and> u \<in> \<T> Q \<or> t \<in> \<D> Q \<and> u \<in> \<T> P\<close>
+    unfolding D_Sync by blast
+  show \<open>d \<in> ?rhs\<close>
+  proof (cases \<open>tF v\<close>)
+    from "*" show \<open>tF v \<Longrightarrow> d \<in> ?rhs\<close> by blast
+  next
+    assume \<open>\<not> tF v\<close>
+    with "*"(1, 3) have \<open>w = []\<close> \<open>d = v\<close> by simp_all
+    from D_imp_front_tickFree \<open>d = v\<close> \<open>d \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+    have \<open>ftF v\<close> by blast
+    with \<open>\<not> tF v\<close> obtain r v' where \<open>v = v' @ [\<checkmark>(r)]\<close>
+      by (meson nonTickFree_n_frontTickFree)
+    with "*"(4) obtain t' u'
+      where ** : \<open>t = t' @ [\<checkmark>(r)]\<close> \<open>u = u' @ [\<checkmark>(r)]\<close>
+        \<open>v' setinterleaves ((t', u'), range tick \<union> ev ` A)\<close>
+      by (simp add: \<open>v = v' @ [\<checkmark>(r)]\<close>)
+        (meson "*"(5) D_imp_front_tickFree SyncWithTick_imp_NTF T_imp_front_tickFree)
+    have \<open>t' \<in> \<D> P \<and> u' \<in> \<T> Q \<or> t' \<in> \<D> Q \<and> u' \<in> \<T> P\<close>
+      by (metis "*"(5) "**"(1,2) is_processT3_TR_append is_processT9)
+    with "**"(3) \<open>d = v\<close> \<open>ftF v\<close> \<open>v = v' @ [\<checkmark>(r)]\<close>
+      front_tickFree_nonempty_append_imp show \<open>d \<in> ?rhs\<close> by blast
+  qed
+qed
+
+lemma tickFree_interleave_iff :
+  \<open>t setinterleaves ((u, v), S) \<Longrightarrow> tF t \<longleftrightarrow> tF u \<and> tF v\<close>
+  by (induct \<open>(u, S, v)\<close> arbitrary: t u v rule: setinterleaving.induct)
+    (auto split: if_split_asm option.split_asm)
+
+lemma interleave_subsetL :
+  \<open>tF t \<Longrightarrow> {a. ev a \<in> set u} \<subseteq> A \<Longrightarrow>
+   t setinterleaves ((u, v), range tick \<union> ev ` A) \<Longrightarrow> t = v\<close>
+  for t u v :: \<open>('a, 'r) traceptick\<close>
+proof (induct \<open>(u, range tick \<union> ev ` A :: ('a, 'r) refusalptick, v)\<close>
+    arbitrary: t u v rule: setinterleaving.induct)
+  case 1 thus ?case by simp
+next
+  case (2 y v) thus ?case by (auto simp add: image_iff split: if_split_asm)
+next
+  case (3 x u) thus ?case
+    by (simp add: image_iff subset_iff split: if_split_asm)
+      (metis (mono_tags, lifting) eventptick.exhaust)
+next
+  case (4 x u y v)
+  from "4.prems" show ?case
+    apply (simp add: subset_iff split: if_split_asm)
+       apply (metis (no_types, lifting) "4.hyps"(1) Un_iff
+        mem_Collect_eq subsetI tickFree_Cons_iff)
+      apply (metis (no_types, lifting) "4.hyps"(2,4) "4.prems"(2,3) SyncHd_Tl
+        SyncSameHdTl list.sel(1) setinterleaving_sym tickFree_Cons_iff)
+    by (metis eventptick.exhaust imageI rangeI)+
+qed
+
+lemma interleave_subsetR :
+  \<open>tF t \<Longrightarrow> {a. ev a \<in> set v} \<subseteq> A \<Longrightarrow>
+   t setinterleaves ((u, v), range tick \<union> ev ` A) \<Longrightarrow> t = u\<close>
+  by (simp add: interleave_subsetL setinterleaving_sym)
+
+
+lemma interleave_imp_lengthLR_le :
+  \<open>t setinterleaves ((u, v), S) \<Longrightarrow>
+   length u \<le> length t \<and> length v \<le> length t\<close>
+  by (induct \<open>(u, S, v)\<close> arbitrary: t u v rule: setinterleaving.induct;
+      simp split: if_split_asm; use nat_le_linear not_less_eq_eq in fastforce)
+
+lemma interleave_le_prefixLR :
+  \<open>t setinterleaves ((u, v), S) \<Longrightarrow> u' \<le> u \<Longrightarrow> v' \<le> v \<Longrightarrow>
+   (\<exists>t' \<le> t. \<exists>v'' \<le> v'. t' setinterleaves ((u', v''), S)) \<or>
+   (\<exists>t' \<le> t. \<exists>u'' \<le> u'. t' setinterleaves ((u'', v'), S))\<close>
+proof (induct \<open>(u, S, v)\<close>
+    arbitrary: t u u' v v' rule: setinterleaving.induct)
+  case 1
+  then show ?case by simp
+next
+  case (2 y v)
+  thus ?case by (simp split: if_split_asm)
+      (metis si_empty1 insert_iff nil_le)
+next
+  case (3 x u)
+  thus ?case by (simp split: if_split_asm)
+      (metis si_empty1 insert_iff nil_le)
+next
+  case (4 x u y v)
+  show ?case
+  proof (cases \<open>u' = [] \<or> v' = []\<close>)
+    show \<open>u' = [] \<or> v' = [] \<Longrightarrow> ?case\<close> by force
+  next
+    assume \<open>\<not> (u' = [] \<or> v' = [])\<close>
+    with "4.prems"(2, 3)
+    obtain u'' v'' where \<open>u' = x # u''\<close> \<open>u'' \<le> u\<close> \<open>v' = y # v''\<close> \<open>v'' \<le> v\<close>
+      by (meson Prefix_Order.prefix_Cons)
+    with "4.prems"(1) consider (both_in)   t' where \<open>x \<in> S\<close> \<open>y \<in> S\<close> \<open>x = y\<close> \<open>t = x # t'\<close>
+      \<open>t' setinterleaves ((u, v), S)\<close>
+    |      (inR_mvL)   t' where \<open>x \<notin> S\<close> \<open>y \<in> S\<close> \<open>t = x # t'\<close>
+      \<open>t' setinterleaves ((u, y # v), S)\<close>
+    |      (inL_mvR)   t' where \<open>x \<in> S\<close> \<open>y \<notin> S\<close> \<open>t = y # t'\<close>
+      \<open>t' setinterleaves ((x # u, v), S)\<close>
+    |      (notin_mvL) t' where \<open>x \<notin> S\<close> \<open>y \<notin> S\<close> \<open>t = x # t'\<close>
+      \<open>t' setinterleaves ((u, y # v), S)\<close>
+    |      (notin_mvR) t' where \<open>x \<notin> S\<close> \<open>y \<notin> S\<close> \<open>t = y # t'\<close>
+      \<open>t' setinterleaves ((x # u, v), S)\<close>
+      by (auto split: if_split_asm)
+    thus ?case
+    proof cases
+      case both_in
+      from "4.hyps"(1)[OF both_in(1-3, 5) \<open>u'' \<le> u\<close> \<open>v'' \<le> v\<close>]
+      show ?thesis
+      proof (elim disjE exE conjE)
+        fix t'' v'''
+        assume \<open>t'' \<le> t'\<close> \<open>v''' \<le> v''\<close> \<open>t'' setinterleaves ((u'', v'''), S)\<close>
+        hence \<open>y # t'' \<le> t \<and> y # v''' \<le> v' \<and>
+              (y # t'') setinterleaves ((u', y # v'''), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> both_in(2-4))
+        thus ?thesis by blast
+      next
+        fix t'' u'''
+        assume \<open>t'' \<le> t'\<close> \<open>u''' \<le> u''\<close> \<open>t'' setinterleaves ((u''', v''), S)\<close>
+        hence \<open>x # t'' \<le> t \<and> x # u''' \<le> u' \<and>
+              (x # t'') setinterleaves ((x # u''', v'), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> both_in(2-4))
+        thus ?thesis by blast
+      qed
+    next
+      case inR_mvL
+      from "4.hyps"(5)[simplified, OF inR_mvL(1, 2 ,4) \<open>u'' \<le> u\<close> \<open>v' \<le> y # v\<close>]
+      show ?thesis
+      proof (elim disjE exE conjE)
+        fix t'' v'''
+        assume \<open>t'' \<le> t'\<close> \<open>v''' \<le> v'\<close> \<open>t'' setinterleaves ((u'', v'''), S)\<close>
+        hence \<open>x # t'' \<le> t \<and> v''' \<le> v' \<and>
+              (x # t'') setinterleaves ((u', v'''), S)\<close>
+          by (cases v''') (simp_all add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> inR_mvL(1, 3))
+        thus ?thesis by blast
+      next
+        fix t'' u'''
+        assume \<open>t'' \<le> t'\<close> \<open>u''' \<le> u''\<close> \<open>t'' setinterleaves ((u''', v'), S)\<close>
+        hence \<open>x # t'' \<le> t \<and> x # u''' \<le> u' \<and>
+              (x # t'') setinterleaves ((x # u''', v'), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> inR_mvL(1, 3))
+        thus ?thesis by blast
+      qed
+    next
+      case inL_mvR
+      from "4.hyps"(2)[OF inL_mvR(1, 2, 4) \<open>u' \<le> x # u\<close> \<open>v'' \<le> v\<close>]
+      show ?thesis
+      proof (elim disjE exE conjE)
+        fix t'' v'''
+        assume \<open>t'' \<le> t'\<close> \<open>v''' \<le> v''\<close> \<open>t'' setinterleaves ((u', v'''), S)\<close>
+        hence \<open>y # t'' \<le> t \<and> y # v''' \<le> v' \<and>
+               (y # t'') setinterleaves ((u', y # v'''), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> inL_mvR(2, 3))
+        thus ?thesis by blast
+      next
+        fix t'' u'''
+        assume \<open>t'' \<le> t'\<close> \<open>u''' \<le> u'\<close> \<open>t'' setinterleaves ((u''', v''), S)\<close>
+        hence \<open>y # t'' \<le> t \<and> u''' \<le> u' \<and>
+               (y # t'') setinterleaves ((u''', v'), S)\<close>
+          by (cases u''') (simp_all add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> inL_mvR(2, 3))
+        thus ?thesis by blast
+      qed
+    next
+      case notin_mvL
+      from "4.hyps"(3)[OF notin_mvL(1, 2, 4) \<open>u'' \<le> u\<close> \<open>v' \<le> y # v\<close>]
+      show ?thesis
+      proof (elim disjE exE conjE)
+        fix t'' v'''
+        assume \<open>t'' \<le> t'\<close> \<open>v''' \<le> v'\<close> \<open>t'' setinterleaves ((u'', v'''), S)\<close>
+        hence \<open>x # t'' \<le> t \<and> v''' \<le> v' \<and>
+               (x # t'') setinterleaves ((u', v'''), S)\<close>
+          by (cases v''') (simp_all add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> notin_mvL(1, 3))
+        thus ?thesis by blast
+      next
+        fix t'' u'''
+        assume \<open>t'' \<le> t'\<close> \<open>u''' \<le> u''\<close> \<open>t'' setinterleaves ((u''', v'), S)\<close>
+        hence \<open>x # t'' \<le> t \<and> x # u''' \<le> u' \<and>
+               (x # t'') setinterleaves ((x # u''', v'), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> notin_mvL(1, 3))
+        thus ?thesis by blast
+      qed
+    next
+      case notin_mvR
+      from "4.hyps"(4)[OF notin_mvR(1, 2, 4) \<open>u' \<le> x # u\<close> \<open>v'' \<le> v\<close>]
+      show ?thesis
+      proof (elim disjE exE conjE)
+        fix t'' v'''
+        assume \<open>t'' \<le> t'\<close> \<open>v''' \<le> v''\<close> \<open>t'' setinterleaves ((u', v'''), S)\<close>
+        hence \<open>y # t'' \<le> t \<and> y # v''' \<le> v' \<and>
+               (y # t'') setinterleaves ((u', y # v'''), S)\<close>
+          by (simp add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> notin_mvR(2, 3))
+        thus ?thesis by blast
+      next
+        fix t'' u'''
+        assume \<open>t'' \<le> t'\<close> \<open>u''' \<le> u'\<close> \<open>t'' setinterleaves ((u''', v''), S)\<close>
+        hence \<open>y # t'' \<le> t \<and> u''' \<le> u' \<and>
+               (y # t'') setinterleaves ((u''', v'), S)\<close>
+          by (cases u''') (simp_all add: \<open>u' = x # u''\<close> \<open>v' = y # v''\<close> notin_mvR(2, 3))
+        thus ?thesis by blast
+      qed
+    qed
+  qed
+qed
+
+
+lemma restriction_processptick_Sync_FD_div_oneside :
+  assumes \<open>tF u\<close> \<open>ftF v\<close> \<open>t_P \<in> \<D> (P \<down> n)\<close> \<open>t_Q \<in> \<T> (Q \<down> n)\<close>
+    \<open>u setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+  shows \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+proof (insert assms(3, 4), elim D_restriction_processptickE T_restriction_processptickE)
+  from assms(1, 2, 5) show \<open>t_P \<in> \<D> P \<Longrightarrow> t_Q \<in> \<T> Q \<Longrightarrow> u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    by (auto simp add: D_restriction_processptick D_Sync)
+next
+  fix t_Q' t_Q''
+  assume * : \<open>t_P \<in> \<D> P\<close> \<open>length t_P \<le> n\<close> \<open>t_Q = t_Q' @ t_Q''\<close>
+    \<open>t_Q' \<in> \<T> Q\<close> \<open>length t_Q' = n\<close> \<open>tF t_Q'\<close> \<open>ftF t_Q''\<close>
+  from \<open>t_Q = t_Q' @ t_Q''\<close> have \<open>t_Q' \<le> t_Q\<close> by simp
+  from interleave_le_right[OF assms(5) this]
+  obtain t_P' t_P'' u' u''
+    where ** : \<open>u = u' @ u''\<close> \<open>t_P = t_P' @ t_P''\<close>
+      \<open>u' setinterleaves ((t_P', t_Q'), range tick \<union> ev ` A)\<close>
+    by (meson Prefix_Order.prefixE)
+  from assms(1) \<open>u = u' @ u''\<close> have \<open>tF u'\<close> by auto
+  moreover from "*"(1,4) "**"(2,3) have \<open>u' \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+    by (simp add: T_Sync) (metis D_T is_processT3_TR_append)
+  moreover have \<open>length t_Q' \<le> length u'\<close>
+    using "**"(3) interleave_imp_lengthLR_le by blast
+  ultimately have \<open>u' \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    by (metis "*"(5) D_restriction_processptickI nless_le)
+  with "**"(1) assms(1, 2) show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    by (metis is_processT7 tickFree_append_iff tickFree_imp_front_tickFree)
+next
+  fix t_P' t_P''
+  assume * : \<open>t_P = t_P' @ t_P''\<close> \<open>t_P' \<in> \<T> P\<close> \<open>length t_P' = n\<close>
+    \<open>tF t_P'\<close> \<open>ftF t_P''\<close> \<open>t_Q \<in> \<T> Q\<close> \<open>length t_Q \<le> n\<close>
+  from \<open>t_P = t_P' @ t_P''\<close> have \<open>t_P' \<le> t_P\<close> by simp
+  from interleave_le_left[OF assms(5) this]
+  obtain t_Q' t_Q'' u' u''
+    where ** : \<open>u = u' @ u''\<close> \<open>t_Q = t_Q' @ t_Q''\<close>
+      \<open>u' setinterleaves ((t_P', t_Q'), range tick \<union> ev ` A)\<close>
+    by (meson Prefix_Order.prefixE)
+  from assms(1) \<open>u = u' @ u''\<close> have \<open>tF u'\<close> by auto
+  moreover from "*"(2,6) "**"(2,3) have \<open>u' \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+    by (simp add: T_Sync) (metis is_processT3_TR_append)
+  moreover have \<open>length t_P' \<le> length u'\<close>
+    using "**"(3) interleave_imp_lengthLR_le by blast
+  ultimately have \<open>u' \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    by (metis "*"(3) D_restriction_processptickI nless_le)
+  with "**"(1) assms(1, 2) show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    by (metis is_processT7 tickFree_append_iff tickFree_imp_front_tickFree)
+next
+  fix t_P' t_P'' t_Q' t_Q''
+  assume $ : \<open>t_P = t_P' @ t_P''\<close> \<open>t_P' \<in> \<T> P\<close> \<open>length t_P' = n\<close>
+    \<open>tF t_P'\<close> \<open>ftF t_P''\<close> \<open>t_Q = t_Q' @ t_Q''\<close> \<open>t_Q' \<in> \<T> Q\<close>
+    \<open>length t_Q' = n\<close> \<open>tF t_Q'\<close> \<open>ftF t_Q''\<close>
+  from "$"(1, 6) have \<open>t_P' \<le> t_P\<close> \<open>t_Q' \<le> t_Q\<close> by simp_all
+  from interleave_le_prefixLR[OF assms(5) this]
+  show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+  proof (elim disjE conjE exE)
+    fix u' t_Q''' assume $$ : \<open>u' \<le> u\<close> \<open>t_Q''' \<le> t_Q'\<close>
+      \<open>u' setinterleaves ((t_P', t_Q'''), range tick \<union> ev ` A)\<close>
+    from "$"(7) "$$"(2) is_processT3_TR have \<open>t_Q''' \<in> \<T> Q\<close> by blast
+    with $$(3) \<open>t_P' \<in> \<T> P\<close> have \<open>u' \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+      by (auto simp add: T_Sync)
+    moreover have \<open>n \<le> length u'\<close>
+      using "$"(3) "$$"(3) interleave_imp_lengthLR_le by blast
+    ultimately have \<open>u' \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      by (metis "$$"(1) D_restriction_processptickI Prefix_Order.prefixE
+          assms(1) nless_le tickFree_append_iff)
+    thus \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      by (metis "$$"(1) Prefix_Order.prefixE assms(1,2) is_processT7
+          tickFree_append_iff tickFree_imp_front_tickFree)
+  next
+    fix u' t_P''' assume $$ : \<open>u' \<le> u\<close> \<open>t_P''' \<le> t_P'\<close>
+      \<open>u' setinterleaves ((t_P''', t_Q'), range tick \<union> ev ` A)\<close>
+    from "$"(2) "$$"(2) is_processT3_TR have \<open>t_P''' \<in> \<T> P\<close> by blast
+    with $$(3) \<open>t_Q' \<in> \<T> Q\<close> have \<open>u' \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+      by (auto simp add: T_Sync)
+    moreover have \<open>n \<le> length u'\<close>
+      using "$"(8) "$$"(3) interleave_imp_lengthLR_le by blast
+    ultimately have \<open>u' \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      by (metis "$$"(1) D_restriction_processptickI Prefix_Order.prefixE
+          assms(1) nless_le tickFree_append_iff)
+    thus \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      by (metis "$$"(1) Prefix_Order.prefixE assms(1,2) is_processT7
+          tickFree_append_iff tickFree_imp_front_tickFree)
+  qed
+qed
+
+
+subsection \<open>Refinement\<close>
+
+
+
+lemma restriction_processptick_Sync_FD :
+  \<open>P \<lbrakk>A\<rbrakk> Q \<down> n \<sqsubseteq>FD (P \<down> n) \<lbrakk>A\<rbrakk> (Q \<down> n)\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+proof (unfold refine_defs, safe)
+  show \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    by (unfold D_Sync_optimized, safe)
+      (solves \<open>simp add: restriction_processptick_Sync_FD_div_oneside\<close>,
+        metis Sync_commute restriction_processptick_Sync_FD_div_oneside)
+  thus \<open>(t, X) \<in> \<F> ((P \<down> n) \<lbrakk>A\<rbrakk> (Q \<down> n)) \<Longrightarrow> (t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close> for t X
+    by (meson is_processT8 le_approx2 mono_Sync restriction_processptick_approx_self)
+qed
+
+text \<open>The equality does not hold in general, but we can establish it
+      by adding an assumption over the strict alphabets of the processes.\<close>
+
+lemma strict_events_of_subset_restriction_processptick_Sync :
+  \<open>P \<lbrakk>A\<rbrakk> Q \<down> n = (P \<down> n) \<lbrakk>A\<rbrakk> (Q \<down> n)\<close> (is \<open>?lhs = ?rhs\<close>)
+  if \<open>\<alpha>(P) \<subseteq> A \<or> \<alpha>(Q) \<subseteq> A\<close>
+proof (rule FD_antisym)
+  show \<open>?lhs \<sqsubseteq>FD ?rhs\<close> by (fact restriction_processptick_Sync_FD)
+next
+  have div : \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q) \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+    by (auto simp add: D_Sync restriction_processptick_projs)
+
+  { fix t u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    from this(2) consider \<open>u \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+      | t_P t_Q where \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+        \<open>u setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+      unfolding Sync_projs by blast
+    hence \<open>t \<in> \<D> ?rhs\<close>
+    proof cases
+      show \<open>u \<in> \<D> (P \<lbrakk>A\<rbrakk> Q) \<Longrightarrow> t \<in> \<D> ?rhs\<close>
+        by (simp add: \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> div is_processT7)
+    next
+      fix t_P t_Q assume \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+        and setinter : \<open>u setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+      consider \<open>t_P \<in> \<D> P \<or> t_Q \<in> \<D> Q\<close> | \<open>t_P \<notin> \<D> P\<close> \<open>t_Q \<notin> \<D> Q\<close> by blast
+      thus \<open>t \<in> \<D> ?rhs\<close>
+      proof cases
+        assume \<open>t_P \<in> \<D> P \<or> t_Q \<in> \<D> Q\<close>
+        with \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> setinter \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close>
+        have \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          using setinterleaving_sym by (simp add: D_Sync) blast
+        thus \<open>t \<in> \<D> ?rhs\<close> by (fact div)
+      next
+        assume \<open>t_P \<notin> \<D> P\<close> \<open>t_Q \<notin> \<D> Q\<close>
+        with \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> \<open>\<alpha>(P) \<subseteq> A \<or> \<alpha>(Q) \<subseteq> A\<close>
+        have \<open>{a. ev a \<in> set t_P} \<subseteq> A \<or> {a. ev a \<in> set t_Q} \<subseteq> A\<close>
+          by (auto dest: subsetD intro: strict_events_of_memI)
+        with interleave_subsetL[OF \<open>tF u\<close> _ setinter]
+          interleave_subsetR[OF \<open>tF u\<close> _ setinter]
+        have \<open>u = t_P \<or> u = t_Q\<close> by blast
+        with \<open>length u = n\<close> have \<open>length t_P = n \<or> length t_Q = n\<close> by auto
+        moreover from \<open>tF u\<close> tickFree_interleave_iff[OF setinter]
+        have \<open>tF t_P\<close> \<open>tF t_Q\<close> by simp_all
+        ultimately have \<open>t_P \<in> \<D> (P \<down> n) \<or> t_Q \<in> \<D> (Q \<down> n)\<close>
+          using \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> by (metis D_restriction_processptickI)
+        moreover from \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+        have \<open>t_P \<in> \<T> (P \<down> n)\<close> \<open>t_Q \<in> \<T> (Q \<down> n)\<close>
+          by (simp_all add: T_restriction_processptickI)
+        ultimately show \<open>t \<in> \<D> ?rhs\<close>
+          using \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> setinter
+          by (simp add: D_Sync_optimized)
+            (metis setinterleaving_sym)
+      qed
+    qed
+  } note * = this
+
+  show \<open>?rhs \<sqsubseteq>FD ?lhs\<close>
+  proof (unfold refine_defs, safe)
+    show \<open>t \<in> \<D> ?lhs \<Longrightarrow> t \<in> \<D> ?rhs\<close> for t
+    proof (elim D_restriction_processptickE)
+      show \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q) \<Longrightarrow> t \<in> \<D> ?rhs\<close> by (fact div)
+    next
+      show \<open>\<lbrakk>t = u @ v; u \<in> \<T> (P \<lbrakk>A\<rbrakk> Q); length u = n; tF u; ftF v\<rbrakk>
+            \<Longrightarrow> t \<in> \<D> ?rhs\<close> for u v by (fact "*")
+    qed
+  next
+    show \<open>(t, X) \<in> \<F> ?lhs \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for t X
+    proof (elim F_restriction_processptickE)
+      assume \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q)\<close>
+      then consider \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+        | (fail) t_P t_Q X_P X_Q where \<open>(t_P, X_P) \<in> \<F> P\<close> \<open>(t_Q, X_Q) \<in> \<F> Q\<close>
+          \<open>t setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+          \<open>X = (X_P \<union> X_Q) \<inter> (range tick \<union> ev ` A) \<union> X_P \<inter> X_Q\<close>
+        unfolding Sync_projs by blast
+      thus \<open>(t, X) \<in> \<F> ?rhs\<close>
+      proof cases
+        from div D_F show \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q) \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> by blast
+      next
+        case fail
+        thus \<open>(t, X) \<in> \<F> ?rhs\<close>
+          by (auto simp add: F_Sync F_restriction_processptick)
+      qed
+    next
+      show \<open>\<lbrakk>t = u @ v; u \<in> \<T> (P \<lbrakk>A\<rbrakk> Q); length u = n; tF u; ftF v\<rbrakk>
+            \<Longrightarrow> (t, X) \<in> \<F> ?rhs\<close> for u v by (simp add: "*" is_processT8)
+    qed
+  qed
+qed
+
+
+corollary restriction_processptick_MultiSync_FD :
+  \<open>\<lbrakk>A\<rbrakk> m \<in># M. P l \<down> n \<sqsubseteq>FD \<lbrakk>A\<rbrakk> m \<in># M. (P l \<down> n)\<close>
+proof (induct M rule: induct_subset_mset_empty_single)
+  case 1 show ?case by simp
+next
+  case (2 m) show ?case by simp
+next
+  case (3 N m)
+  show ?case
+    by (simp add: \<open>N \<noteq> {#}\<close>)
+      (fact trans_FD[OF restriction_processptick_Sync_FD mono_Sync_FD[OF idem_FD "3.hyps"(4)]])
+qed
+
+
+
+text \<open>In the following corollary, we could be more precise by having
+      the condition on at least term\<open>size M - 1\<close> processes.\<close>
+
+corollary strict_events_of_subset_restriction_processptick_MultiSync :
+  \<open>\<lbrakk>A\<rbrakk> m \<in># M. P m \<down> n = (if n = 0 then \<bottom> else \<lbrakk>A\<rbrakk> m \<in># M. (P m \<down> n))\<close>
+  \<comment>\<open>\<open>if n = 0 then \<bottom> else _\<close> is necessary because we can have term\<open>M = {#}\<close>.\<close>
+  if \<open>\<And>m. m \<in># M \<Longrightarrow> \<alpha>(P m) \<subseteq> A\<close>
+proof (split if_split, intro conjI impI)
+  show \<open>n = 0 \<Longrightarrow> \<lbrakk>A\<rbrakk> m \<in># M. P m \<down> n = \<bottom>\<close> by simp
+next
+  show \<open>\<lbrakk>A\<rbrakk> m \<in># M. P m \<down> n = \<lbrakk>A\<rbrakk> m \<in># M. (P m \<down> n)\<close> if \<open>n \<noteq> 0\<close>
+  proof (induct M rule: induct_subset_mset_empty_single)
+    case 1 from \<open>n \<noteq> 0\<close> show ?case by simp
+  next
+    case (2 m) show ?case by simp
+  next
+    case (3 N m)
+    have \<open>(\<lbrakk>A\<rbrakk> n\<in>#add_mset m N. P n) \<down> n = (P m \<lbrakk>A\<rbrakk> \<lbrakk>A\<rbrakk> n\<in>#N. P n) \<down> n\<close>
+      by (simp add: \<open>N \<noteq> {#}\<close>)
+    also have \<open>\<dots> = (P m \<down> n) \<lbrakk>A\<rbrakk> (\<lbrakk>A\<rbrakk> n\<in>#N. P n \<down> n)\<close>
+      by (rule strict_events_of_subset_restriction_processptick_Sync)
+        (simp add: "3.hyps"(1) \<open>\<And>m. m \<in># M \<Longrightarrow> \<alpha>(P m) \<subseteq> A\<close>)
+    also have \<open>(\<lbrakk>A\<rbrakk> m\<in>#N. P m) \<down> n = \<lbrakk>A\<rbrakk> m\<in>#N. (P m \<down> n)\<close> by (fact "3.hyps"(4))
+    finally show ?case by (simp add: \<open>N \<noteq> {#}\<close>)
+  qed
+qed
+
+
+
+corollary restriction_processptick_Par :
+  \<open>P || Q \<down> n = (P \<down> n) || (Q \<down> n)\<close>
+  by (simp add: strict_events_of_subset_restriction_processptick_Sync)
+
+corollary restriction_processptick_MultiPar :
+  \<open>|| m \<in># M. P l \<down> n = (if n = 0 then \<bottom> else || m \<in># M. (P l \<down> n))\<close>
+  by (simp add: strict_events_of_subset_restriction_processptick_MultiSync)
+
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma Sync_non_destructive :
+  \<open>non_destructive (\<lambda>(P, Q). P \<lbrakk>A\<rbrakk> Q)\<close>
+proof (rule order_non_destructiveI, clarify)
+  fix P P' Q Q' :: \<open>('a, 'r) processptick\<close> and n
+  assume \<open>(P, Q) \<down> n = (P', Q') \<down> n\<close>
+  hence \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+    by (simp_all add: restriction_prod_def)
+  show \<open>P \<lbrakk>A\<rbrakk> Q \<down> n \<sqsubseteq>FD P' \<lbrakk>A\<rbrakk> Q' \<down> n\<close>
+  proof (rule leFD_restriction_processptickI)
+    show div : \<open>t \<in> \<D> (P' \<lbrakk>A\<rbrakk> Q') \<Longrightarrow> t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close> if \<open>length t \<le> n\<close> for t
+    proof (unfold D_Sync_optimized, safe)
+      fix u v t_P t_Q
+      assume * : \<open>t = u @ v\<close> \<open>tF u\<close> \<open>ftF v\<close>
+        \<open>u setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+        \<open>t_P \<in> \<D> P'\<close> \<open>t_Q \<in> \<T> Q'\<close>
+      from "*"(1) \<open>length t \<le> n\<close> have \<open>length u \<le> n\<close> by simp
+      from \<open>length u \<le> n\<close> interleave_imp_lengthLR_le[OF "*"(4)]
+      have \<open>length t_P \<le> n\<close> \<open>length t_Q \<le> n\<close> by simp_all
+      from \<open>t_Q \<in> \<T> Q'\<close> \<open>length t_Q \<le> n\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>t_Q \<in> \<T> Q\<close>
+        by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+      show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      proof (cases \<open>length u = n\<close>)
+        assume \<open>length u = n\<close>
+        from \<open>t_P \<in> \<D> P'\<close> \<open>length t_P \<le> n\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>t_P \<in> \<T> P\<close>
+          by (simp add: D_T D_restriction_processptickI length_le_in_T_restriction_processptick)
+        with \<open>t_Q \<in> \<T> Q\<close> "*"(4) have \<open>u \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          unfolding T_Sync by blast
+        with \<open>length u = n\<close> "*"(2, 3) show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+          by (simp add: D_restriction_processptickI is_processT7)
+      next
+        assume \<open>length u \<noteq> n\<close>
+        with \<open>length u \<le> n\<close> have \<open>length u < n\<close> by simp
+        with interleave_imp_lengthLR_le[OF "*"(4)]
+        have \<open>length t_P < n\<close> by simp
+        with \<open>t_P \<in> \<D> P'\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>t_P \<in> \<D> P\<close>
+          by (metis D_restriction_processptickI
+              length_less_in_D_restriction_processptick)
+        with \<open>t_Q \<in> \<T> Q\<close> "*"(2-4) have \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          unfolding D_Sync by blast
+        thus \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+          by (simp add: D_restriction_processptickI)
+      qed
+    next
+      fix u v t_P t_Q
+      assume * : \<open>t = u @ v\<close> \<open>tF u\<close> \<open>ftF v\<close>
+        \<open>u setinterleaves ((t_Q, t_P), range tick \<union> ev ` A)\<close>
+        \<open>t_P \<in> \<T> P'\<close> \<open>t_Q \<in> \<D> Q'\<close>
+      from "*"(1) \<open>length t \<le> n\<close> have \<open>length u \<le> n\<close> by simp
+      from \<open>length u \<le> n\<close> interleave_imp_lengthLR_le[OF "*"(4)]
+      have \<open>length t_P \<le> n\<close> \<open>length t_Q \<le> n\<close> by simp_all
+      from \<open>t_P \<in> \<T> P'\<close> \<open>length t_P \<le> n\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>t_P \<in> \<T> P\<close>
+        by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+      show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      proof (cases \<open>length u = n\<close>)
+        assume \<open>length u = n\<close>
+        from \<open>t_Q \<in> \<D> Q'\<close> \<open>length t_Q \<le> n\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>t_Q \<in> \<T> Q\<close>
+          by (simp add: D_T D_restriction_processptickI length_le_in_T_restriction_processptick)
+        with \<open>t_P \<in> \<T> P\<close> "*"(4) have \<open>u \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          unfolding T_Sync using setinterleaving_sym by blast
+        with \<open>length u = n\<close> "*"(2, 3) show \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+          by (simp add: D_restriction_processptickI is_processT7)
+      next
+        assume \<open>length u \<noteq> n\<close>
+        with \<open>length u \<le> n\<close> have \<open>length u < n\<close> by simp
+        with interleave_imp_lengthLR_le[OF "*"(4)]
+        have \<open>length t_Q < n\<close> by simp
+        with \<open>t_Q \<in> \<D> Q'\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>t_Q \<in> \<D> Q\<close>
+          by (metis D_restriction_processptickI
+              length_less_in_D_restriction_processptick)
+        with \<open>t_P \<in> \<T> P\<close> "*"(2-4) have \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          unfolding D_Sync by blast
+        thus \<open>u @ v \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+          by (simp add: D_restriction_processptickI)
+      qed
+    qed
+
+    fix t X assume \<open>(t, X) \<in> \<F> (P' \<lbrakk>A\<rbrakk> Q')\<close> \<open>length t \<le> n\<close>
+    then consider \<open>t \<in> \<D> (P' \<lbrakk>A\<rbrakk> Q')\<close>
+      | (fail) t_P t_Q X_P X_Q
+      where \<open>(t_P, X_P) \<in> \<F> P'\<close> \<open>(t_Q, X_Q) \<in> \<F> Q'\<close>
+        \<open>t setinterleaves ((t_P, t_Q), range tick \<union> ev ` A)\<close>
+        \<open>X = (X_P \<union> X_Q) \<inter> (range tick \<union> ev ` A) \<union> X_P \<inter> X_Q\<close>
+      unfolding Sync_projs by blast
+    thus \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+    proof cases
+      from div \<open>length t \<le> n\<close> D_F
+      show \<open>t \<in> \<D> (P' \<lbrakk>A\<rbrakk> Q') \<Longrightarrow> (t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close> by blast
+    next
+      case fail
+      show \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+      proof (cases \<open>length t = n\<close>)
+        assume \<open>length t = n\<close>
+        from \<open>length t \<le> n\<close> interleave_imp_lengthLR_le[OF fail(3)]
+        have \<open>length t_P \<le> n\<close> \<open>length t_Q \<le> n\<close> by simp_all
+        with fail(1, 2) \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+          by (metis F_T T_restriction_processptickI
+              length_le_in_T_restriction_processptick)+
+        from F_imp_front_tickFree \<open>(t, X) \<in> \<F> (P' \<lbrakk>A\<rbrakk> Q')\<close>
+        have \<open>ftF t\<close> by blast
+        with fail(3) consider \<open>tF t\<close>
+          | r s t_P' t_Q' where \<open>t_P = t_P' @ [\<checkmark>(r)]\<close> \<open>t_Q = t_Q' @ [\<checkmark>(s)]\<close>
+          by (metis F_imp_front_tickFree SyncWithTick_imp_NTF
+              fail(1-2) nonTickFree_n_frontTickFree)
+        thus \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+        proof cases
+          assume \<open>tF t\<close>
+          from \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> fail(3)
+          have \<open>t \<in> \<T> (P \<lbrakk>A\<rbrakk> Q)\<close> unfolding T_Sync by blast
+          hence \<open>t \<in> \<D> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+            by (simp add: D_restriction_processptickI \<open>length t = n\<close> \<open>tF t\<close>)
+          thus \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close> by (simp add: is_processT8)
+        next
+          fix r s t_P' t_Q' assume \<open>t_P = t_P' @ [\<checkmark>(r)]\<close> \<open>t_Q = t_Q' @ [\<checkmark>(s)]\<close>
+          have \<open>(t_P, X_P) \<in> \<F> P\<close>
+            by (metis \<open>t_P \<in> \<T> P\<close> \<open>t_P = t_P' @ [\<checkmark>(r)]\<close> tick_T_F)
+          moreover have \<open>(t_Q, X_Q) \<in> \<F> Q\<close>
+            by (metis \<open>t_Q \<in> \<T> Q\<close> \<open>t_Q = t_Q' @ [\<checkmark>(s)]\<close> tick_T_F)
+          ultimately have \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q)\<close>
+            using fail(3, 4) unfolding F_Sync by fast
+          thus \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+            by (simp add: F_restriction_processptickI)
+        qed
+      next
+        assume \<open>length t \<noteq> n\<close>
+        with \<open>length t \<le> n\<close> have \<open>length t < n\<close> by simp
+        with interleave_imp_lengthLR_le[OF fail(3)]
+        have \<open>length t_P < n\<close> \<open>length t_Q < n\<close> by simp_all
+        with fail(1, 2) \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>(t_P, X_P) \<in> \<F> P\<close> \<open>(t_Q, X_Q) \<in> \<F> Q\<close>
+          by (metis F_restriction_processptickI
+              length_less_in_F_restriction_processptick)+
+        with fail(3, 4) have \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q)\<close>
+          unfolding F_Sync by fast
+        thus \<open>(t, X) \<in> \<F> (P \<lbrakk>A\<rbrakk> Q \<down> n)\<close>
+          by (simp add: F_restriction_processptickI)
+      qed
+    qed
+  qed
+qed
+
+
+(*>*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/Throw_Non_Destructive.thy
--- /dev/null
+++ thys/HOL-CSP_RS/Throw_Non_Destructive.thy
@@ -0,0 +1,529 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Non Destructiveness of Throw\<close>
+
+
+(*<*)
+theory Throw_Non_Destructive
+  imports Process_Restriction_Space "HOL-CSP_OpSem.Initials"
+begin
+  (*>*)
+
+
+subsection \<open>Equality\<close>
+
+lemma Depth_Throw_1_is_constant: \<open>P \<Theta> a \<in> A. Q1 a \<down> 1 = P \<Theta> a \<in> A. Q2 a \<down> 1\<close>
+proof (rule FD_antisym)
+  show \<open>P \<Theta> a \<in> A. Q2 a \<down> 1 \<sqsubseteq>FD P \<Theta> a \<in> A. Q1 a \<down> 1\<close> for Q1 Q2
+  proof (unfold refine_defs, safe)
+    show div :  \<open>t \<in> \<D> (Throw P A Q1 \<down> 1) \<Longrightarrow> t \<in> \<D> (Throw P A Q2 \<down> 1)\<close> for t
+    proof (elim D_restriction_processptickE) 
+      assume \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q1 a)\<close> and \<open>length t \<le> 1\<close>
+      from \<open>length t \<le> 1\<close> consider \<open>t = []\<close> | e where \<open>t = [e]\<close> by (cases t) simp_all
+      thus \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+      proof cases
+        from \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q1 a)\<close> show \<open>t = [] \<Longrightarrow> t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+          by (simp add: D_restriction_processptick D_Throw)
+      next
+        fix e assume \<open>t = [e]\<close>
+        with \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q1 a)\<close>
+        consider \<open>[] \<in> \<D> P\<close> | a where \<open>t = [ev a]\<close> \<open>[ev a] \<in> \<D> P\<close> \<open>a \<notin> A\<close>
+          | a where \<open>t = [ev a]\<close> \<open>[ev a] \<in> \<T> P\<close> \<open>a \<in> A\<close>
+          by (auto simp add: D_Throw disjoint_iff image_iff)
+            (metis D_T append_Nil eventptick.exhaust process_charn,
+              metis append_Nil empty_iff empty_set hd_append2 hd_in_set in_set_conv_decomp set_ConsD)
+        thus \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+        proof cases
+          show \<open>t = [ev a] \<Longrightarrow> [ev a] \<in> \<T> P \<Longrightarrow> a \<in> A \<Longrightarrow> t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close> for a
+            by (simp add: D_restriction_processptick T_Throw)
+              (metis append_Nil append_self_conv front_tickFree_charn inf_bot_left
+                is_ev_def is_processT1_TR length_0_conv length_Cons list.set(1)
+                tickFree_Cons_iff tickFree_Nil)
+        next
+          show \<open>[] \<in> \<D> P \<Longrightarrow> t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+            by (simp flip: BOT_iff_Nil_D add: D_BOT)
+              (use \<open>t = [e]\<close> front_tickFree_single in blast)
+        next
+          show \<open>t = [ev a] \<Longrightarrow> [ev a] \<in> \<D> P \<Longrightarrow> a \<notin> A \<Longrightarrow> t \<in> \<D> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close> for a  
+            by (simp add: D_restriction_processptick D_Throw disjoint_iff image_iff)
+              (metis append.right_neutral empty_set eventptick.disc(1) eventptick.sel(1)
+                front_tickFree_Nil list.simps(15) singletonD tickFree_Cons_iff tickFree_Nil)
+        qed
+      qed
+    next
+      fix u v assume * : \<open>t = u @ v\<close> \<open>u \<in> \<T> (Throw P A Q1)\<close> \<open>length u = 1\<close> \<open>tF u\<close> \<open>ftF v\<close>
+      from \<open>length u = 1\<close> \<open>tF u\<close> obtain a where \<open>u = [ev a]\<close>
+        by (cases u) (auto simp add: is_ev_def)
+      with "*"(2) show \<open>t \<in> \<D> (Throw P A Q2 \<down> 1)\<close>
+        by (simp add: \<open>t = u @ v\<close> D_restriction_processptick Throw_projs Cons_eq_append_conv)
+          (metis (no_types) "*"(3-5) One_nat_def append_Nil empty_set inf_bot_left
+            insert_disjoint(2) is_processT1_TR list.simps(15))
+    qed
+
+    show \<open>(t, X) \<in> \<F> (Throw P A Q1 \<down> 1) \<Longrightarrow> (t, X) \<in> \<F> (Throw P A Q2 \<down> 1)\<close> for t X
+    proof (elim F_restriction_processptickE)
+      assume \<open>(t, X) \<in> \<F> (P \<Theta> a \<in> A. Q1 a)\<close> \<open>length t \<le> 1\<close>
+      then consider \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q1 a)\<close> | \<open>(t, X) \<in> \<F> P\<close> \<open>set t \<inter> ev ` A = {}\<close>
+        | a where \<open>t = [ev a]\<close> \<open>[ev a] \<in> \<T> P\<close> \<open>a \<in> A\<close>
+        by (auto simp add: F_Throw D_Throw)
+      thus \<open>(t, X) \<in> \<F> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+      proof cases
+        from D_F div \<open>length t \<le> 1\<close>
+        show \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q1 a) \<Longrightarrow> (t, X) \<in> \<F> (P \<Theta> a \<in> A. Q2 a \<down> 1)\<close>
+          using D_restriction_processptickI by blast
+      next
+        show \<open>(t, X) \<in> \<F> P \<Longrightarrow> set t \<inter> ev ` A = {} \<Longrightarrow> (t, X) \<in> \<F> (Throw P A Q2 \<down> 1)\<close>
+          by (simp add: F_restriction_processptick F_Throw)
+      next
+        show \<open>\<lbrakk>t = [ev a]; [ev a] \<in> \<T> P; a \<in> A\<rbrakk> \<Longrightarrow> (t, X) \<in> \<F> (Throw P A Q2 \<down> 1)\<close> for a
+          by (simp add: F_restriction_processptick T_Throw)
+            (metis append.right_neutral append_Nil empty_set eventptick.disc(1)
+              front_tickFree_Nil inf_bot_left is_processT1_TR length_Cons
+              list.size(3) tickFree_Cons_iff tickFree_Nil)
+      qed
+    next
+      fix u v assume * : \<open>t = u @ v\<close> \<open>u \<in> \<T> (Throw P A Q1)\<close> \<open>length u = 1\<close> \<open>tF u\<close> \<open>ftF v\<close>
+      from \<open>length u = 1\<close> \<open>tF u\<close> obtain a where \<open>u = [ev a]\<close>
+        by (cases u) (auto simp add: is_ev_def)
+      with "*"(2) show \<open>(t, X) \<in> \<F> (Throw P A Q2 \<down> 1)\<close>
+        by (simp add: \<open>t = u @ v\<close> F_restriction_processptick Throw_projs Cons_eq_append_conv)
+          (metis (no_types) "*"(3-5) One_nat_def append_Nil empty_set inf_bot_left
+            insert_disjoint(2) is_processT1_TR list.simps(15))
+    qed
+  qed
+
+  thus \<open>P \<Theta> a \<in> A. Q2 a \<down> 1 \<sqsubseteq>FD P \<Theta> a \<in> A. Q1 a \<down> 1\<close> by simp
+qed
+
+
+
+subsection \<open>Refinement\<close>
+
+lemma restriction_processptick_Throw_FD :
+  \<open>(P \<Theta> a \<in> A. Q a) \<down> n \<sqsubseteq>FD (P \<down> n) \<Theta> a \<in> A. (Q a \<down> n)\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+proof (unfold refine_defs, safe)
+  show \<open>t \<in> \<D> ?lhs\<close> if \<open>t \<in> \<D> ?rhs\<close> for t
+  proof -
+    from \<open>t \<in> \<D> ?rhs\<close>
+    consider t1 t2 where \<open>t = t1 @ t2\<close> \<open>t1 \<in> \<D> (P \<down> n)\<close> \<open>tF t1\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>ftF t2\<close>
+      | t1 a t2 where \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> (P \<down> n)\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<D> (Q a \<down> n)\<close>
+      unfolding D_Throw by blast
+    thus \<open>t \<in> \<D> ?lhs\<close>
+    proof cases
+      show \<open>\<lbrakk>t = t1 @ t2; t1 \<in> \<D> (P \<down> n); tF t1; set t1 \<inter> ev ` A = {}; ftF t2\<rbrakk> \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t1 t2
+        by (elim D_restriction_processptickE, simp_all add: Throw_projs D_restriction_processptick)
+          (blast, metis (no_types, lifting) front_tickFree_append inf_sup_aci(8) inf_sup_distrib2 sup_bot_right)
+    next
+      fix t1 a t2
+      assume \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> (P \<down> n)\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<D> (Q a \<down> n)\<close>
+      from \<open>t2 \<in> \<D> (Q a \<down> n)\<close> show \<open>t \<in> \<D> ?lhs\<close>
+      proof (elim D_restriction_processptickE)
+        from \<open>t1 @ [ev a] \<in> \<T> (P \<down> n)\<close> show \<open>t2 \<in> \<D> (Q a) \<Longrightarrow> length t2 \<le> n \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+        proof (elim T_restriction_processptickE)
+          from \<open>a \<in> A\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>t = t1 @ ev a # t2\<close>
+          show \<open>t2 \<in> \<D> (Q a) \<Longrightarrow> t1 @ [ev a] \<in> \<T> P \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+            by (auto simp add: D_restriction_processptick D_Throw)
+        next
+          fix u v assume \<open>t2 \<in> \<D> (Q a)\<close> \<open>length t2 \<le> n\<close> \<open>t1 @ [ev a] = u @ v\<close>
+            \<open>u \<in> \<T> P\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+          from \<open>t1 @ [ev a] = u @ v\<close> \<open>ftF v\<close> consider \<open>t1 @ [ev a] = u\<close>
+            | v' where \<open>t1 = u @ v'\<close> \<open>v = v' @ [ev a]\<close> \<open>ftF v'\<close>
+            by (cases v rule: rev_cases) (simp_all add: front_tickFree_append_iff)
+          thus \<open>t \<in> \<D> ?lhs\<close>
+          proof cases
+            from \<open>a \<in> A\<close> \<open>length u = n\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>t2 \<in> \<D> (Q a)\<close> \<open>tF u\<close> \<open>u \<in> \<T> P\<close>
+            show \<open>t1 @ [ev a] = u \<Longrightarrow> t \<in> \<D> ?lhs\<close>
+              by (simp add: D_restriction_processptick T_Throw \<open>t = t1 @ ev a # t2\<close>)
+                (metis Cons_eq_appendI D_imp_front_tickFree append_Nil append_assoc is_processT1_TR)
+          next
+            from \<open>ftF v\<close> \<open>length u = n\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>t = t1 @ ev a # t2\<close> \<open>u \<in> \<T> P\<close>
+            show \<open>t1 = u @ v' \<Longrightarrow> v = v' @ [ev a] \<Longrightarrow> ftF v' \<Longrightarrow> t \<in> \<D> ?lhs\<close> for v'
+              by (simp add: D_restriction_processptick T_Throw)
+                (metis D_imp_front_tickFree Int_assoc Un_Int_eq(3) append_assoc
+                  front_tickFree_append front_tickFree_nonempty_append_imp
+                  inf_bot_right list.distinct(1) same_append_eq set_append \<open>t \<in> \<D> ?rhs\<close>)
+          qed
+        qed
+      next
+        from \<open>t1 @ [ev a] \<in> \<T> (P \<down> n)\<close>
+        show \<open>\<lbrakk>t2 = u @ v; u \<in> \<T> (Q a); length u = n; tF u; ftF v\<rbrakk> \<Longrightarrow> t \<in> \<D> ?lhs\<close> for u v
+        proof (elim T_restriction_processptickE)
+          assume \<open>t2 = u @ v\<close> \<open>u \<in> \<T> (Q a)\<close> \<open>length u = n\<close> \<open>tF u\<close>
+            \<open>ftF v\<close> \<open>t1 @ [ev a] \<in> \<T> P\<close> \<open>length (t1 @ [ev a]) \<le> n\<close>
+          from \<open>a \<in> A\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>t1 @ [ev a] \<in> \<T> P\<close> \<open>u \<in> \<T> (Q a)\<close>
+          have \<open>t1 @ ev a # u \<in> \<T> (P \<Theta> a\<in>A. Q a)\<close> by (auto simp add: T_Throw)
+          moreover have \<open>n < length (t1 @ ev a # u)\<close> by (simp add: \<open>length u = n\<close>)
+          ultimately have \<open>t1 @ ev a # u \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+          moreover have \<open>t = (t1 @ ev a # u) @ v\<close> by (simp add: \<open>t = t1 @ ev a # t2\<close> \<open>t2 = u @ v\<close>)
+          moreover from \<open>t1 @ [ev a] \<in> \<T> P\<close> \<open>tF u\<close> append_T_imp_tickFree
+          have \<open>tF (t1 @ ev a # u)\<close> by auto
+          ultimately show \<open>t \<in> \<D> ?lhs\<close> using \<open>ftF v\<close> is_processT7 by blast
+        next
+          fix w x assume \<open>t2 = u @ v\<close> \<open>u \<in> \<T> (Q a)\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+            \<open>t1 @ [ev a] = w @ x\<close> \<open>w \<in> \<T> P\<close> \<open>length w = n\<close> \<open>tF w\<close> \<open>ftF x\<close>
+          from \<open>t1 @ [ev a] = w @ x\<close> consider \<open>t1 @ [ev a] = w\<close>
+            | x' where \<open>t1 = w @ x'\<close> \<open>x = x' @ [ev a]\<close>
+            by (cases x rule: rev_cases) simp_all
+          thus \<open>t \<in> \<D> ?lhs\<close>
+          proof cases
+            assume \<open>t1 @ [ev a] = w\<close>
+            with \<open>a \<in> A\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>u \<in> \<T> (Q a)\<close> \<open>w \<in> \<T> P\<close>
+            have \<open>t1 @ ev a # u \<in> \<T> (P \<Theta> a\<in>A. Q a)\<close> by (auto simp add: T_Throw)
+            moreover have \<open>n < length (t1 @ ev a # u)\<close> by (simp add: \<open>length u = n\<close>)
+            ultimately have \<open>t1 @ ev a # u \<in> \<D> ?lhs\<close> by (blast intro: D_restriction_processptickI)
+            moreover have \<open>t = (t1 @ ev a # u) @ v\<close>
+              by (simp add: \<open>t = t1 @ ev a # t2\<close> \<open>t2 = u @ v\<close>)
+            moreover from \<open>t1 @ [ev a] = w\<close> \<open>tF u\<close> \<open>tF w\<close> have \<open>tF (t1 @ ev a # u)\<close> by auto
+            ultimately show \<open>t \<in> \<D> ?lhs\<close> using \<open>ftF v\<close> is_processT7 by blast
+          next
+            fix x' assume \<open>t1 = w @ x'\<close> \<open>x = x' @ [ev a]\<close>
+            from \<open>set t1 \<inter> ev ` A = {}\<close> \<open>t1 = w @ x'\<close> \<open>w \<in> \<T> P\<close>
+            have \<open>w \<in> \<T> P \<and> set w \<inter> ev ` A = {}\<close> by auto
+            hence \<open>w \<in> \<T> (P \<Theta> a\<in>A. Q a)\<close> by (simp add: T_Throw)
+            with \<open>length w = n\<close> \<open>tF w\<close> have \<open>w \<in> \<D> ?lhs\<close>
+              by (blast intro: D_restriction_processptickI)
+            moreover have \<open>t = w @ x @ t2\<close>
+              by (simp add: \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] = w @ x\<close>)
+            moreover from D_imp_front_tickFree[OF \<open>t \<in> \<D> ?rhs\<close>] \<open>t = t1 @ ev a # t2\<close> \<open>t1 = w @ x'\<close>
+              front_tickFree_nonempty_append_imp that tickFree_append_iff
+            have \<open>ftF (x @ t2)\<close> by (simp add: \<open>t = w @ x @ t2\<close> front_tickFree_append_iff)
+            ultimately show \<open>t \<in> \<D> ?lhs\<close> by (simp add: \<open>tF w\<close> is_processT7)
+          qed
+        qed
+      qed
+    qed
+  qed
+
+  thus \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+    by (meson is_processT8 le_approx2 mono_Throw restriction_processptick_approx_self)
+qed
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma Throw_non_destructive :
+  \<open>non_destructive (\<lambda>(P :: ('a, 'r) processptick, Q). P \<Theta> a \<in> A. Q a)\<close>
+proof (rule order_non_destructiveI, clarify)
+  fix P P' :: \<open>('a, 'r) processptick\<close> and Q Q' :: \<open>'a \<Rightarrow> ('a, 'r) processptick\<close> and n
+  assume \<open>(P, Q) \<down> n = (P', Q') \<down> n\<close> \<open>0 < n\<close>
+  hence \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+    by (simp_all add: restriction_prod_def)
+
+  { let ?lhs = \<open>P \<Theta> a \<in> A. Q a \<down> n\<close>
+    fix t u v assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (Throw P' A Q')\<close> \<open>length u = n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    from this(2) consider \<open>u \<in> \<T> P'\<close> \<open>set u \<inter> ev ` A = {}\<close>
+      | (divL) t1 t2 where \<open>u = t1 @ t2\<close> \<open>t1 \<in> \<D> P'\<close> \<open>tF t1\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>ftF t2\<close>
+      | (traces) t1 a t2 where \<open>u = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<T> (Q' a)\<close>
+      unfolding T_Throw by blast
+    hence \<open>t \<in> \<D> ?lhs\<close>
+    proof cases
+      assume \<open>u \<in> \<T> P'\<close> \<open>set u \<inter> ev ` A = {}\<close>
+      from \<open>u \<in> \<T> P'\<close> \<open>length u = n\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>u \<in> \<T> P\<close>
+        by (metis T_restriction_processptickI dual_order.refl
+            length_le_in_T_restriction_processptick)
+      with \<open>set u \<inter> ev ` A = {}\<close> have \<open>u \<in> \<T> (Throw P A Q)\<close>
+        by (simp add: T_Throw)
+      with \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> \<open>length u = n\<close> show \<open>t \<in> \<D> ?lhs\<close>
+        by (simp add: D_restriction_processptickI is_processT7)
+    next
+      case divL
+      show \<open>t \<in> \<D> ?lhs\<close>
+      proof (cases \<open>length t1 = n\<close>)
+        assume \<open>length t1 = n\<close>
+        with \<open>P \<down> n = P' \<down> n\<close> divL(2,4) have \<open>t1 \<in> \<T> (Throw P A Q)\<close>
+          by (simp add: T_Throw)
+            (metis D_T T_restriction_processptickI dual_order.refl
+              length_le_in_T_restriction_processptick)
+        with \<open>ftF v\<close> \<open>length t1 = n\<close> \<open>t = u @ v\<close> \<open>tF u\<close>
+          divL(1,3,5) \<open>length u = n\<close> show \<open>t \<in> \<D> ?lhs\<close>
+          by (auto simp add: D_restriction_processptick is_processT7)
+      next
+        assume \<open>length t1 \<noteq> n\<close>
+        with \<open>length u = n\<close> divL(1) have \<open>length t1 < n\<close> by simp
+        with \<open>P \<down> n = P' \<down> n\<close> divL(2,3,4) have \<open>t1 \<in> \<D> P\<close>
+          by (simp add: D_Throw)
+            (metis D_restriction_processptickI length_less_in_D_restriction_processptick)
+        with \<open>ftF v\<close> \<open>t = u @ v\<close> divL(1,3,4) show \<open>t \<in> \<D> ?lhs\<close>
+          by (simp add: D_restriction_processptick T_Throw)
+            (metis \<open>length u = n\<close> \<open>tF u\<close> append.assoc divL(5))
+      qed
+    next
+      case traces
+      from \<open>length u = n\<close> traces(1)
+      have \<open>length (t1 @ [ev a]) \<le> n\<close> \<open>length t2 \<le> n\<close> by simp_all
+      from \<open>P \<down> n = P' \<down> n\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close> \<open>length (t1 @ [ev a]) \<le> n\<close>
+      have \<open>t1 @ [ev a] \<in> \<T> P\<close>
+        by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+      moreover from \<open>length t2 \<le> n\<close> \<open>t2 \<in> \<T> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+      have \<open>t2 \<in> \<T> (Q a)\<close>
+        by (metis T_restriction_processptickI restriction_fun_def
+            length_le_in_T_restriction_processptick)
+      ultimately have \<open>u \<in> \<T> (P \<Theta> a \<in> A. Q a)\<close>
+        using traces(1,3,4) unfolding T_Throw by blast
+      with \<open>ftF v\<close> \<open>length u = n\<close> \<open>t = u @ v\<close> \<open>tF u\<close>
+      show \<open>t \<in> \<D> ?lhs\<close> by (auto simp add: D_restriction_processptick)
+    qed
+  } note * = this
+
+  show \<open>P \<Theta> a \<in> A. Q a \<down> n \<sqsubseteq>FD P' \<Theta> a \<in> A. Q' a \<down> n\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+  proof (unfold refine_defs, safe)
+    show div : \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    proof (elim D_restriction_processptickE)
+      assume \<open>t \<in> \<D> (P' \<Theta> a \<in> A. Q' a)\<close> \<open>length t \<le> n\<close>
+      from this(1) consider (divL) t1 t2 where \<open>t = t1 @ t2\<close> \<open>t1 \<in> \<D> P'\<close>
+        \<open>tF t1\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>ftF t2\<close>
+      | (divR) t1 a t2 where \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<D> (Q' a)\<close>
+        unfolding D_Throw by blast
+      thus \<open>t \<in> \<D> ?lhs\<close>
+      proof cases
+        case divL
+        show \<open>t \<in> \<D> ?lhs\<close>
+        proof (cases \<open>length t1 = n\<close>)
+          assume \<open>length t1 = n\<close>
+          have \<open>t1 \<in> \<T> P'\<close> by (simp add: D_T divL(2))
+          with \<open>P \<down> n = P' \<down> n\<close> \<open>length t1 = n\<close> have \<open>t1 \<in> \<T> P\<close>
+            by (metis T_restriction_processptickI dual_order.eq_iff
+                length_le_in_T_restriction_processptick)
+          with divL(4) have \<open>t1 \<in> \<T> (Throw P A Q)\<close> by (simp add: T_Throw)
+          with \<open>length t1 = n\<close> divL(1,3,5) show \<open>t \<in> \<D> ?lhs\<close>
+            by (simp add: D_restriction_processptickI is_processT7)
+        next
+          assume \<open>length t1 \<noteq> n\<close>
+          with \<open>length t \<le> n\<close> divL(1) have \<open>length t1 < n\<close> by simp
+          with \<open>t1 \<in> \<D> P'\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>t1 \<in> \<D> P\<close>
+            by (metis D_restriction_processptickI length_less_in_D_restriction_processptick)
+          with divL(3, 4) front_tickFree_Nil have \<open>t1 \<in> \<D> (Throw P A Q)\<close>
+            by (simp (no_asm) add: D_Throw) blast
+          with divL(1,3,5) show \<open>t \<in> \<D> ?lhs\<close>
+            by (simp add: D_restriction_processptickI is_processT7)
+        qed
+      next
+        case divR
+        from \<open>length t \<le> n\<close> divR(1)
+        have \<open>length (t1 @ [ev a]) \<le> n\<close> \<open>length t2 < n\<close> by simp_all
+        from \<open>P \<down> n = P' \<down> n\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close> \<open>length (t1 @ [ev a]) \<le> n\<close>
+        have \<open>t1 @ [ev a] \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+        moreover from \<open>length t2 < n\<close> \<open>t2 \<in> \<D> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>t2 \<in> \<D> (Q a)\<close>
+          by (metis D_restriction_processptickI restriction_fun_def
+              length_less_in_D_restriction_processptick)
+        ultimately have \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q a)\<close>
+          using divR(1,3,4) unfolding D_Throw by blast
+        thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (Throw P' A Q') \<Longrightarrow> length u = n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> t \<in> \<D> ?lhs\<close> for u v by (fact "*")
+    qed
+
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+    proof (elim F_restriction_processptickE)
+      assume \<open>(t, X) \<in> \<F> (Throw P' A Q')\<close> \<open>length t \<le> n\<close>
+      from this(1) consider \<open>t \<in> \<D> (Throw P' A Q')\<close> | \<open>(t, X) \<in> \<F> P'\<close> \<open>set t \<inter> ev ` A = {}\<close>
+        | (failR) t1 a t2 where \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close>
+          \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>(t2, X) \<in> \<F> (Q' a)\<close>
+        unfolding Throw_projs by auto
+      thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+      proof cases
+        assume \<open>t \<in> \<D> (Throw P' A Q')\<close>
+        hence \<open>t \<in> \<D> ?rhs\<close> by (simp add: D_restriction_processptickI)
+        with D_F div show \<open>(t, X) \<in> \<F> ?lhs\<close> by blast
+      next
+        assume \<open>(t, X) \<in> \<F> P'\<close> \<open>set t \<inter> ev ` A = {}\<close>
+        from \<open>(t, X) \<in> \<F> P'\<close> \<open>P \<down> n = P' \<down> n\<close> \<open>length t \<le> n\<close> have \<open>t \<in> \<T> P\<close>
+          by (metis F_T T_restriction_processptickI length_le_in_T_restriction_processptick)
+        with \<open>set t \<inter> ev ` A = {}\<close> have \<open>t \<in> \<T> (Throw P A Q)\<close> by (simp add: T_Throw)
+        from F_imp_front_tickFree \<open>(t, X) \<in> \<F> P'\<close> have \<open>ftF t\<close> by blast
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+        proof (elim front_tickFreeE)
+          show \<open>(t, X) \<in> \<F> ?lhs\<close> if \<open>tF t\<close>
+          proof (cases \<open>length t = n\<close>)
+            assume \<open>length t = n\<close>
+            with \<open>t \<in> \<T> (Throw P A Q)\<close> \<open>tF t\<close> front_tickFree_charn show \<open>(t, X) \<in> \<F> ?lhs\<close>
+              by (simp add: F_restriction_processptick) blast
+          next
+            assume \<open>length t \<noteq> n\<close>
+            with \<open>length t \<le> n\<close> have \<open>length t < n\<close> by simp
+            with \<open>(t, X) \<in> \<F> P'\<close> \<open>P \<down> n = P' \<down> n\<close> have \<open>(t, X) \<in> \<F> P\<close>
+              by (simp add: F_restriction_processptick length_less_in_F_restriction_processptick)
+            with \<open>set t \<inter> ev ` A = {}\<close> have \<open>(t, X) \<in> \<F> (Throw P A Q)\<close> by (simp add: F_Throw)
+            thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+          qed
+        next
+          fix t' r assume \<open>t = t' @ [\<checkmark>(r)]\<close>
+          with \<open>t \<in> \<T> (Throw P A Q)\<close> have \<open>(t, X) \<in> \<F> (Throw P A Q)\<close>
+            by (simp add: tick_T_F)
+          thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+        qed
+      next
+        case failR
+        from \<open>length t \<le> n\<close> failR(1)
+        have \<open>length (t1 @ [ev a]) \<le> n\<close> \<open>length t2 < n\<close> by simp_all
+        from \<open>P \<down> n = P' \<down> n\<close> \<open>t1 @ [ev a] \<in> \<T> P'\<close> \<open>length (t1 @ [ev a]) \<le> n\<close>
+        have \<open>t1 @ [ev a] \<in> \<T> P\<close>
+          by (metis T_restriction_processptickI length_le_in_T_restriction_processptick)
+        moreover from \<open>length t2 < n\<close> \<open>(t2, X) \<in> \<F> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+        have \<open>(t2, X) \<in> \<F> (Q a)\<close>
+          by (metis F_restriction_processptickI restriction_fun_def
+              length_less_in_F_restriction_processptick)
+        ultimately have \<open>(t, X) \<in> \<F> (P \<Theta> a \<in> A. Q a)\<close>
+          using failR(1,3,4) unfolding F_Throw by blast
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (Throw P' A Q') \<Longrightarrow> length u = n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for u v
+        by (simp add: "*" is_processT8)
+    qed
+  qed
+qed
+
+
+
+lemma ThrowR_constructive_if_disjoint_initials :
+  \<open>constructive (\<lambda>Q :: 'a \<Rightarrow> ('a, 'r) processptick. P \<Theta> a \<in> A. Q a)\<close>
+  if \<open>A \<inter> {e. ev e \<in> P0} = {}\<close>
+proof (rule order_constructiveI)
+  fix Q Q' :: \<open>'a \<Rightarrow> ('a, 'r) processptick\<close> and n assume \<open>Q \<down> n = Q' \<down> n\<close>
+
+  { let ?lhs = \<open>Throw P A Q \<down> Suc n\<close>
+    fix t u v
+    assume \<open>t = u @ v\<close> \<open>u \<in> \<T> (Throw P A Q')\<close> \<open>length u = Suc n\<close> \<open>tF u\<close> \<open>ftF v\<close>
+    from \<open>u \<in> \<T> (Throw P A Q')\<close> consider \<open>u \<in> \<T> P\<close> \<open>set u \<inter> ev ` A = {}\<close>
+      | (divL) t1 t2 where \<open>u = t1 @ t2\<close> \<open>t1 \<in> \<D> P\<close> \<open>tF t1\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>ftF t2\<close>
+      | (traces) t1 a t2 where \<open>u = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<T> (Q' a)\<close>
+      unfolding T_Throw by blast
+    hence \<open>u \<in> \<D> ?lhs\<close>
+    proof cases
+      assume \<open>u \<in> \<T> P\<close> \<open>set u \<inter> ev ` A = {}\<close>
+      hence \<open>u \<in> \<T> (Throw P A Q)\<close> by (simp add: T_Throw)
+      with \<open>length u = Suc n\<close> \<open>tF u\<close> show \<open>u \<in> \<D> ?lhs\<close>
+        by (simp add: D_restriction_processptickI)
+    next
+      case divL
+      hence \<open>u \<in> \<D> (Throw P A Q)\<close> by (auto simp add: D_Throw)
+      thus \<open>u \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+    next
+      case traces
+      from \<open>length u = Suc n\<close> traces(1) have \<open>length t2 \<le> n\<close> by simp
+      with \<open>t2 \<in> \<T> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>t2 \<in> \<T> (Q a)\<close>
+        by (metis T_restriction_processptickI restriction_fun_def
+            length_le_in_T_restriction_processptick)
+      with traces(1-4) have \<open>u \<in> \<T> (Throw P A Q)\<close> by (auto simp add: T_Throw)
+      thus \<open>u \<in> \<D> ?lhs\<close>
+        by (simp add: D_restriction_processptickI \<open>length u = Suc n\<close> \<open>tF u\<close>)
+    qed
+    hence \<open>t \<in> \<D> ?lhs\<close> by (simp add: \<open>ftF v\<close> \<open>t = u @ v\<close> \<open>tF u\<close> is_processT7)
+  } note * = this
+
+  show \<open>(P \<Theta> a \<in> A. Q a) \<down> Suc n \<sqsubseteq>FD P \<Theta> a \<in> A. Q' a \<down> Suc n\<close> (is \<open>?lhs \<sqsubseteq>FD ?rhs\<close>)
+  proof (unfold refine_defs, safe)
+    show div : \<open>t \<in> \<D> ?rhs \<Longrightarrow> t \<in> \<D> ?lhs\<close> for t
+    proof (elim D_restriction_processptickE)
+      assume \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q' a)\<close> \<open>length t \<le> Suc n\<close>
+      from this(1) consider (divL) t1 t2 where \<open>t = t1 @ t2\<close> \<open>t1 \<in> \<D> P\<close>
+        \<open>tF t1\<close> \<open>set t1 \<inter> ev ` A = {}\<close> \<open>ftF t2\<close>
+      | (divR) t1 a t2 where \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P\<close>
+        \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>t2 \<in> \<D> (Q' a)\<close>
+        unfolding D_Throw by blast
+      thus \<open>t \<in> \<D> ?lhs\<close>
+      proof cases
+        case divL
+        hence \<open>t \<in> \<D> (P \<Theta> a \<in> A. Q a)\<close> by (auto simp add: D_Throw)
+        thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+      next
+        case divR
+        from divR(2,4) that have \<open>t1 \<noteq> []\<close>
+          by (cases t1) (auto intro: initials_memI)
+        with divR(1) \<open>length t \<le> Suc n\<close> nat_less_le have \<open>length t2 < n\<close> by force
+        with \<open>t2 \<in> \<D> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>t2 \<in> \<D> (Q a)\<close>
+          by (metis D_restriction_processptickI restriction_fun_def
+              length_less_in_D_restriction_processptick)
+        with divR(1-4) have \<open>t \<in> \<D> (Throw P A Q)\<close> by (auto simp add: D_Throw)
+        thus \<open>t \<in> \<D> ?lhs\<close> by (simp add: D_restriction_processptickI)
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (Throw P A Q') \<Longrightarrow> length u = Suc n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> t \<in> \<D> ?lhs\<close> for u v by (fact "*")
+    qed
+
+    show \<open>(t, X) \<in> \<F> ?rhs \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for t X
+    proof (elim F_restriction_processptickE)
+      assume \<open>(t, X) \<in> \<F> (Throw P A Q')\<close> \<open>length t \<le> Suc n\<close>
+      from this(1) consider \<open>t \<in> \<D> (Throw P A Q')\<close> | \<open>(t, X) \<in> \<F> P\<close> \<open>set t \<inter> ev ` A = {}\<close>
+        | (failR) t1 a t2 where \<open>t = t1 @ ev a # t2\<close> \<open>t1 @ [ev a] \<in> \<T> P\<close>
+          \<open>set t1 \<inter> ev ` A = {}\<close> \<open>a \<in> A\<close> \<open>(t2, X) \<in> \<F> (Q' a)\<close>
+        unfolding Throw_projs by auto
+      thus \<open>(t, X) \<in> \<F> ?lhs\<close>
+      proof cases
+        assume \<open>t \<in> \<D> (Throw P A Q')\<close>
+        hence \<open>t \<in> \<D> ?rhs\<close> by (simp add: D_restriction_processptickI)
+        with D_F div show \<open>(t, X) \<in> \<F> ?lhs\<close> by blast
+      next
+        assume \<open>(t, X) \<in> \<F> P\<close> \<open>set t \<inter> ev ` A = {}\<close>
+        hence \<open>(t, X) \<in> \<F> (Throw P A Q)\<close> by (simp add: F_Throw)
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+      next
+        case failR
+        from failR(2, 4) that have \<open>t1 \<noteq> []\<close>
+          by (cases t1) (auto intro: initials_memI)
+        with failR(1) \<open>length t \<le> Suc n\<close> nat_less_le have \<open>length t2 < n\<close> by force
+        with \<open>(t2, X) \<in> \<F> (Q' a)\<close> \<open>Q \<down> n = Q' \<down> n\<close> have \<open>(t2, X) \<in> \<F> (Q a)\<close>
+          by (metis F_restriction_processptickI restriction_fun_def
+              length_less_in_F_restriction_processptick)
+        with failR(1-4) have \<open>(t, X) \<in> \<F> (Throw P A Q)\<close> by (auto simp add: F_Throw)
+        thus \<open>(t, X) \<in> \<F> ?lhs\<close> by (simp add: F_restriction_processptickI)
+      qed
+    next
+      show \<open>t = u @ v \<Longrightarrow> u \<in> \<T> (Throw P A Q') \<Longrightarrow> length u = Suc n \<Longrightarrow>
+            tF u \<Longrightarrow> ftF v \<Longrightarrow> (t, X) \<in> \<F> ?lhs\<close> for u v
+        by (simp add: "*" is_processT8)
+    qed
+  qed
+qed
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/HOL-CSP_RS/document/root.tex
--- /dev/null
+++ thys/HOL-CSP_RS/document/root.tex
@@ -0,0 +1,78 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{pifont}
+
+\usepackage{mathpartir}
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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}
+
+%\pagestyle{plain}
+
+
+\begin{document}
+
+\title{HOL-CSP\_RS: CSP Semantics over Restriction Spaces}
+\author{Benoît Ballenghien \and	Burkhart Wolff}
+\maketitle
+
+\abstract{
+We use the \verb'Restriction_Spaces' library as a semantic foundation for the process algebra framework \verb'HOL-CSP',
+offering a complementary backend to the existing \verb'HOLCF' infrastructure.
+The type of processes is instantiated as a restriction space, and we prove that it is complete in this setting.
+This enables the construction of fixed points for recursive process definitions without having to
+rely exclusively on a pointed complete partial order.
+Notably, some operators are constructive without being Scott-continuous, and vice versa,
+illustrating the genuine complementarity between the two approaches.
+We also show that key CSP operators are either constructive or non-destructive,
+and verify the admissibility of several predicates, thereby supporting automated reasoning over recursive specifications.
+}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Deadlock/Deadlock.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Deadlock/Deadlock.thy
@@ -0,0 +1,2369 @@
+section \<open>Deadlock Checking\<close>
+
+theory Deadlock
+  imports
+    Timed_Automata.Timed_Automata Timed_Automata.CTL Difference_Bound_Matrices.DBM_Operations
+    Timed_Automata.Normalized_Zone_Semantics
+    Munta_Base.Normalized_Zone_Semantics_Impl
+    Munta_Base.Normalized_Zone_Semantics_Impl_Refine
+    Difference_Bound_Matrices.FW_More
+    Munta_Base.TA_Syntax_Bundles
+begin
+
+unbundle no_library_syntax
+
+subsection \<open>Notes\<close>
+
+text \<open>
+If we want to run the deadlock checker together with a reachability check for property \<open>P\<close>, we can:
+  item run a reachability check with \<open>\<not> check_deadlock\<close> first; if we find a deadlock, we are done;
+    else we can check whether any of the reachable states satisfies \<open>P\<close>;
+  item or run a reachability check with \<open>P\<close> first, which might give us earlier termination;
+    if we find that \<open>P\<close> is satisfied, can we directly report that the original formula is satisfied?
+
+Generally, it seems advantageous to compute \<open>\<not> check_deadlock\<close> on the final set of states, and
+not intermediate subsumed ones, as the operation is expensive.
+\<close>
+
+
+subsection \<open>Abstract Reachability Checking\<close>
+
+definition zone_time_pre :: "('c, ('t::time)) zone \<Rightarrow> ('c, 't) zone"
+("_\<down>" [71] 71)
+where
+  "Z\<down> = {u | u d. (u \<oplus> d) \<in> Z \<and> d \<ge> (0::'t)}"
+
+definition zone_set_pre :: "('c, 't::time) zone \<Rightarrow> 'c list \<Rightarrow> ('c, 't) zone"
+where
+  "zone_set_pre Z r = {u . ([r \<rightarrow> (0::'t)]u) \<in> Z}"
+
+definition zone_pre :: "('c, 't::time) zone \<Rightarrow> 'c list \<Rightarrow> ('c, 't) zone"
+  where
+  "zone_pre Z r = (zone_set_pre Z r)\<down>"
+
+lemma zone_time_pre_mono:
+  "A\<down> \<subseteq> B\<down>" if "A \<subseteq> B"
+  using that unfolding zone_time_pre_def by auto
+
+lemma clock_set_split:
+  "P (([r \<rightarrow> 0]u) x) \<longleftrightarrow> (x \<notin> set r \<longrightarrow> P (u x)) \<and> (x \<in> set r \<longrightarrow> P 0)"
+  by (cases "x \<in> set r") auto
+
+
+
+
+
+
+context Regions_TA
+begin
+
+definition
+  "check_deadlock l Z \<equiv> Z \<subseteq>
+    \<Union> {(zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l})\<down> | g a r l'.
+        A \<turnstile> l \<longrightarrow>g,a,r l'}"
+
+lemma V_zone_time_pre:
+  "x \<in> (Z \<inter> V)\<down>" if "x \<in> Z\<down>" "x \<in> V"
+  using that unfolding zone_time_pre_def by (auto simp: V_def cval_add_def)
+
+lemma check_deadlock_alt_def:
+  "check_deadlock l Z = (Z \<subseteq> \<Union> {
+    (zone_set_pre ({u. u \<turnstile> inv_of A l'} \<inter> V) r \<inter> {u. \<forall> x \<in> set r. u x \<ge> 0}
+       \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l})\<down> \<inter> V
+    | g a r l'. A \<turnstile> l \<longrightarrow>g,a,r l'})" (is "_ = (?L \<subseteq> ?R)") if "Z \<subseteq> V"
+proof -
+  { fix g a r l' x
+    assume t: "A \<turnstile> l \<longrightarrow>g,a,r l'"
+    assume x: "x \<in> (zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l})\<down>"
+    assume "x \<in> V"
+    let ?A = "zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l}"
+    let ?B = "zone_set_pre ({u. u \<turnstile> inv_of A l'} \<inter> V) r \<inter> {u. \<forall> x \<in> set r. u x \<ge> 0}
+              \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l}"
+    from valid_abstraction have "collect_clkvt (trans_of A) \<subseteq> X"
+      by (auto elim: valid_abstraction.cases)
+    have *: "0 \<le> u c" if "c \<in> set r" "u \<in> V" for c u
+    proof -
+      from t \<open>c \<in> set r\<close> have "c \<in> collect_clkvt (trans_of A)"
+        unfolding collect_clkvt_def by force
+      with \<open>_ \<subseteq> X\<close> have "c \<in> X"
+        by auto
+      with \<open>u \<in> V\<close> show ?thesis
+        by (auto simp: V_def)
+    qed
+    have **: "u \<in> zone_set_pre ({u. u \<turnstile> inv_of A l'} \<inter> V) r"
+      if "u \<in> zone_set_pre {u. u \<turnstile> inv_of A l'} r" "u \<in> V" for u
+      using that unfolding zone_set_pre_def V_def by (auto split: clock_set_split)
+    from x \<open>x \<in> V\<close> have "x \<in> (?A \<inter> V)\<down>"
+      by (rule V_zone_time_pre)
+    moreover have "y \<in> ?B" if "y \<in> ?A \<inter> V" for y
+      using that by (auto intro: * **)
+    ultimately have "x \<in> ?B\<down>"
+      unfolding zone_time_pre_def by auto
+  } note * = this
+  have "zone_set_pre (Z \<inter> V) r \<subseteq> zone_set_pre Z r" for Z r
+    unfolding zone_set_pre_def by auto
+  with \<open>Z \<subseteq> V\<close> show ?thesis
+    unfolding check_deadlock_def
+    apply safe
+    subgoal for x
+      apply rotate_tac
+      apply (drule subsetD, assumption)
+      apply (drule subsetD, assumption)
+      apply clarsimp
+      apply (frule *, assumption+)
+      subgoal for g a r l'
+        by (inst_existentials
+            "(zone_set_pre ({u. u \<turnstile> inv_of A l'} \<inter> V) r \<inter> {u. \<forall>x\<in>set r. 0 \<le> u x} \<inter> {u. u \<turnstile> g} \<inter>
+             {u. u \<turnstile> inv_of A l})\<down> \<inter> V") auto
+      done
+    apply (drule subsetD, assumption)+
+    apply safe
+    subgoal for x X g a r l'
+      by (drule
+          subsetD[OF zone_time_pre_mono,
+            where B1 = "zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l}",
+              rotated]; force)
+    done
+qed
+
+lemma step_trans1:
+  assumes "A \<turnstile>t \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>"
+  shows "u \<in> zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g}" "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  using assms by (auto elim!: step_trans.cases simp: zone_set_pre_def)
+
+lemma step_trans2:
+  assumes "u \<in> zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g}" "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  shows "\<exists> u'. A \<turnstile>t \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>"
+  using assms unfolding zone_set_pre_def by auto
+
+lemma time_pre_zone:
+  "u \<in> (Z \<inter> {u. u \<turnstile> inv_of A l})\<down>" if "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l', u'\<rangle>" "u' \<in> Z"
+  using that by (auto elim!: step_t.cases simp: zone_time_pre_def)
+
+lemma time_pre_zone':
+  "\<exists> d u'. u' \<in> Z \<and> A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l, u'\<rangle>" if "u \<in> (Z \<inter> {u. u \<turnstile> inv_of A l})\<down>"
+  using that unfolding zone_time_pre_def by auto
+
+lemma step_trans3:
+  assumes "A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>"
+  shows "u \<in> (zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l})\<down>"
+        "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  using assms by (auto dest: step_trans1 time_pre_zone step_delay_loc elim: step_trans'.cases)
+
+lemma step_trans4:
+  assumes "u \<in> (zone_set_pre {u. u \<turnstile> inv_of A l'} r \<inter> {u. u \<turnstile> g} \<inter> {u. u \<turnstile> inv_of A l})\<down>"
+          "A \<turnstile> l \<longrightarrow>g,a,r l'"
+    shows "\<exists> u'. A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>"
+  using assms by (fast dest: time_pre_zone' step_trans2[rotated])
+
+lemma check_deadlock_correct:
+  "check_deadlock l Z \<longleftrightarrow> (\<forall>u \<in> Z. \<exists>l' u' g a r. A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>)"
+  unfolding check_deadlock_def
+  apply safe
+  subgoal for x
+    using step_trans4 by blast
+  subgoal for x
+    using step_trans3 by fast
+  done
+
+lemma step'_step_trans'_iff:
+  "A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle> \<longleftrightarrow> (\<exists>g a r. A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow>(g,a,r) \<langle>l', u'\<rangle>)"
+  by (metis prod_cases3 step'.cases step'.intros step_a.cases step_a.simps step_trans'.cases
+            step_trans'.intros step_trans.cases step_trans.simps
+     )
+
+lemma check_deadlock_correct_step':
+  "check_deadlock l Z \<longleftrightarrow> (\<forall>u \<in> Z. \<exists>l' u'. A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  using check_deadlock_correct step'_step_trans'_iff by simp
+
+paragraph \<open>Unused\<close>
+
+lemma delay_step_zone:
+  "u' \<in> Z\<up> \<inter> {u. u \<turnstile> inv_of A l}" if "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l, u'\<rangle>" "u \<in> Z"
+  using that by (auto elim!: step_t.cases simp: zone_delay_def)
+
+lemma delay_step_zone':
+  "\<exists> d u. u \<in> Z \<and> A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l, u'\<rangle>" if "u' \<in> Z\<up> \<inter> {u. u \<turnstile> inv_of A l}"
+  using that by (auto simp: zone_delay_def)
+
+lemma delay_step_zone'':
+  "(\<exists> d u. u \<in> Z \<and> A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l, u'\<rangle>) \<longleftrightarrow> u' \<in> Z\<up> \<inter> {u. u \<turnstile> inv_of A l}"
+  using delay_step_zone delay_step_zone' by blast
+
+lemma delay_step_zone''':
+  "{u' | u' d u. u \<in> Z \<and> A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l, u'\<rangle>} = Z\<up> \<inter> {u. u \<turnstile> inv_of A l}"
+  using delay_step_zone'' by auto
+
+end (* Regions TA *)
+
+
+context Regions_TA_Start_State
+begin
+
+lemma check_deadlock_deadlocked:
+  "\<not> check_deadlock l Z \<longleftrightarrow> (\<exists>u\<in>Z. sim.sim.deadlocked (l, u))"
+  unfolding check_deadlock_correct_step' sim.sim.deadlocked_def by simp
+
+lemma deadlock_check':
+  "(\<exists>x0\<in>a0. \<exists>l u. sim.sim.reaches x0 (l, u) \<and> sim.sim.deadlocked (l, u)) \<longleftrightarrow>
+   (\<exists>l Z. reaches (l0, Z0) (l, Z) \<and> \<not> check_deadlock l Z)"
+  apply (subst ta_reaches_ex_iff)
+  subgoal for l u u' R
+    by (rule sim_complete_bisim'.P1_deadlocked_compatible[where a = "from_R l R"];
+       (rule sim_complete_bisim'.P1_P1')?) (auto intro: sim_complete_bisim'.P1_P1')
+  using check_deadlock_deadlocked by auto
+
+lemma deadlock_check:
+  "(\<exists>x0\<in>a0. sim.sim.deadlock x0) \<longleftrightarrow> (\<exists>l Z. reaches (l0, Z0) (l, Z) \<and> \<not> check_deadlock l Z)"
+  unfolding deadlock_check'[symmetric] sim.sim.deadlock_def by simp
+
+end (* Regions TA Start State *)
+
+
+subsection \<open>Operations\<close>
+
+subsubsection \<open>Subset inclusion check for federations on DBMs\<close>
+
+lemma
+  "S \<subseteq> R \<longleftrightarrow> S \<inter> -R = {}"
+  by auto
+
+lemma
+  "A \<subseteq> B \<union> C \<longleftrightarrow> A \<inter> -B \<inter> -C = {}"
+  by auto
+
+lemma
+  "(A \<union> B) \<inter> (C \<union> D) = A \<inter> C \<union> A \<inter> D \<union> B \<inter> C \<union> B \<inter> D"
+  by auto
+
+lemma Le_le_inf[intro]:
+  "Le (x :: _ :: linordered_cancel_ab_monoid_add) \<preceq> \<infinity>"
+  by (auto intro: linordered_monoid.order.strict_implies_order)
+
+lemma dbm_entry_val_mono:
+  "dbm_entry_val u a b e'" if "dbm_entry_val u a b e" "e \<le> e'"
+using that
+  by cases
+  (auto simp: DBM.less_eq intro: dbm_entry_val_mono1 dbm_entry_val_mono2 dbm_entry_val_mono3
+    | simp add: DBM.less_eq dbm_entry_val.simps dbm_le_def
+  )+
+
+definition and_entry ::
+  "nat \<Rightarrow> nat \<Rightarrow> ('t::{linordered_cancel_ab_monoid_add,uminus}) DBMEntry \<Rightarrow> 't DBM \<Rightarrow> 't DBM" where
+  "and_entry a b e M = (\<lambda>i j. if i = a \<and> j = b then min (M i j) e else M i j)"
+
+lemma and_entry_mono:
+  "and_entry a b e M i j \<le> M i j"
+  by (auto simp: and_entry_def)
+
+abbreviation "clock_to_option a \<equiv> (if a > 0 then Some a else None)"
+
+definition
+  "dbm_entry_val' u a b e \<equiv> dbm_entry_val u (clock_to_option a) (clock_to_option b) e"
+
+definition
+  "dbm_minus n xs m = concat (map (\<lambda>(i, j). map (\<lambda> M. and_entry j i (neg_dbm_entry (m i j)) M) xs)
+  [(i, j). i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m i j \<noteq> \<infinity>])"
+
+locale Default_Nat_Clock_Numbering =
+  fixes n :: nat and v :: "nat \<Rightarrow> nat"
+  assumes v_is_id: "\<forall> c. c > 0 \<and> c \<le> n \<longrightarrow> v c = c" "\<forall> c. c > n \<longrightarrow> v c = n + 1" "v 0 = n + 1"
+begin
+
+lemma v_id:
+  "v c = c" if "v c \<le> n"
+  using v_is_id that
+  apply (cases "c = 0")
+  apply (simp; fail)
+  apply (cases "c \<le> n"; auto)
+  done
+
+lemma le_n:
+  "c \<le> n" if "v c \<le> n"
+  using that v_is_id by auto
+
+lemma gt_0:
+  "c > 0" if "v c \<le> n"
+  using that v_is_id by auto
+
+lemma le_n_iff:
+  "v c \<le> n \<longleftrightarrow> c \<le> n \<and> c > 0"
+  using v_is_id by auto
+
+lemma v_0:
+  "v c = 0 \<longleftrightarrow> False"
+  using v_is_id by (cases "c > 0"; simp; cases "c > n"; simp)
+
+lemma surj_on: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+  using v_is_id by blast
+
+abbreviation zone_of ("\<lbrakk>_\<rbrakk>") where "zone_of M \<equiv> [M]v,n"
+
+abbreviation
+  "dbm_fed S \<equiv> \<Union> m \<in> S. \<lbrakk>m\<rbrakk>"
+
+abbreviation
+  "dbm_list xs \<equiv> dbm_fed (set xs)"
+
+lemma dbm_fed_singleton:
+  "dbm_fed {m} = [m]v,n"
+  by auto
+
+lemma dbm_list_single:
+  "dbm_list xs = [m]v,n" if "set xs = {m}"
+  using that by auto
+
+lemma dbm_fed_superset_fold:
+  "S \<subseteq> dbm_list xs \<longleftrightarrow> fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs S = {}"
+proof (induction xs arbitrary: S)
+  case Nil
+  then show ?case
+    by auto
+next
+  case (Cons m xs)
+  have "S \<subseteq> dbm_list (m # xs) \<longleftrightarrow> S \<inter> - ([m]v,n) \<subseteq> dbm_list xs"
+    by auto
+  moreover have "\<dots> \<longleftrightarrow> fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs (S \<inter> - ([m]v,n)) = {}"
+    by fact
+  ultimately show ?case
+    by simp
+qed
+
+lemma dbm_fed_superset_fold':
+  "dbm_fed S \<subseteq> dbm_list xs \<longleftrightarrow> dbm_fed (fold f xs S) = {}" if
+  "\<And> m S. m \<in> set xs \<Longrightarrow> dbm_fed (f m S) = dbm_fed S \<inter> - ([m]v,n)"
+proof -
+  from that have "fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs (dbm_fed S) = dbm_fed (fold f xs S)"
+  proof (induction xs arbitrary: S)
+    case Nil
+    then show ?case
+      by simp
+  next
+    case (Cons a xs)
+    from Cons.prems have
+      "dbm_fed (fold f xs (f a S)) = fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs (dbm_fed (f a S))"
+      by - (rule sym, rule Cons.IH, auto)
+    then show ?case
+      by (simp add: Cons.prems)
+  qed
+  then show ?thesis
+    by (simp add: dbm_fed_superset_fold)
+qed
+
+lemma dbm_fed_superset_fold'':
+  "dbm_list S \<subseteq> dbm_list xs \<longleftrightarrow> dbm_list (fold f xs S) = {}" if
+  "\<And> m S. m \<in> set xs \<Longrightarrow> dbm_list (f m S) = dbm_list S \<inter> - ([m]v,n)"
+proof -
+  from that have "fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs (dbm_list S) = dbm_list (fold f xs S)"
+  proof (induction xs arbitrary: S)
+    case Nil
+    then show ?case
+      by simp
+  next
+    case (Cons a xs)
+    from Cons.prems have
+      "dbm_list (fold f xs (f a S)) = fold (\<lambda>m S. S \<inter> - ([m]v,n)) xs (dbm_list (f a S))"
+      by - (rule sym, rule Cons.IH, auto)
+    then show ?case
+      by (simp add: Cons.prems)
+  qed
+  then show ?thesis
+    by (simp add: dbm_fed_superset_fold)
+qed
+
+lemma neg_inf:
+  "{u. \<not> dbm_entry_val u a b e} = {}" if "e = (\<infinity> :: _ DBMEntry)"
+  using that by auto
+
+lemma dbm_entry_val'_diff_shift:
+  "dbm_entry_val' (u \<oplus> d) c1 c2 (M c1 c2)" if "dbm_entry_val' u c1 c2 (M c1 c2)" "0 < c1" "0 < c2"
+  using that unfolding dbm_entry_val'_def cval_add_def
+  by (auto elim!: dbm_entry_val.cases intro!: dbm_entry_val.intros)
+
+lemma dbm_entry_val_iff_bounded_Le1:
+  "dbm_entry_val u (Some c1) None e \<longleftrightarrow> Le (u c1) \<le> e"
+  by (cases e) (auto simp: any_le_inf)
+
+lemma dbm_entry_val_iff_bounded_Le2:
+  "dbm_entry_val u None (Some c2) e \<longleftrightarrow> Le (- u c2) \<le> e"
+  by (cases e) (auto simp: any_le_inf)
+
+lemma dbm_entry_val_iff_bounded_Le3:
+  "dbm_entry_val u (Some c1) (Some c2) e \<longleftrightarrow> Le (u c1 - u c2) \<le> e"
+  by (cases e) (auto simp: any_le_inf)
+
+lemma dbm_entry_val'_iff_bounded:
+  "dbm_entry_val' u c1 c2 e \<longleftrightarrow> Le((if c1 > 0 then u c1 else 0) - (if c2 > 0 then u c2 else 0)) \<le> e"
+  if "c1 > 0 \<or> c2 > 0"
+  using that unfolding dbm_entry_val'_def
+  by (auto simp:
+      dbm_entry_val_iff_bounded_Le1 dbm_entry_val_iff_bounded_Le2 dbm_entry_val_iff_bounded_Le3
+     )
+
+context
+  notes [simp] = dbm_entry_val'_def
+begin
+
+lemma neg_entry':
+  "{u. \<not> dbm_entry_val' u a b e} = {u. dbm_entry_val' u b a (neg_dbm_entry e)}"
+  if "e \<noteq> (\<infinity> :: _ DBMEntry)" "a > 0 \<or> b > 0"
+  (* using that by (auto simp: neg_entry) *)
+  using that by (cases e; cases "a > 0"; cases "b > 0"; auto 4 3 simp: le_minus_iff less_minus_iff)
+
+lemma neg_unbounded:
+  "{u. \<not> dbm_entry_val' u i j e} = {}" if "e = (\<infinity> :: _ DBMEntry)"
+  using that by auto
+
+lemma and_entry_sound:
+  "u \<turnstile>v,n and_entry a b e M" if "dbm_entry_val' u a b e" "u \<turnstile>v,n M"
+  using that unfolding DBM_val_bounded_def
+  by (cases a; cases b; auto simp: le_n_iff v_is_id(1) min_def v_0 and_entry_def)
+
+lemma DBM_val_bounded_mono:
+  "u \<turnstile>v,n M" if "u \<turnstile>v,n M'" "\<forall> i \<le> n. \<forall> j \<le> n. M' i j \<le> M i j"
+  using that unfolding DBM_val_bounded_def
+  apply (safe; clarsimp simp: le_n_iff v_is_id(1) DBM.less_eq[symmetric])
+     apply force
+    apply (blast intro: dbm_entry_val_mono)+
+  done
+
+lemma and_entry_entry:
+  "dbm_entry_val' u a b e" if "u \<turnstile>v,n and_entry a b e M" "a \<le> n" "b \<le> n" "a > 0 \<or> b > 0"
+proof -
+  from that have "dbm_entry_val' u a b (min (M a b) e)"
+    unfolding DBM_val_bounded_def by (fastforce simp: le_n_iff v_is_id(1) and_entry_def)
+  then show ?thesis
+    by (auto intro: dbm_entry_val_mono)
+qed
+
+lemma and_entry_correct:
+  "[and_entry a b e M]v,n = [M]v,n \<inter> {u. dbm_entry_val' u a b e}"
+  if "a \<le> n" "b \<le> n" "a > 0 \<or> b > 0"
+  unfolding DBM_zone_repr_def using that
+  by (blast intro: and_entry_entry and_entry_sound DBM_val_bounded_mono and_entry_mono)
+
+lemma dbm_list_Int_entry_iff_map:
+  "dbm_list xs \<inter> {u. dbm_entry_val' u i j e} = dbm_list (map (\<lambda> m. and_entry i j e m) xs)"
+  if "i \<le> n" "j \<le> n" "i > 0 \<or> j > 0"
+  unfolding dbm_entry_val'_def
+  by (induction xs;
+      simp add: and_entry_correct[OF that, symmetric, unfolded dbm_entry_val'_def] Int_Un_distrib2
+     )
+
+context
+  fixes m :: "nat \<Rightarrow> nat \<Rightarrow> ('a :: {time}) DBMEntry"
+  assumes "Le 0 \<preceq> m 0 0"
+begin
+
+private lemma A:
+  "- ([m]v,n) =
+  (\<Union> (i, j) \<in> {(i, j). i > 0 \<and> j > 0 \<and> i \<le> n \<and> j \<le> n}.
+    {u. \<not> dbm_entry_val u (Some i) (Some j) (m i j)})
+  \<union> (\<Union> i \<in> {i. i > 0 \<and> i \<le> n}. {u. \<not> dbm_entry_val u (Some i) None (m i 0)})
+  \<union> (\<Union> j \<in> {i. i > 0 \<and> i \<le> n}. {u. \<not> dbm_entry_val u None (Some j) (m 0 j)})"
+  unfolding DBM_zone_repr_def
+  apply safe
+  subgoal for u
+    unfolding DBM_val_bounded_def
+    apply (intro conjI impI allI)
+    subgoal
+      by (rule \<open>Le 0 \<preceq> m 0 0\<close>)
+    subgoal for c
+      by (auto simp: le_n_iff v_is_id(1))
+    subgoal for c
+      by (auto simp: le_n_iff v_is_id(1))
+    subgoal for c1 c2
+      by (auto elim: allE[where x = c1] simp: le_n_iff v_is_id(1))
+    done
+  unfolding DBM_val_bounded_def by (simp add: le_n_iff v_is_id(1))+
+
+private lemma B:
+  "S \<inter> - ([m]v,n) =
+  (\<Union> (i, j) \<in> {(i, j). i > 0 \<and> j > 0 \<and> i \<le> n \<and> j \<le> n}.
+    S \<inter> {u. \<not> dbm_entry_val u (Some i) (Some j) (m i j)})
+  \<union> (\<Union> i \<in> {i. i > 0 \<and> i \<le> n}. S \<inter> {u. \<not> dbm_entry_val u (Some i) None (m i 0)})
+  \<union> (\<Union> j \<in> {i. i > 0 \<and> i \<le> n}. S \<inter> {u. \<not> dbm_entry_val u None (Some j) (m 0 j)})"
+  by (subst A) auto
+
+private lemma UNION_cong:
+  "(\<Union> x \<in> S. f x) = (\<Union> x \<in> T. g x)" if "S = T" "\<And> x. x \<in> T \<Longrightarrow> f x = g x"
+  by (simp add: that)
+
+private lemma 1:
+  "S \<inter> - ([m]v,n) =
+  (\<Union> (i, j) \<in> {(i, j). (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n}. S \<inter> {u. \<not> dbm_entry_val' u i j (m i j)})
+  "
+proof -
+  have *: "{(i, j). (0 < i \<or> 0 < j) \<and> i \<le> n \<and> j \<le> n}
+  = {(i, j). 0 < i \<and> 0 < j \<and> i \<le> n \<and> j \<le> n}
+  \<union> {(i, j). 0 < i \<and> 0 = j \<and> i \<le> n \<and> j \<le> n}
+  \<union> {(i, j). 0 = i \<and> 0 < j \<and> i \<le> n \<and> j \<le> n}"
+    by auto
+  show ?thesis
+    by (simp only: B UN_Un *) (intro arg_cong2[where f = "(\<union>)"] UNION_cong; force)
+qed
+
+private lemma UNION_remove:
+  "(\<Union> x \<in> S. f x) = (\<Union> x \<in> T. g x)"
+  if "T \<subseteq> S" "\<And> x. x \<in> T \<Longrightarrow> f x = g x" "\<And> x. x \<in> S - T \<Longrightarrow> f x = {}"
+  using that by fastforce
+
+private lemma 2:
+  "(\<Union>(i, j)\<in>{(i, j).(i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n}. S \<inter> {u. \<not> dbm_entry_val' u i j (m i j)})
+ = (\<Union>(i, j)\<in>{(i, j).(i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m i j \<noteq> \<infinity>}.
+    S \<inter> {u. dbm_entry_val' u j i (neg_dbm_entry (m i j))})"
+  apply (rule UNION_remove)
+    apply force
+  subgoal for x
+    by (cases x; simp add: neg_entry'[simplified])
+  by auto
+
+lemma dbm_list_subtract:
+  "dbm_list xs \<inter> - ([m]v,n) = dbm_list (dbm_minus n xs m)"
+proof -
+  have *:
+    "set [(i, j). i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m i j \<noteq> \<infinity>]
+  = {(i, j).(i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m i j \<noteq> \<infinity>}"
+    by (auto simp del: upt.upt_Suc)
+  show ?thesis
+    unfolding dbm_minus_def
+    apply (subst set_concat)
+    apply (subst set_map)
+    apply (subst *)
+    apply (subst 1, subst 2)
+    apply (subst UN_UN_flatten)
+    apply (subst UN_simps)
+    apply (rule UNION_cong[OF HOL.refl])
+    apply (simp split del: if_split split: prod.splits)
+    apply (subst dbm_list_Int_entry_iff_map[simplified])
+       apply auto
+    done
+qed
+
+end \<comment> \<open>Context for fixed DBM\<close>
+
+end \<comment> \<open>Simplifier setup\<close>
+
+lemma dbm_list_empty_check:
+  "dbm_list xs = {} \<longleftrightarrow> list_all (\<lambda>m. [m]v,n = {}) xs"
+  unfolding list_all_iff by auto
+
+lemmas dbm_list_superset_op =
+  dbm_fed_superset_fold''[OF dbm_list_subtract[symmetric], unfolded dbm_list_empty_check]
+
+end (* Trivial clock numbering *)
+
+context TA_Start_No_Ceiling
+begin
+
+sublocale dbm: Default_Nat_Clock_Numbering n v
+  by unfold_locales (auto simp: v_def)
+
+end
+
+
+subsubsection \<open>Down\<close>
+
+paragraph \<open>Auxiliary\<close>
+
+lemma dbm_entry_le_iff:
+  "Le a \<le> Le b \<longleftrightarrow> a \<le> b"
+  "Le a \<le> Lt b \<longleftrightarrow> a < b"
+  "Lt a \<le> Le b \<longleftrightarrow> a \<le> b"
+  "Lt a \<le> Lt b \<longleftrightarrow> a \<le> b"
+  "0 \<le> Le a \<longleftrightarrow> 0 \<le> a"
+  "0 \<le> Lt a \<longleftrightarrow> 0 < a"
+  "Le a \<le> 0 \<longleftrightarrow> a \<le> 0"
+  "Lt a \<le> 0 \<longleftrightarrow> a \<le> 0"
+  "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
+  "x \<le> \<infinity> \<longleftrightarrow> True"
+proof -
+  show "\<infinity> \<le> x \<longleftrightarrow> x = \<infinity>"
+    by (cases x; auto)
+qed (auto simp: any_le_inf DBM.neutral)
+
+lemma dbm_entry_lt_iff:
+  "Le a < Le b \<longleftrightarrow> a < b"
+  "Le a < Lt b \<longleftrightarrow> a < b"
+  "Lt a < Le b \<longleftrightarrow> a \<le> b"
+  "Lt a < Lt b \<longleftrightarrow> a < b"
+  "0 < Le a \<longleftrightarrow> 0 < a"
+  "0 < Lt a \<longleftrightarrow> 0 < a"
+  "Le a < 0 \<longleftrightarrow> a < 0"
+  "Lt a < 0 \<longleftrightarrow> a \<le> 0"
+  "x < \<infinity> \<longleftrightarrow> x \<noteq> \<infinity>"
+  "\<infinity> < x \<longleftrightarrow> False"
+  by (auto simp: any_le_inf DBM.neutral DBM.less)
+
+lemmas [dbm_entry_simps] = dbm_entry_le_iff(1-9) dbm_entry_lt_iff(1-9)
+
+lemma Le_le_sum_iff:
+  "Le (y :: _ :: time) \<le> e \<longleftrightarrow> 0 \<le> e + Le (- y)"
+  by (cases e) (auto simp: DBM.add dbm_entry_le_iff)
+
+lemma dense':
+  "\<exists>c\<ge>a. c \<le> b" if "a \<le> b" for a :: "_ :: time"
+  using dense \<open>a \<le> b\<close> by auto
+
+lemma aux1:
+  "- c \<le> (a :: _ :: time)" if "a \<ge> 0" "c \<ge> 0"
+  using that using dual_order.trans neg_le_0_iff_le by blast
+
+lemma dbm_entries_dense:
+  "\<exists>d\<ge>0. Le (- d) \<le> l \<and> Le (d :: _ :: time) \<le> r" if "0 \<le> l" "l \<le> r"
+  using that by (cases l; cases r; auto simp: dbm_entry_le_iff intro: aux1)
+
+lemma dbm_entries_dense'_aux:
+  "\<exists>d\<ge>0. l + Le d \<ge> 0 \<and> 0 \<le> r + Le (- d :: _ :: time)" if "l \<le> 0" "l + r \<ge> 0" "r \<ge> 0"
+proof ((cases l; cases r), goal_cases)
+  case (2 x1 x2)
+  have "\<exists>d\<ge>0. 0 \<le> x + d \<and> d < y" if "x \<le> 0" "0 < x + y" "0 < y" for x y :: 'a
+    using that by (metis add.right_inverse add_le_cancel_left leD leI linear)
+  with 2 that show ?case
+    by (auto simp: dbm_entry_le_iff DBM.add)
+next
+  case (3 x1)
+  have "\<exists>d\<ge>0. 0 \<le> x + d" if "x \<le> 0" for x :: 'a
+    using that by (metis add.right_inverse eq_iff neg_0_le_iff_le)
+  with that 3 show ?case 
+    by (auto simp: dbm_entry_le_iff DBM.add)
+next
+  case (5 a b)
+  have "\<exists>d\<ge>0. 0 < x + d \<and> d < y" if "x \<le> 0" "0 < x + y" for x y :: 'a
+    using that by (smt add.right_inverse add_less_cancel_left leD le_less_trans linear neg_0_le_iff_le time_class.dense)
+  with 5 that show ?case 
+    by (auto simp: dbm_entry_le_iff DBM.add)
+next
+  case (6 x2)
+  have "\<exists>d\<ge>0. 0 < x + d" if "x \<le> 0" for x :: 'a
+    by (metis add.inverse_neutral add_minus_cancel add_strict_increasing2 eq_iff less_le less_minus_iff non_trivial_neg not_less_iff_gr_or_eq)
+  with that 6 show ?case 
+    by (auto simp: dbm_entry_le_iff DBM.add)
+qed (use that in \<open>auto simp: dbm_entry_le_iff DBM.add\<close>)
+
+lemma dbm_entries_dense':
+  "\<exists>d\<ge>0. l + Le d \<ge> 0 \<and> 0 \<le> r + Le (- d :: _ :: time)" if "l \<le> 0" "l + r \<ge> 0"
+proof -
+  from that have "r \<ge> 0"
+    by (meson add_decreasing order_refl order_trans)
+  with that show ?thesis
+    by (rule dbm_entries_dense'_aux)
+qed
+
+lemma (in time) non_trivial_pos: "\<exists> x. x > 0"
+  by (meson leD le_less_linear neg_le_0_iff_le non_trivial_neg)
+
+lemma dbm_entries_dense_pos:
+  "\<exists>d>0. Le (d :: _ :: time) \<le> e" if "e > 0"
+  supply dbm_entry_simps[simp]
+proof (cases e)
+case (Le d)
+  with that show ?thesis
+    by auto
+next
+  case prems: (Lt d)
+  with that have "d > 0"
+    by auto
+  from dense[OF this] obtain z where "z > 0" "z < d"
+    by auto
+  then show ?thesis
+    by (auto simp: prems)
+next
+  case prems: INF
+  obtain d :: 'a where "d > 0"
+    by atomize_elim (rule non_trivial_pos)
+  then show ?thesis
+    by (auto simp: prems)
+qed
+
+lemma le_minus_iff:
+  "- x \<le> (y :: _ :: time) \<longleftrightarrow> 0 \<le> y + x"
+  by (metis add.commute add.right_inverse add_le_cancel_left)
+
+lemma lt_minus_iff:
+  "- x < (y :: _ :: time) \<longleftrightarrow> 0 < y + x"
+  by (metis add.commute add_less_cancel_right neg_eq_iff_add_eq_0)
+
+context Default_Nat_Clock_Numbering
+begin
+
+lemma DBM_val_bounded_alt_def1:
+  "u \<turnstile>v,n m \<equiv>
+     Le 0 \<preceq> m 0 0 \<and>
+     (\<forall>c. c > 0 \<and> c \<le> n \<longrightarrow>
+          dbm_entry_val u None (Some c) (m 0 c) \<and>
+          dbm_entry_val u (Some c) None (m c 0)) \<and>
+     (\<forall>c1 c2. c1 > 0 \<and> c1 \<le> n \<and> c2 > 0 \<and> c2 \<le> n \<longrightarrow> dbm_entry_val u (Some c1) (Some c2) (m c1 c2))"
+  unfolding DBM_val_bounded_def by (rule eq_reflection) (auto simp: v_id le_n_iff)
+
+lemma DBM_val_bounded_alt_def2:
+  "u \<turnstile>v,n m \<equiv>
+     Le 0 \<le> m 0 0 \<and>
+     (\<forall>c1 c2. (c1 \<noteq> 0 \<or> c2 \<noteq> 0) \<and> c1 \<le> n \<and> c2 \<le> n \<longrightarrow> dbm_entry_val' u c1 c2 (m c1 c2))"
+  unfolding DBM_val_bounded_alt_def1 dbm_entry_val'_def DBM.less_eq
+  by (rule eq_reflection; clarsimp; safe; blast)
+
+lemma DBM_val_bounded_altI:
+  assumes
+    "Le 0 \<le> m 0 0"
+    "\<And> c1 c2. (c1 \<noteq> 0 \<or> c2 \<noteq> 0) \<and> c1 \<le> n \<and> c2 \<le> n \<Longrightarrow> dbm_entry_val' u c1 c2 (m c1 c2)"
+  shows
+    "u \<in> \<lbrakk>m\<rbrakk>"
+  unfolding DBM_zone_repr_def DBM_val_bounded_alt_def2 using assms by auto
+
+lemma dbm_entry_val'_delay1:
+  "dbm_entry_val' u c1 c2 (m c1 c2)" if "dbm_entry_val' (u \<oplus> d) c1 c2 (m c1 c2)" "d \<ge> 0" "c1 > 0"
+  using that unfolding dbm_entry_val'_def
+  by (cases "m c1 c2")
+     (auto 0 2
+        dest: add_strict_increasing2 add_increasing intro!: dbm_entry_val.intros
+        simp: cval_add_def
+     )
+
+lemma dbm_entry_val'_delay2:
+  "dbm_entry_val' u (0 :: nat) c2 (m c1 c2)" if
+  "dbm_entry_val' (u \<oplus> d) c1 c2 (m c1 c2)" "d \<ge> 0"
+  "c1 > 0" "c2 > 0" "c1 \<le> n" "c2 \<le> n"
+  "\<forall> c \<le> n. c > 0 \<longrightarrow> u c \<ge> 0"
+proof -
+  have "- u c2 \<le> da"
+    if "0 \<le> d"
+       "0 < c1"
+       "0 < c2"
+       "c1 \<le> n"
+       "c2 \<le> n"
+       "\<forall>c\<le>n. 0 < c \<longrightarrow> 0 \<le> u c"
+       "m c1 c2 = Le da"
+       "u c1 - u c2 \<le> da"
+    for da :: 'a
+    using that by (auto simp: algebra_simps le_minus_iff)
+  moreover have "- u c2 < da"
+    if "0 \<le> d"
+       "0 < c1"
+       "0 < c2"
+       "c1 \<le> n"
+       "c2 \<le> n"
+       "\<forall>c\<le>n. 0 < c \<longrightarrow> 0 \<le> u c"
+       "m c1 c2 = Lt da"
+       "u c1 - u c2 < da"
+    for da :: 'a
+    using that by (auto simp: algebra_simps lt_minus_iff intro: dual_order.strict_trans2)
+  ultimately show ?thesis
+    using that  unfolding dbm_entry_val'_def
+    by (auto elim!: dbm_entry_val.cases intro!: dbm_entry_val.intros simp: cval_add_def)
+qed
+
+lemma dbm_entry_val'_nonneg_bound:
+  "dbm_entry_val' u (0 :: nat) c (Le 0)" if "u c \<ge> 0" "c > 0"
+  using that unfolding dbm_entry_val'_def by auto
+
+lemma neg_diag_empty_spec:
+  "\<lbrakk>M\<rbrakk> = {}" if "i \<le> n" "M i i < 0"
+  using that by (meson neg_diag_empty v_is_id(1))
+
+lemma in_DBM_D:
+  "dbm_entry_val' u c1 c2 (M c1 c2)" if "u \<in> \<lbrakk>M\<rbrakk>" "c1 \<noteq> 0 \<or> c2 \<noteq> 0" "c1 \<le> n" "c2 \<le> n"
+  using that unfolding zone_time_pre_def DBM_zone_repr_def DBM_val_bounded_alt_def2 by auto
+
+context
+  fixes M :: "('t::time) DBM"
+  assumes "\<lbrakk>M\<rbrakk> \<noteq> {}"
+begin
+
+lemma non_empty_diag_0_0: "M 0 0 \<ge> 0"
+  using \<open>\<lbrakk>M\<rbrakk> \<noteq> {}\<close> neg_diag_empty_spec[of 0 M] leI by auto
+
+lemma M_k_0: "M k 0 \<ge> 0" if "\<forall> u \<in> \<lbrakk>M\<rbrakk>. \<forall> c \<le> n. c > 0 \<longrightarrow> u c \<ge> 0" "k \<le> n"
+proof (cases "k = 0")
+  case True with non_empty_diag_0_0 show ?thesis
+    by auto
+next
+  case False
+  from \<open>\<lbrakk>M\<rbrakk> \<noteq> {}\<close> obtain u where "u \<in> \<lbrakk>M\<rbrakk>"
+    by auto
+  with False that(1) \<open>k \<le> n\<close> have "u k \<ge> 0"
+    by auto
+  from \<open>u \<in> _\<close> \<open>k \<noteq> 0\<close> \<open>k \<le> n\<close> have "dbm_entry_val' u k 0 (M k 0)"
+    unfolding DBM_zone_repr_def DBM_val_bounded_alt_def2 by auto
+  with \<open>k \<noteq> 0\<close> have "Le (u k) \<le> M k 0"
+    by (simp add: dbm_entry_val'_iff_bounded)
+  with \<open>u k \<ge> 0\<close> show "M k 0 \<ge> 0"
+    by (cases "M k 0") (auto simp: dbm_entry_le_iff)
+qed
+
+lemma non_empty_cycle_free:
+  "cycle_free M n"
+  using \<open>\<lbrakk>M\<rbrakk> \<noteq> {}\<close> non_empty_cycle_free v_is_id(1) by blast
+
+lemma canonical_saturated_2:
+  assumes "Le r \<le> M 0 c"
+    and "Le (- r) \<le> M c 0"
+    and "cycle_free M n"
+    and "canonical M n"
+    and "c \<le> n"
+    and "c > 0"
+  obtains u where "u \<in> \<lbrakk>M\<rbrakk>" "u c = - r"
+  using assms v_0 by (auto simp: v_is_id intro: canonical_saturated_2[of r M v c n])
+
+lemma M_0_k: "M 0 k \<le> 0"
+  if "canonical M n" "M 0 0 \<le> 0" "\<forall> u \<in> \<lbrakk>M\<rbrakk>. \<forall> c \<le> n. c > 0 \<longrightarrow> u c \<ge> 0" "k \<le> n"
+proof (cases "k = 0")
+  case True
+  with \<open>M 0 0 \<le> 0\<close> show ?thesis
+    by auto
+next
+  case False
+  show ?thesis
+  proof (rule ccontr)
+    assume "\<not> M 0 k \<le> 0"
+    then have "M 0 k > 0"
+      by auto
+    from that(3) \<open>k \<le> n\<close> have "M k 0 \<ge> 0"
+      by (rule M_k_0)
+    from \<open>M 0 k > 0\<close> obtain d where
+      "Le d \<le> M 0 k" "d > 0"
+      by (rule dbm_entries_dense_pos[elim_format]) auto
+    with \<open>M k 0 \<ge> 0\<close> have "Le (-d) \<le> M k 0"
+      by (auto simp: dbm_entry_le_iff intro: order.trans[rotated])
+    with \<open>canonical M n\<close> \<open>Le d \<le> M 0 k\<close> obtain u where
+      "u \<in> \<lbrakk>M\<rbrakk>" "u k = -d"
+      using v_0 False \<open>k \<le> n\<close>
+      by - (rule canonical_saturated_2[of d], auto simp: non_empty_cycle_free)
+    with \<open>d > 0\<close> that(3) False \<open>k \<le> n\<close> show False
+      by fastforce
+  qed
+qed
+
+end (* Fixed non-empty DBM *)
+
+end (* Default Clock Numbering *)
+
+paragraph \<open>Definition\<close>
+
+definition
+  down :: "nat \<Rightarrow> ('t::linordered_cancel_ab_monoid_add) DBM \<Rightarrow> 't DBM"
+where
+  "down n M \<equiv>
+    \<lambda> i j. if i = 0 \<and> j > 0 then Min ({Le 0} \<union> {M k j | k. 1 \<le> k \<and> k \<le> n}) else M i j"
+
+
+paragraph \<open>Correctness\<close>
+
+context Default_Nat_Clock_Numbering
+begin
+
+sublocale Alpha_defs "{1..n}" .
+
+context
+  fixes M :: "('t::time) DBM"
+begin
+
+lemma down_complete: "u \<in> \<lbrakk>down n M\<rbrakk>" if "u \<in> \<lbrakk>M\<rbrakk>\<down>" "\<forall> c \<le> n. c > 0 \<longrightarrow> u c \<ge> 0"
+proof (rule DBM_val_bounded_altI, goal_cases)
+  case 1
+  with \<open>u \<in> _\<close> show ?case
+    unfolding down_def zone_time_pre_def by (auto intro: non_empty_diag_0_0 simp: neutral[symmetric])
+next
+  case prems: (2 c1 c2)
+  then consider "c1 > 0" | "c1 = 0" "c2 > 0"
+    by auto
+  then show ?case
+  proof cases
+    case 1
+    with prems \<open>u \<in> _\<close> show ?thesis
+      unfolding zone_time_pre_def down_def by (auto intro: dbm_entry_val'_delay1 dest: in_DBM_D)
+  next
+    case 2
+    from \<open>u \<in> \<lbrakk>M\<rbrakk>\<down>\<close> obtain d where d: "0 \<le> d" "u \<oplus> d \<in> \<lbrakk>M\<rbrakk>"
+      unfolding zone_time_pre_def by auto
+    let ?e = "Min ({Le 0} \<union> {M k c2 |k. 1 \<le> k \<and> k \<le> n})"
+    have "?e \<in> {Le 0} \<union> {M k c2 |k. 1 \<le> k \<and> k \<le> n}"
+      by (intro Min_in) auto
+    then consider "?e = Le 0" | k where "?e = M k c2" "k > 0" "k \<le> n"
+      by auto
+    then show ?thesis
+      using prems that(2) d 2 unfolding down_def
+      by cases (auto intro: dbm_entry_val'_delay2 dbm_entry_val'_nonneg_bound in_DBM_D)
+  qed
+qed
+
+lemma down_sound: "u \<in> \<lbrakk>M\<rbrakk>\<down>" if "u \<in> \<lbrakk>down n M\<rbrakk>" "canonical M n"
+proof -
+  note [simp] = dbm_entry_simps and [intro] = order.trans add_right_mono
+  from \<open>u \<in> _\<close> non_empty_diag_0_0[of "down n M"] have "Le 0 \<le> M 0 0"
+    by (auto simp: down_def neutral)
+  note * = in_DBM_D[OF \<open>u \<in> _\<close>]
+  define l where "l = Min ({M 0 c + Le (u c)   | c. 0 < c \<and> c \<le> n} \<union> {Le 0})"
+    \<comment> \<open>maximum current violation of the future bounds\<close>
+  define r where "r = Min ({M c 0 + Le (- u c) | c. 0 < c \<and> c \<le> n} \<union> {\<infinity>})"
+    \<comment> \<open>slack for shifting upwards\<close>
+  have "0 \<le> l + r" "l \<le> 0"
+  proof -
+    have
+      "l \<in> {M 0 c + Le (u c)   | c. 0 < c \<and> c \<le> n} \<union> {Le 0}"
+      "r \<in> {M c 0 + Le (- u c) | c. 0 < c \<and> c \<le> n} \<union> {\<infinity>}"
+      unfolding l_def r_def by (intro Min_in; simp)+
+    from \<open>l \<in> _\<close> show "l \<le> 0"
+      unfolding l_def by (auto intro: Min_le simp: DBM.neutral)
+    from \<open>l \<in> _\<close> \<open>r \<in> _\<close> show "0 \<le> l + r"
+    proof (safe, goal_cases)
+      case prems: (1 c1 c2)
+      with \<open>u \<in> _\<close> have "Le (u c2 - u c1) \<le> M c2 c1"
+        by (auto 0 2 dest: in_DBM_D simp: dbm_entry_val'_iff_bounded down_def)
+      also from prems \<open>canonical M n\<close> have "M c2 0 + M 0 c1 \<ge> M c2 c1"
+        by auto
+      finally have "0 \<le> M c2 0 + M 0 c1 + (Le (u c1) + Le (- u c2))"
+        by (simp add: DBM.add Le_le_sum_iff)
+      then show ?case
+        by (simp add: algebra_simps)
+    next
+      case (3 c)
+      with \<open>u \<in> _\<close> have "Le (u c) \<le> M c 0"
+        by (auto 0 2 dest: in_DBM_D simp: dbm_entry_val'_iff_bounded down_def)
+      then show ?case
+        by (auto simp: DBM.add Le_le_sum_iff)
+    qed auto
+  qed
+  from dbm_entries_dense'[OF this(2,1)] obtain d where
+    "d \<ge> 0" "0 \<le> l + Le d" "0 \<le> r + Le (- d)"
+    by auto
+  have "u \<oplus> d \<in> \<lbrakk>M\<rbrakk>"
+  proof (rule DBM_val_bounded_altI, goal_cases)
+    case 1
+    from \<open>Le 0 \<le> M 0 0\<close> show ?case .
+  next
+    case (2 c1 c2)
+    with * have **: "dbm_entry_val' u c1 c2 (down n M c1 c2)"
+      by auto
+    from 2 consider
+      "c1 \<le> n" "c2 \<le> n" "c1 > 0" "c2 > 0"
+      | "c1 = 0" "c2 \<le> n" "c2 > 0" | "c2 = 0" "c1 \<le> n" "c1 > 0"
+      by auto
+    then show ?case
+    proof cases
+      case 1
+      then show ?thesis
+        using ** unfolding down_def by (auto intro: dbm_entry_val'_diff_shift)
+    next
+      case 2
+      then have "l \<le> (M 0 c2 + Le (u c2))"
+        unfolding l_def by (auto intro: Min_le)
+      with \<open>0 \<le> l + Le d\<close> have "0 \<le> M 0 c2 + Le (u c2) + Le d"
+        by auto
+      with 2 show ?thesis
+        unfolding down_def dbm_entry_val'_def
+        by (cases "M 0 c2")
+           (auto 4 3 simp: cval_add_def DBM.add algebra_simps lt_minus_iff le_minus_iff)
+    next
+      case 3
+      then have "r \<le> M c1 0 + Le (- u c1)"
+        unfolding r_def by (auto intro: Min_le)
+      with \<open>0 \<le> r + Le (- d)\<close> have "0 \<le> M c1 0 + Le (- u c1) + Le ( -d)"
+        by auto
+      with 3 ** show ?thesis
+        unfolding down_def dbm_entry_val'_def
+        by (auto elim!: dbm_entry_val.cases simp: cval_add_def algebra_simps DBM.add)
+    qed
+  qed
+  with \<open>d \<ge> 0\<close> show ?thesis
+    unfolding zone_time_pre_def cval_add_def by auto
+qed
+
+lemma down_canonical:
+  "canonical (down n M) n"
+  if assms: "canonical M n" "\<lbrakk>M\<rbrakk> \<noteq> {}" "\<forall> u \<in> \<lbrakk>M\<rbrakk>. \<forall> c \<le> n. c > 0 \<longrightarrow> u c \<ge> 0" "M 0 0 \<le> 0"
+proof -
+  from non_empty_diag_0_0[OF \<open>\<lbrakk>M\<rbrakk> \<noteq> {}\<close>] have "M 0 0 \<ge> 0" .
+  with \<open>M 0 0 \<le> 0\<close> have "M 0 0 = 0"
+    by auto
+  note M_0_k = M_0_k[OF that(2,1,4,3)] and M_k_0 = M_k_0[OF that(2,3)]
+  have Suc_0_le_iff: "Suc 0 \<le> x \<longleftrightarrow> 0 < x" for x
+    by auto
+  define S where "S j = Min ({Le 0} \<union> {M k j |k. 1 \<le> k \<and> k \<le> n})" for j
+  { fix j :: nat
+    consider (0) "S j = 0" "\<forall> i. 1 \<le> i \<and> i \<le> n \<longrightarrow> M i j \<ge> 0"
+      | (entry) i where
+        "S j = M i j" "0 < i" "i \<le> n" "M i j \<le> 0" "\<forall> k. 1 \<le> k \<and> k \<le> n \<longrightarrow> M i j \<le> M k j"
+      unfolding S_def neutral
+      using Min_in[of "{Le 0} \<union> {M k j |k. 1 \<le> k \<and> k \<le> n}"]
+      using Min_le[of "{Le 0} \<union> {M k j |k. 1 \<le> k \<and> k \<le> n}"]
+      by (simp, safe) auto
+  } note S_cases = this
+  show ?thesis
+    apply (intro allI impI; elim conjE)
+    unfolding down_def S_def[symmetric]
+    apply clarsimp
+    apply safe
+    subgoal premises prems for i j k
+      using \<open>M 0 0 \<ge> 0\<close> by (blast intro: add_increasing)
+    subgoal for i j k
+      using \<open>canonical M n\<close>
+      by (cases rule: S_cases[of k]; cases rule: S_cases[of j])
+         (auto intro: order.trans simp: Suc_0_le_iff)
+    subgoal premises prems for i j k
+    proof -
+      from prems have "M 0 k \<le> S k"
+      apply (cases rule: S_cases[of k])
+      subgoal
+        using M_0_k[of k] by auto
+      subgoal for i'
+        using M_0_k \<open>canonical M n\<close> by (metis add.left_neutral add_right_mono dual_order.trans)
+      done
+    from \<open>canonical M n\<close> prems have "M i k \<le> M i 0 + M 0 k"
+      by auto
+    also from \<open>_ \<le> S k\<close> have "\<dots> \<le> M i 0 + S k"
+      by (simp add: add_left_mono)
+    finally show ?thesis .
+  qed
+  subgoal for i j k
+    using \<open>canonical M n\<close> \<open>M 0 0 \<le> 0\<close>
+    by (smt M_k_0 S_cases add_increasing Orderings.order.trans)
+  apply (use \<open>canonical M n\<close> in simp; fail)+
+  subgoal for i j k
+    using \<open>canonical M n\<close> \<open>M 0 0 \<le> 0\<close>
+    by (smt M_k_0 S_cases add_increasing Orderings.order.trans)
+   apply (use \<open>canonical M n\<close> in simp_all)
+  done
+qed
+
+end (* Fixed DBM *)
+
+end (* Default Clock Numbering *)
+
+subsubsection \<open>Free\<close>
+
+paragraph \<open>Definition\<close>
+
+definition
+  free :: "nat \<Rightarrow> ('t::linordered_cancel_ab_monoid_add) DBM \<Rightarrow> nat \<Rightarrow> 't DBM"
+where
+  "free n M x \<equiv>
+    \<lambda> i j. if i = x \<and> j \<noteq> x then \<infinity> else if i \<noteq> x \<and> j = x then M i 0 else M i j"
+
+definition repair_pair where
+  "repair_pair n M a b = FWI (FWI M n b) n a"
+
+definition
+  "and_entry_repair n a b e M \<equiv> repair_pair n (and_entry a b e M) a b"
+
+definition
+  "restrict_zero n M x \<equiv>
+    let
+      M1 = and_entry x 0 (Le 0) M;
+      M2 = and_entry 0 x (Le 0) M1
+    in repair_pair n M2 x 0"
+
+definition
+  "pre_reset n M x \<equiv> free n (restrict_zero n M x) x"
+
+definition
+  "pre_reset_list n M r \<equiv> fold (\<lambda> x M. pre_reset n M x) r M"
+
+paragraph \<open>Auxiliary\<close>
+
+lemma repair_pair_characteristic:
+  assumes "canonical_subs n I M"
+    and "I \<subseteq> {0..n}"
+    and "a \<le> n" "b \<le> n"
+  shows "canonical_subs n (I \<union> {a,b}) (repair_pair n M a b) \<or> (\<exists>i\<le>n. repair_pair n M a b i i < 0)"
+proof -
+  from fwi_characteristic[OF assms(1,2,4)] have
+    "canonical_subs n (I \<union> {b}) (FWI M n b) \<or> (\<exists>i\<le>n. FWI M n b i i < 0)"
+    by auto
+  then show ?thesis
+  proof
+    assume "canonical_subs n (I \<union> {b}) (FWI M n b)"
+    from fwi_characteristic[OF this _ \<open>a \<le> n\<close>] assms(2) \<open>b \<le> n\<close> show ?thesis
+      unfolding repair_pair_def by simp
+  next
+    assume "\<exists>i\<le>n. FWI M n b i i < 0"
+    then have "\<exists>i\<le>n. repair_pair n M a b i i < 0"
+      unfolding repair_pair_def
+      apply safe
+      subgoal for i
+        apply (inst_existentials i)
+         apply assumption
+        apply (frule FWI_mono[where M = "FWI M n b" and k = a])
+         apply auto
+        done
+      done
+    then show ?thesis ..
+  qed
+qed
+
+lemma repair_pair_mono:
+  assumes "i \<le> n"
+      and "j \<le> n"
+    shows "repair_pair n M a b i j \<le> M i j"
+  unfolding repair_pair_def by (auto intro: FWI_mono assms order.trans)
+
+context Default_Nat_Clock_Numbering
+begin
+
+lemmas FWI_zone_equiv = FWI_zone_equiv[OF surj_on, symmetric]
+
+lemma repair_pair_zone_equiv:
+  "\<lbrakk>repair_pair n M a b\<rbrakk> = \<lbrakk>M\<rbrakk>" if "a \<le> n" "b \<le> n"
+  using that unfolding repair_pair_def by (simp add: FWI_zone_equiv)
+
+context
+  fixes c1 c2 c x :: nat
+  notes [simp] = dbm_entry_val'_iff_bounded dbm_entry_simps DBM.add algebra_simps
+begin
+
+lemma dbm_entry_val'_diag_iff: "dbm_entry_val' u c c e \<longleftrightarrow> e \<ge> 0" if "c > 0"
+  using that by (cases e) auto
+
+lemma dbm_entry_val'_inf: "dbm_entry_val' u c1 c2 \<infinity> \<longleftrightarrow> True"
+  unfolding dbm_entry_val'_def by auto
+
+lemma dbm_entry_val'_reset_1:
+  "dbm_entry_val' (u(x := d)) x c e \<longleftrightarrow> dbm_entry_val' u 0 c (e + Le (-d))"
+  if "d \<ge> 0" "c \<noteq> x" "c > 0" "x > 0"
+  using that \<open>d \<ge> 0\<close> by (cases e) auto
+
+lemma dbm_entry_val'_reset_2:
+  "dbm_entry_val' (u(x := d)) c x e \<longleftrightarrow> dbm_entry_val' u c (0 :: nat) (e + Le d)"
+  if "d \<ge> 0" "c \<noteq> x" "c > 0" "x > 0"
+  using that \<open>d \<ge> 0\<close> by (cases e) auto
+
+lemma dbm_entry_val'_reset_2':
+  "dbm_entry_val' (u(x := d)) 0 x e \<longleftrightarrow> Le (- d) \<le> e" if "d \<ge> 0" "x > 0"
+  using that \<open>d \<ge> 0\<close> by (cases e) auto
+
+lemma dbm_entry_val'_reset_3:
+  "dbm_entry_val' (u(x := d)) c1 c2 e \<longleftrightarrow> dbm_entry_val' u c1 c2 e" if "c1 \<noteq> x" "c2 \<noteq> x" for e
+  using that unfolding dbm_entry_val'_def by (cases e) auto
+
+end (* Simplifier setup *)
+
+
+paragraph \<open>Correctness\<close>
+
+context
+  fixes M :: "('t::time) DBM"
+begin
+
+lemma free_complete: "u(x := d) \<in> \<lbrakk>free n M x\<rbrakk>"
+  if assms: "u \<in> \<lbrakk>M\<rbrakk>" "d \<ge> 0" "x > 0" "\<forall>c \<le> n. M c c \<ge> 0"
+proof (rule DBM_val_bounded_altI, goal_cases)
+  case 1
+  with \<open>_ \<in> \<lbrakk>M\<rbrakk>\<close> show ?case
+    unfolding free_def by (auto simp: neutral[symmetric] intro: non_empty_diag_0_0)
+next
+  case prems: (2 c1 c2)
+  then have "c1 \<le> n" "c2 \<le> n"
+    by auto
+  note [simp] = dbm_entry_simps
+  have *: "Le (u c1) \<le> M c1 0 + Le d" if "c1 > 0"
+  proof -
+    from \<open>_ \<in> \<lbrakk>M\<rbrakk>\<close> \<open>c1 > 0\<close> \<open>c1 \<le> n\<close> have "Le (u c1) \<le> M c1 0"
+      by (auto 0 2 simp: dbm_entry_val'_iff_bounded dest: in_DBM_D)
+    with \<open>d \<ge> 0\<close> show ?thesis
+      by (simp add: algebra_simps add_increasing)
+  qed
+  have "dbm_entry_val' (u(x := d)) c1 x (M c1 0)" if "c1 \<noteq> x"
+  proof (cases "c1 = 0")
+    case True
+    with that show ?thesis
+      using assms(4) \<open>d \<ge> 0\<close> by (auto intro: order.trans[rotated] simp: dbm_entry_val'_reset_2')
+  next
+    case False
+    with that \<open>x > 0\<close> show ?thesis
+      by (subst dbm_entry_val'_reset_2[OF \<open>d \<ge> 0\<close>]) (auto simp: dbm_entry_val'_iff_bounded *)
+  qed
+  with prems in_DBM_D[OF \<open>_ \<in> \<lbrakk>M\<rbrakk>\<close>] that(4) show ?case
+    by (auto simp: free_def dbm_entry_val'_diag_iff dbm_entry_val'_inf dbm_entry_val'_reset_3)
+qed
+
+lemma free_sound: "\<exists>d \<ge> 0. u(x := d) \<in> \<lbrakk>M\<rbrakk>" "u x \<ge> 0"
+  if assms: "u \<in> \<lbrakk>free n M x\<rbrakk>" "x > 0" "x \<le> n" "canonical M n" "M 0 x \<le> 0" "M 0 0 \<le> 0"
+proof -
+  define l where "l = Min ({M c x + Le (- u c) | c. 0 < c \<and> c \<le> n \<and> c \<noteq> x} \<union> {M 0 x})"
+  define r where "r = Min ({M x c + Le (u c)   | c. 0 < c \<and> c \<le> n \<and> c \<noteq> x} \<union> {M x 0})"
+  from non_empty_diag_0_0 \<open>u \<in> _\<close> \<open>x > 0\<close> have "0 \<le> M 0 0"
+    unfolding free_def by fastforce
+  note [simp]  = dbm_entry_simps and [intro] = order.trans add_right_mono
+  have "0 \<le> l + r" "l \<le> 0"
+  proof -
+    have
+      "l \<in> {M c x + Le (- u c)   | c. 0 < c \<and> c \<le> n \<and> c \<noteq> x} \<union> {M 0 x}"
+      "r \<in> {M x c + Le (u c)     | c. 0 < c \<and> c \<le> n \<and> c \<noteq> x} \<union> {M x 0}"
+      unfolding l_def r_def by (intro Min_in; simp)+
+    from \<open>l \<in> _\<close> \<open>M 0 x \<le> 0\<close> show "l \<le> 0"
+      unfolding l_def by - (rule order.trans[rotated], auto intro: Min_le simp: DBM.neutral)
+    from \<open>l \<in> _\<close> \<open>r \<in> _\<close> show "0 \<le> l + r"
+    proof (safe, goal_cases)
+      case prems: (1 c1 c2)
+      with \<open>canonical M n\<close> \<open>x \<le> n\<close> have "M c1 x + M x c2 \<ge> M c1 c2"
+        by auto
+      from prems \<open>u \<in> _\<close> have "Le (u c1 - u c2) \<le> M c1 c2"
+        unfolding free_def by (auto 0 2 simp: dbm_entry_val'_iff_bounded dest: in_DBM_D)
+      with \<open>M c1 c2 \<le> M c1 x + M x c2\<close> have "Le (u c1 - u c2) \<le> M c1 x + M x c2"
+        by auto
+      then have "0 \<le> M c1 x + M x c2 + (Le (u c2) + Le (- u c1))"
+        by (simp add: DBM.add Le_le_sum_iff)
+      then show ?case
+        by (simp add: algebra_simps)
+    next
+      case prems: (2 c)
+      from prems \<open>u \<in> _\<close> have "Le (u c) \<le> M c 0"
+        unfolding free_def by (auto 0 2 simp: dbm_entry_val'_iff_bounded dest: in_DBM_D)
+      also from prems \<open>canonical M n\<close> \<open>x \<le> n\<close> have "\<dots> \<le> M c x + M x 0"
+        by auto
+      finally show ?case
+        by (simp add: algebra_simps Le_le_sum_iff)
+    next
+      case prems: (3 c)
+      with \<open>u \<in> _\<close> \<open>x > 0\<close> have "Le (- u c) \<le> M 0 c"
+        unfolding free_def by (auto simp: dbm_entry_val'_iff_bounded dest: in_DBM_D[of _ _ 0 c])
+      also from prems \<open>canonical M n\<close> \<open>x \<le> n\<close> have "\<dots> \<le> M 0 x + M x c"
+        by auto
+      finally show ?case
+        by (simp add: algebra_simps Le_le_sum_iff)
+    next
+      case 4
+      from \<open>0 \<le> M 0 0\<close> \<open>canonical M n\<close> \<open>x \<le> n\<close> show ?case
+        by auto
+    qed
+  qed
+  from dbm_entries_dense'[OF this(2,1)] obtain d where
+    "d \<ge> 0" "0 \<le> l + Le d" "0 \<le> r + Le (- d)"
+    by auto
+  have "u(x := d) \<in> \<lbrakk>M\<rbrakk>"
+  proof (rule DBM_val_bounded_altI, goal_cases)
+    case 1
+    from \<open>0 \<le> M 0 0\<close> show ?case unfolding DBM.neutral .
+  next
+    case prems: (2 c1 c2)
+    then have **: "dbm_entry_val' u c1 c2 (free n M x c1 c2)"
+      by (auto intro: in_DBM_D[OF \<open>u \<in> _\<close>])
+    show ?case
+    proof -
+      have "Le (d - u c2) \<le> M x c2"
+        if "0 < x" "c1 = x" "c2 \<noteq> x" "x \<le> n" "c2 \<le> n" "0 < c2"
+      proof -
+        from that have "r \<le> M x c2 + Le (u c2)"
+          unfolding r_def by (intro Min_le) auto
+        with \<open>0 \<le> r + _\<close> have "0 \<le> M x c2 + Le (u c2) + Le (- d)"
+          by auto
+        moreover have "Le (d - u c2) \<le> M x c2 \<longleftrightarrow> 0 \<le> M x c2 + Le (u c2) + Le (- d)"
+          by (cases "M x c2") (auto simp: DBM.add algebra_simps)
+        ultimately show "Le (d - u c2) \<le> M x c2"
+          by simp
+      qed
+      moreover have "Le d \<le> M x 0"
+        if "c1 = x" "0 < x" "x \<le> n" "c2 = 0"
+      proof -
+        from \<open>0 \<le> r + _\<close> have "Le d \<le> r"
+          by (simp add: Le_le_sum_iff)
+        also have "r \<le> M x 0"
+          unfolding r_def by auto
+        finally show "Le d \<le> M x 0" .
+      qed
+      moreover have "Le (u c1 - d) \<le> M c1 x"
+        if "0 < x" "c1 \<le> n" "x \<le> n" "c1 \<noteq> x" "c2 = x" "0 < c1" "Le (u c1 - u x) \<le> M c1 0"
+      proof -
+        from that have "l \<le> M c1 x + Le (- u c1)"
+          unfolding l_def by (intro Min_le) auto
+        with \<open>0 \<le> l + Le d\<close> have "0 \<le> M c1 x + Le (- u c1) + Le d"
+          by auto
+        moreover have "Le (u c1 - d) \<le> M c1 x \<longleftrightarrow> 0 \<le> M c1 x + Le (- u c1) + Le d"
+          by (cases "M c1 x") (auto simp: DBM.add algebra_simps)
+        ultimately show "Le (u c1 - d) \<le> M c1 x"
+          by simp
+      qed
+      moreover have "Le (- d) \<le> M 0 x"
+        if "x \<le> n" "0 < x" "c2 = x" "c1 = 0" "Le (- u x) \<le> M 0 0"
+      proof -
+        from \<open>0 \<le> l + Le d\<close> have "Le (- d) \<le> l"
+          by (simp add: Le_le_sum_iff)
+        also have "l \<le> M 0 x"
+          unfolding l_def by auto
+        finally show "Le (- d) \<le> M 0 x" .
+      qed
+      ultimately show ?thesis
+        using prems \<open>x > 0\<close> **
+        by (auto simp: dbm_entry_val'_iff_bounded free_def split: if_split_asm)
+    qed
+  qed
+  with \<open>d \<ge> 0\<close> show "\<exists>d\<ge>0. u(x := d) \<in> \<lbrakk>M\<rbrakk>"
+    by auto
+  from \<open>x > 0\<close> \<open>x \<le> n\<close> have "dbm_entry_val' u 0 x (free n M x 0 x)"
+    by (auto intro: in_DBM_D[OF \<open>u \<in> _\<close>])
+  with \<open>0 < x\<close> have "Le (- u x) \<le> M 0 0"
+    by (auto simp: free_def dbm_entry_val'_iff_bounded)
+  with \<open>M 0 0 \<le> 0\<close> have "Le (- u x) \<le> 0"
+    by blast
+  then show "0 \<le> u x"
+    by auto
+qed
+
+lemma free_correct:
+  "\<lbrakk>free n M x\<rbrakk> = {u(x := d) | u d. u \<in> \<lbrakk>M\<rbrakk> \<and> d \<ge> 0}"
+  if "x > 0" "x \<le> n" "\<forall>c \<le> n. M c c \<ge> 0" "\<forall> u \<in> \<lbrakk>M\<rbrakk>. u x \<ge> 0" "canonical M n"
+    "M 0 x \<le> 0" "M 0 0 \<le> 0"
+  using that
+  apply safe
+  subgoal for u'
+    apply (frule free_sound, assumption+)
+    apply (frule free_sound(2), assumption+)
+    apply (erule exE)
+    subgoal for d
+      by (inst_existentials "u'(x := d)" "u' x"; simp)
+    done
+  subgoal for u d
+    by (auto intro: free_complete)
+  done
+
+lemma pre_reset_correct_aux:
+  "{u. (u(x := (0::'t))) \<in> \<lbrakk>M\<rbrakk>} \<inter> {u. u x \<ge> 0} = {u(x := d) | u d. u \<in> \<lbrakk>M\<rbrakk> \<and> u x = 0 \<and> d \<ge> 0}"
+  apply safe
+  subgoal for u
+    by (inst_existentials "u(x := (0::'t))" "u x") auto
+  apply clarsimp
+  subgoal for u d
+    by (subgoal_tac "u = u(x := 0)") auto
+  by auto
+
+lemma restrict_zero_correct:
+  "\<lbrakk>restrict_zero n M x\<rbrakk> = {u. u \<in> \<lbrakk>M\<rbrakk> \<and> u x = 0}" if "0 < x" "x \<le> n"
+  using that unfolding restrict_zero_def
+  by (auto simp: repair_pair_zone_equiv and_entry_correct dbm_entry_val'_iff_bounded
+      dbm_entry_simps)
+
+lemma restrict_zero_canonical:
+  "canonical (restrict_zero n M x) n \<or> check_diag n (uncurry (restrict_zero n M x))"
+  if "canonical M n" "x \<le> n"
+proof -
+  from \<open>x \<le> n\<close> have *: "{0..n} - {0, x} \<union> {x, 0} = {0..n}"
+    by auto
+  define M1 and M2 where "M1 = and_entry x 0 (Le 0) M" "M2 = and_entry 0 x (Le 0) M1"
+  from \<open>canonical M n\<close> have "canonical_subs n {0..n} M"
+    unfolding canonical_alt_def .
+  with * have "canonical_subs n ({0..n} - {0, x}) M2"
+    unfolding and_entry_def M1_M2_def canonical_subs_def by (auto simp: min.coboundedI1)
+  from repair_pair_characteristic[OF this, of x 0] \<open>x \<le> n\<close> have
+    "canonical (repair_pair n M2 x 0) n \<or> check_diag n (uncurry (repair_pair n M2 x 0))"
+    unfolding canonical_alt_def check_diag_def * neutral by auto
+  then show ?thesis
+    unfolding restrict_zero_def M1_M2_def Let_def .
+qed
+
+end (* Fixed DBM *)
+
+subsection \<open>Structural Properties\<close>
+
+lemma free_canonical:
+  "canonical (free n M x) n" if "canonical M n" "M x x \<ge> 0"
+  unfolding free_def using that by (auto simp: add_increasing2 any_le_inf)
+
+lemma free_diag:
+  "free n M x i i = M i i"
+  unfolding free_def by auto
+
+lemma check_diag_free:
+  "check_diag n (uncurry (free n M x))" if "check_diag n (uncurry M)"
+  using that unfolding check_diag_def by (auto simp: free_diag)
+
+lemma
+  "\<forall>i\<le>n. (free n M x) i i \<le> 0" if "\<forall>i\<le>n. M i i \<le> 0"
+  using that by (auto simp: free_diag)
+
+lemma canonical_nonneg_diag_non_empty:
+  assumes "canonical M n" "\<forall>i\<le>n. 0 \<le> M i i"
+  shows "[M]v,n \<noteq> {}"
+  using v_0 by (intro canonical_nonneg_diag_non_empty[OF assms]) force
+
+lemma V_structuralI:
+  "\<lbrakk>M\<rbrakk> \<subseteq> V" if "\<forall> i \<le> n. i > 0 \<longrightarrow> M 0 i \<le> 0"
+  using that
+  unfolding V_def
+proof safe
+  fix u i assume "u \<in> \<lbrakk>M\<rbrakk>" "i \<in> {1..n}"
+  then have "Suc 0 \<le> i" "i \<le> n" by simp+
+  from in_DBM_D[OF \<open>u \<in> _\<close>, of 0 i] \<open>_ \<le> i\<close> \<open>i \<le> n\<close> have "dbm_entry_val' u 0 i (M 0 i)"
+    by auto
+  with \<open>_ \<le> i\<close> \<open>i \<le> n\<close> have "Le (- u i) \<le> M 0 i"
+    by (auto simp: dbm_entry_val'_iff_bounded)
+  also from that \<open>_ \<le> i\<close> \<open>i \<le> n\<close> have "\<dots> \<le> 0"
+    by simp
+  finally show "0 \<le> u i"
+    by (auto simp: dbm_entry_simps)
+qed
+
+lemma canonical_V_non_empty_iff:
+  assumes "canonical M n" "M 0 0 \<le> 0"
+  shows "\<lbrakk>M\<rbrakk> \<subseteq> V \<and> \<lbrakk>M\<rbrakk> \<noteq> {} \<longleftrightarrow> (\<forall> i \<le> n. i > 0 \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall> i \<le> n. M i i \<ge> 0)"
+proof (safe, goal_cases)
+  case (1 u i)
+  with \<open>M 0 0 \<le> 0\<close> show ?case
+    unfolding V_def by - (rule M_0_k[OF _ \<open>canonical M n\<close>], auto)
+next
+  case (2 x i)
+  then show ?case
+    using neg_diag_empty_spec[of i M] by fastforce
+next
+  case prems: (3 u)
+  then show ?case
+    by (auto dest: subsetD[OF V_structuralI])
+next
+  case 4
+  with canonical_nonneg_diag_non_empty[OF \<open>canonical M n\<close>] show ?case
+    by simp
+qed
+
+lemma
+  assumes "(\<forall> i \<le> n. i > 0 \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall> i \<le> n. M i i \<ge> 0)" "M 0 0 \<le> 0" "x > 0"
+  shows "(\<forall> i \<le> n. i > 0 \<longrightarrow> free n M x 0 i \<le> 0) \<and> (\<forall> i \<le> n. free n M x i i \<ge> 0)"
+  using assms by (auto simp: free_def)
+
+lemma
+  assumes "(\<forall> i \<le> n. i > 0 \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall> i \<le> n. M i i \<ge> 0) \<Longrightarrow>
+           (\<forall> i \<le> n. i > 0 \<longrightarrow> f M 0 i \<le> 0) \<and> (\<forall> i \<le> n. f M i i \<ge> 0)"
+  assumes "canonical M n" "canonical (f M) n"
+  assumes "M 0 0 \<le> 0" "f M 0 0 \<le> 0"
+  assumes check_diag: "check_diag n (uncurry M) \<Longrightarrow> check_diag n (uncurry (f M))"
+  assumes "\<lbrakk>M\<rbrakk> \<subseteq> V"
+  shows "\<lbrakk>f M\<rbrakk> \<subseteq> V"
+proof (cases "\<lbrakk>M\<rbrakk> = {}")
+  case True
+  then have "check_diag n (uncurry M)"
+  using canonical_nonneg_diag_non_empty[OF \<open>canonical M n\<close>] by (force simp: neutral check_diag_def)
+  then have "check_diag n (uncurry (f M))"
+    by (rule check_diag)
+  then have "\<lbrakk>f M\<rbrakk> = {}"
+    by (auto dest: neg_diag_empty_spec simp: check_diag_def neutral)
+  then show ?thesis
+    by auto
+next
+  case False
+  with \<open>\<lbrakk>M\<rbrakk> \<subseteq> V\<close> canonical_V_non_empty_iff[OF \<open>canonical M n\<close> \<open>M 0 0 \<le> 0\<close>] have
+    "(\<forall>i\<le>n. 0 < i \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall>i\<le>n. 0 \<le> M i i)"
+    by auto
+  then have "(\<forall> i \<le> n. i > 0 \<longrightarrow> f M 0 i \<le> 0) \<and> (\<forall> i \<le> n. f M i i \<ge> 0)"
+    by (rule assms(1))
+  with \<open>canonical (f M) n\<close> have "\<lbrakk>f M\<rbrakk> \<subseteq> V \<and> \<lbrakk>f M\<rbrakk> \<noteq> {}"
+    using \<open>f M 0 0 \<le> 0\<close> by (subst canonical_V_non_empty_iff) (auto simp: free_diag)
+  then show ?thesis ..
+qed
+
+lemma
+  "\<lbrakk>free n M x\<rbrakk> \<subseteq> V" if assms: "x > 0" "canonical M n" "M 0 0 \<le> 0" "0 \<le> M x x" "\<lbrakk>M\<rbrakk> \<subseteq> V"
+proof (cases "\<lbrakk>M\<rbrakk> = {}")
+  case True
+  then obtain i where "M i i < 0" "i \<le> n"
+    using canonical_nonneg_diag_non_empty[OF \<open>canonical M n\<close>] by atomize_elim force
+  then have "free n M x i i < 0"
+    by (auto simp: free_diag)
+  with \<open>i \<le> n\<close> have "\<lbrakk>free n M x\<rbrakk> = {}"
+    by (intro neg_diag_empty_spec)
+  then show ?thesis
+    by auto
+next
+  case False
+  with \<open>\<lbrakk>M\<rbrakk> \<subseteq> V\<close> canonical_V_non_empty_iff[OF that(2,3)] have
+    "(\<forall>i\<le>n. 0 < i \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall>i\<le>n. 0 \<le> M i i)"
+    by auto
+  with that have "(\<forall> i \<le> n. i > 0 \<longrightarrow> free n M x 0 i \<le> 0) \<and> (\<forall> i \<le> n. free n M x i i \<ge> 0)"
+    by (auto simp: free_def)
+  moreover have "canonical (free n M x) n"
+    apply (rule free_canonical)
+     apply fact
+    apply fact
+    done
+  ultimately have "\<lbrakk>free n M x\<rbrakk> \<subseteq> V \<and> \<lbrakk>free n M x\<rbrakk> \<noteq> {}"
+    using \<open>M 0 0 \<le> 0\<close> by (subst canonical_V_non_empty_iff) (auto simp: free_diag)
+  then show ?thesis ..
+qed
+
+lemma
+  "down n M i i = M i i"
+  unfolding down_def by auto
+
+lemma
+  assumes "(\<forall> i \<le> n. i > 0 \<longrightarrow> M 0 i \<le> 0) \<and> (\<forall> i \<le> n. M i i \<ge> 0)" "M 0 0 \<le> 0" "x > 0"
+  shows "(\<forall> i \<le> n. i > 0 \<longrightarrow> down n M 0 i \<le> 0) \<and> (\<forall> i \<le> n. down n M i i \<ge> 0)"
+  using assms by (auto simp: down_def neutral)
+
+lemma check_diag_empty:
+  "\<lbrakk>M\<rbrakk> = {}" if "check_diag n (uncurry M)"
+  using check_diag_empty[of n v "uncurry M"] that v_is_id by auto
+
+lemma restrict_zero_mono:
+  "restrict_zero n M x i j \<le> M i j" if "i \<le> n" "j \<le> n"
+  unfolding restrict_zero_def
+  by simp (rule \<open>i \<le> n\<close> \<open>j \<le> n\<close> repair_pair_mono and_entry_mono order.trans)+
+
+lemma restrict_zero_diag:
+  "check_diag n (uncurry (restrict_zero n M x))" if "check_diag n (uncurry M)"
+  using that unfolding check_diag_def neutral[symmetric]
+  by (elim exE conjE) (frule restrict_zero_mono[where M = M and x = x], auto)
+
+lemma pre_reset_correct:
+  "\<lbrakk>pre_reset n M x\<rbrakk> = {u. (u(x := (0::'t::time))) \<in> \<lbrakk>M\<rbrakk>} \<inter> {u. u x \<ge> 0}"
+  if "x > 0" "x \<le> n" "canonical M n \<or> check_diag n (uncurry M)" "M 0 x \<le> 0" "M 0 0 \<le> 0"
+proof -
+  have check_diag: ?thesis if A: "check_diag n (uncurry (restrict_zero n M x))"
+  proof -
+    from A have "check_diag n (uncurry (pre_reset n M x))"
+      unfolding pre_reset_def by (rule check_diag_free)
+    then have "\<lbrakk>pre_reset n M x\<rbrakk> = {}"
+      by (rule check_diag_empty)
+    from A have "\<lbrakk>restrict_zero n M x\<rbrakk> = {}"
+      by (rule check_diag_empty)
+    then have "{u. (u(x := (0::'t::time))) \<in> \<lbrakk>M\<rbrakk>} \<inter> {u. u x \<ge> 0} = {}"
+      using \<open>0 < x\<close> \<open>x \<le> n\<close> by (auto simp: restrict_zero_correct)
+    with \<open>\<lbrakk>pre_reset n M x\<rbrakk> = {}\<close> show ?thesis
+      by simp
+  qed
+  from that(3) show ?thesis
+  proof
+    assume "canonical M n"
+    from restrict_zero_canonical[OF \<open>canonical M n\<close> \<open>x \<le> n\<close>] have
+      "canonical (restrict_zero n M x) n \<or> check_diag n (uncurry (restrict_zero n M x))"
+      (is "?A \<or> ?B") .
+    then consider ?A "\<not> ?B" | ?B
+      by blast
+    then show ?thesis
+    proof cases
+      case 1
+      assume ?A "\<not> ?B"
+      moreover from \<open>\<not> ?B\<close> have "\<forall>c\<le>n. 0 \<le> restrict_zero n M x c c"
+        unfolding check_diag_def by (auto simp: DBM.neutral)
+      moreover have "\<forall>u\<in>\<lbrakk>restrict_zero n M x\<rbrakk>. 0 \<le> u x"
+        by (simp add: restrict_zero_correct that)
+      moreover from \<open>x \<le> n\<close> \<open>M 0 x \<le> 0\<close> have "restrict_zero n M x 0 x \<le> 0"
+        by (blast intro: order.trans restrict_zero_mono)
+      moreover from \<open>x \<le> n\<close> \<open>M 0 0 \<le> 0\<close> have "restrict_zero n M x 0 0 \<le> 0"
+        by (blast intro: order.trans restrict_zero_mono)
+      ultimately show ?thesis
+        using that
+        by (auto simp: pre_reset_correct_aux restrict_zero_correct free_correct pre_reset_def)
+    next
+      assume ?B then show ?thesis
+        by (rule check_diag)
+    qed
+  next
+    assume "check_diag n (uncurry M)"
+    then have "check_diag n (uncurry (restrict_zero n M x))"
+      by (rule restrict_zero_diag)
+    then show ?thesis
+      by (rule check_diag)
+  qed
+qed
+
+lemma zone_set_pre_Cons:
+  "zone_set_pre \<lbrakk>M\<rbrakk> (x # r) = zone_set_pre {u. (u(x := (0::'t::time))) \<in> \<lbrakk>M\<rbrakk>} r"
+  unfolding zone_set_pre_def by auto
+
+lemma pre_reset_list_Cons:
+  "pre_reset_list n M (x # r) = pre_reset_list n (pre_reset n M x) r"
+  unfolding pre_reset_list_def by simp
+
+lemma pre_reset_diag:
+  "check_diag n (uncurry (pre_reset n M x))" if "check_diag n (uncurry M)"
+  using that unfolding pre_reset_def by (intro check_diag_free restrict_zero_diag)
+
+lemma free_canonical':
+  "canonical (free n (M :: (_ :: time) DBM) x) n \<or> check_diag n (uncurry (free n M x))"
+  if "canonical M n \<or> check_diag n (uncurry M)" "x \<le> n"
+  by (smt check_diag_def check_diag_free dbm_entry_le_iff(5) free_canonical leI
+          order_mono_setup.refl order_trans that uncurry_apply
+     )
+
+lemma pre_reset_canonical':
+  "canonical (pre_reset n (M :: (_ :: time) DBM) x) n \<or> check_diag n (uncurry (pre_reset n M x))"
+  if "canonical M n \<or> check_diag n (uncurry M)" "x \<le> n"
+  using that(1)
+proof standard
+  assume "canonical M n"
+  with \<open>x \<le> n\<close> have
+    "canonical (restrict_zero n M x) n \<or> check_diag n (uncurry (restrict_zero n M x))"
+    by (intro restrict_zero_canonical)
+  with \<open>x \<le> n\<close> show ?thesis
+    unfolding pre_reset_def by (intro free_canonical')
+next
+  assume "check_diag n (uncurry M)"
+  from pre_reset_diag[OF this] show ?thesis ..
+qed
+
+lemma pre_reset_list_correct:
+  "\<lbrakk>pre_reset_list n M r\<rbrakk> = zone_set_pre \<lbrakk>M\<rbrakk> r \<inter> {u. \<forall> x \<in> set r. u x \<ge> 0}"
+  if "\<forall> x \<in> set r. x > 0 \<and> x \<le> n"
+    "canonical M n \<or> check_diag n (uncurry M)" "\<forall> x \<in> set r. M 0 x \<le> 0" "M 0 0 \<le> 0"
+  using that
+  apply (induction r arbitrary: M)
+   apply (simp add: zone_set_pre_def pre_reset_list_def)
+  subgoal premises prems for x r M
+    apply (subst zone_set_pre_Cons)
+    apply (subst pre_reset_list_Cons)
+    apply (subst prems)
+        prefer 5
+        apply (subst pre_reset_correct)
+             prefer 6
+    subgoal
+      unfolding zone_set_pre_def by (cases "x \<in> set r") auto
+    using prems(2-) apply (solves auto)+
+    subgoal
+      using prems(2-) by (intro pre_reset_canonical'; auto)
+    subgoal
+      unfolding pre_reset_def free_def using prems(2-)
+      by (auto 4 3 intro: order.trans restrict_zero_mono)
+    subgoal
+      unfolding pre_reset_def free_def using prems(2-)
+      by (auto 4 3 intro: order.trans restrict_zero_mono)
+    done
+  done
+
+end (* Default Clock Numbering *)
+
+text \<open>
+Computes \<open>dbm_list xs - \<lbrakk>M\<rbrakk> = dbm_list xs \<inter> (- \<lbrakk>M\<rbrakk>)\<close> by negating each entry of \<open>M\<close>
+and intersecting it with each member of \<open>xs\<close>.
+\<close>
+definition
+  "dbm_minus_canonical n xs M =
+   [and_entry_repair n j i (neg_dbm_entry (M i j)) M'.
+     (i, j) \<leftarrow> [(i, j).
+        i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> M i j \<noteq> \<infinity>],
+     M'     \<leftarrow> xs
+   ]"
+
+text \<open>
+Same as @{text dbm_minus_canonical} but filters out empty DBMs.
+\<close>
+definition
+  "dbm_minus_canonical_check n xs M =
+  filter (\<lambda>M. \<not> check_diag n (uncurry M)) (dbm_minus_canonical n xs M)"
+
+text \<open>Checks whether \<open>\<lbrakk>M\<rbrakk> - dbm_list xs = {}\<close>.\<close>
+definition
+  "dbm_subset_fed n M xs \<equiv>
+    let xs = filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs in
+    list_all (\<lambda> M. check_diag n (uncurry M)) (fold (\<lambda>m S. dbm_minus_canonical n S m) xs [M])"
+
+definition
+  "dbm_subset_fed_check n M xs \<equiv>
+    let
+      xs = filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs;
+      is_direct_subset = list_ex (\<lambda> M'. dbm_subset' n (uncurry M) (uncurry M')) xs
+    in is_direct_subset \<or>
+       list_all (\<lambda>M. check_diag n (uncurry M)) (fold (\<lambda>m S. dbm_minus_canonical_check n S m) xs [M])"
+
+definition "canonical' n M \<equiv> canonical M n \<or> check_diag n (uncurry M)"
+
+lemma canonical'I:
+  "canonical' n (f M)" if
+  "canonical' n M"
+  "canonical M n \<Longrightarrow> canonical' n (f M)" "check_diag n (uncurry M) \<Longrightarrow> check_diag n (uncurry (f M))"
+  using that unfolding canonical'_def by metis
+
+lemma check_diag_repair_pair:
+  assumes "check_diag n (uncurry M)"
+  shows "check_diag n (uncurry (repair_pair n M i j))"
+  using assms repair_pair_mono[where M = M and a = i and b = j] unfolding check_diag_def by force
+
+lemma check_diag_and_entry:
+  assumes "check_diag n (uncurry M)"
+  shows "check_diag n (uncurry (and_entry a b e M))"
+  using assms unfolding check_diag_def
+  apply (elim exE)
+  subgoal for i
+    using and_entry_mono[where M = M and a = a and b = b and e = e, of i i] by auto
+  done
+
+lemma canonical'_and_entry_repair:
+  "canonical' n (and_entry_repair n i j e M)" if "canonical' n M" "i \<le> n" "j \<le> n"
+  using that(1)
+proof (rule canonical'I)
+  assume "canonical M n"
+  from \<open>i \<le> n\<close> \<open>j \<le> n\<close> have *: "{0..n} - {i, j} \<union> {i, j} = {0..n}"
+    by auto
+  define M1 where "M1 = and_entry i j e M"
+  from \<open>canonical M n\<close> have "canonical_subs n {0..n} M"
+    unfolding canonical_alt_def .
+  with * have "canonical_subs n ({0..n} - {i, j}) M1"
+    unfolding and_entry_def M1_def canonical_subs_def by (auto simp: min.coboundedI1)
+  from repair_pair_characteristic[OF this, of i j] \<open>i \<le> n\<close> \<open>j \<le> n\<close> have
+    "canonical' n (repair_pair n M1 i j)"
+    unfolding canonical'_def
+    unfolding canonical_alt_def check_diag_def * neutral by auto
+  then show ?thesis
+    unfolding and_entry_repair_def M1_def Let_def .
+next
+  assume "check_diag n (uncurry M)"
+  then show "check_diag n (uncurry (and_entry_repair n i j e M))"
+    unfolding and_entry_repair_def by (intro check_diag_repair_pair check_diag_and_entry)
+qed
+
+lemma dbm_minus_canonical_canonical':
+  "\<forall>M \<in> set (dbm_minus_canonical n xs m). canonical' n M" if "\<forall>M \<in> set xs. canonical' n M"
+  using that unfolding dbm_minus_canonical_def
+  by (auto split: if_split_asm intro: canonical'_and_entry_repair)
+
+lemma dbm_minus_canonical_check_canonical':
+  "\<forall>M \<in> set (dbm_minus_canonical_check n xs m). canonical' n M" if "\<forall>M \<in> set xs. canonical' n M"
+  using dbm_minus_canonical_canonical'[OF that] unfolding dbm_minus_canonical_check_def by auto
+
+subsection \<open>Correctness of @{term dbm_subset_fed}\<close>
+
+paragraph \<open>Misc\<close>
+
+lemma list_all_iffI:
+  assumes "\<forall> x \<in> set xs. \<exists> y \<in> set ys. P x \<longleftrightarrow> Q y"
+      and "\<forall> y \<in> set ys. \<exists> x \<in> set xs. P x \<longleftrightarrow> Q y"
+    shows "list_all P xs \<longleftrightarrow> list_all Q ys"
+  using assms unfolding list_all_def by blast
+
+lemma list_all_iff_list_all2I:
+  assumes "list_all2 (\<lambda> x y. P x \<longleftrightarrow> Q y) xs ys"
+  shows "list_all P xs \<longleftrightarrow> list_all Q ys"
+  using assms by (intro list_all_iffI list_all2_set1 list_all2_set2)
+
+lemma list_all2_mapI:
+  assumes "list_all2 (\<lambda> x y. P (f x) (g y)) xs ys"
+  shows "list_all2 P (map f xs) (map g ys)"
+  using assms by (simp only: list.rel_map)
+
+context Default_Nat_Clock_Numbering
+begin
+
+lemma canonical_empty_zone:
+  "[M]v,n = {} \<longleftrightarrow> (\<exists>i\<le>n. M i i < 0)" if "canonical M n"
+  using v_0 that surj_on by (intro canonical_empty_zone) auto
+
+lemma check_diag_iff_empty:
+  "check_diag n (uncurry M) \<longleftrightarrow> \<lbrakk>M\<rbrakk> = {}" if "canonical' n M"
+proof (safe, goal_cases)
+  case 2
+  show ?case
+  proof -
+    from that
+    show ?thesis
+      unfolding canonical'_def
+    proof
+      assume "canonical M n"
+      from canonical_empty_zone[OF this] \<open>\<lbrakk>M\<rbrakk> = {}\<close> have
+        "\<exists>i\<le>n. M i i < 0"
+        by auto
+      then show ?thesis unfolding check_diag_def neutral
+        by auto
+    next
+      assume "check_diag n (uncurry M)"
+      then show ?thesis unfolding check_diag_def neutral by auto
+    qed
+  qed
+qed (auto dest: check_diag_empty)
+
+lemma and_entry_repair_zone_equiv:
+  "\<lbrakk>and_entry_repair n a b e M\<rbrakk> = \<lbrakk>and_entry a b e M\<rbrakk>" if "a \<le> n" "b \<le> n"
+  unfolding and_entry_repair_def using that by (rule repair_pair_zone_equiv)
+
+lemma dbm_minus_rel:
+  assumes "list_all2 (\<lambda>x y. \<lbrakk>x\<rbrakk> = \<lbrakk>y\<rbrakk>) ms ms'"
+  shows "list_all2 (\<lambda>x y. \<lbrakk>x\<rbrakk> = \<lbrakk>y\<rbrakk>) (dbm_minus n ms m) (dbm_minus_canonical n ms' m)"
+  unfolding dbm_minus_def dbm_minus_canonical_def
+  apply (rule concat_transfer[unfolded rel_fun_def, rule_format])
+  apply (rule list_all2_mapI)
+  apply (rule list.rel_refl_strong)
+  apply (auto 4 3
+      intro: list_all2_mapI list_all2_mono[OF assms]
+      simp: and_entry_repair_zone_equiv and_entry_correct split: if_split_asm
+      )
+  done
+
+lemma dbm_minus_canonical_fold_canonical':
+  "\<forall>M \<in> set (fold (\<lambda>m S. dbm_minus_canonical n S m) xs ms). canonical' n M"
+  if "\<forall>M \<in> set ms. canonical' n M" for ms and xs :: "('t ::time) DBM list"
+  using that by (induction xs arbitrary: ms) (auto dest: dbm_minus_canonical_canonical')
+
+lemma not_check_diag_nonnegD:
+  "M i i \<ge> 0" if "\<not> check_diag n (uncurry M)" "i \<le> n"
+  using that unfolding check_diag_def by (auto simp: DBM.less_eq[symmetric] neutral)
+
+theorem dbm_subset_fed_correct:
+  fixes xs :: "(nat \<Rightarrow> nat \<Rightarrow> ('t ::time) DBMEntry) list"
+    and S :: "(nat \<Rightarrow> nat \<Rightarrow> 't DBMEntry) list"
+  assumes "canonical' n M"
+  shows "\<lbrakk>M\<rbrakk> \<subseteq> (\<Union>m\<in>set xs. \<lbrakk>m\<rbrakk>) \<longleftrightarrow> dbm_subset_fed n M xs"
+proof -
+  have *: "list_all2 (\<lambda>x y. \<lbrakk>x\<rbrakk> = \<lbrakk>y\<rbrakk>)
+     (fold (\<lambda>m S. dbm_minus n S m) xs ms)
+     (fold (\<lambda>m S. dbm_minus_canonical n S m) xs ms')"
+    if "list_all2 (\<lambda>x y. \<lbrakk>x\<rbrakk> = \<lbrakk>y\<rbrakk>) ms ms'" for ms ms' and xs :: "'t DBM list"
+    using that
+  proof (induction xs arbitrary: ms ms')
+    case Nil
+    then show ?case
+      by simp
+  next
+    case prems: (Cons a xs)
+    from this(2) show ?case
+      by simp (intro prems(1) dbm_minus_rel)
+  qed
+  let ?xs = "filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs"
+  have *: "list_all2 (\<lambda>x y. \<lbrakk>x\<rbrakk> = \<lbrakk>y\<rbrakk>)
+     (fold (\<lambda>m S. dbm_minus n S m) ?xs [M])
+     (fold (\<lambda>m S. dbm_minus_canonical n S m) ?xs [M])"
+    by (rule *) simp
+  have **:"(\<Union>m\<in>set xs. \<lbrakk>m\<rbrakk>) = (\<Union>m\<in>set (filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs). \<lbrakk>m\<rbrakk>)"
+    by (auto simp: check_diag_empty)
+  show ?thesis
+    apply (subst **)
+    apply (subst dbm_list_superset_op[where S = "[M]", simplified])
+    subgoal
+      by (auto dest: not_check_diag_nonnegD simp: neutral[symmetric] DBM.less_eq)
+    subgoal
+      unfolding dbm_subset_fed_def Let_def
+      using dbm_minus_canonical_fold_canonical'[of "[M]"] \<open>canonical' n M\<close>
+      by (intro list_all_iff_list_all2I list.rel_mono_strong[OF *])(auto dest: check_diag_iff_empty)
+    done
+qed
+
+lemma dbm_minus_canonical_check_fed_equiv:
+  "dbm_list (dbm_minus_canonical_check n S m) = dbm_list (dbm_minus_canonical n S m)"
+  unfolding dbm_minus_canonical_check_def by (auto simp: check_diag_empty)
+
+lemma dbm_minus_canonical_dbm_minus:
+  "dbm_list (dbm_minus_canonical n xs m) = dbm_list (dbm_minus n xs m)"
+  using dbm_minus_rel[of xs xs m] unfolding list_all2_same
+  by (force dest: list_all2_set1 list_all2_set2)
+
+lemma dbm_minus_canonical_fed_equiv:
+  "dbm_list (dbm_minus_canonical n xs m) = dbm_list (dbm_minus_canonical n xs' m)"
+  if "dbm_list xs = dbm_list xs'" "0 \<le> m 0 0"
+  unfolding dbm_minus_canonical_dbm_minus
+  using that by (auto simp: neutral dbm_list_subtract[symmetric] DBM.less_eq)
+
+theorem dbm_subset_fed_correct':
+  fixes xs :: "(nat \<Rightarrow> nat \<Rightarrow> ('t ::time) DBMEntry) list"
+    and S :: "(nat \<Rightarrow> nat \<Rightarrow> 't DBMEntry) list"
+  assumes "canonical' n M"
+  shows "\<lbrakk>M\<rbrakk> \<subseteq> (\<Union>m\<in>set xs. \<lbrakk>m\<rbrakk>) \<longleftrightarrow> (
+    let xs = filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs in
+    list_all (\<lambda> M. check_diag n (uncurry M)) (fold (\<lambda>m S. dbm_minus_canonical_check n S m) xs [M]))"
+proof -
+  have canonical: "\<forall>M \<in> set (fold (\<lambda>m S. dbm_minus_canonical_check n S m) xs ms). canonical' n M"
+    if "\<forall>M \<in> set ms. canonical' n M" for ms and xs :: "'t DBM list"
+    using that by (induction xs arbitrary: ms) (auto dest: dbm_minus_canonical_check_canonical')
+  have *: "dbm_list (fold (\<lambda>m S. dbm_minus_canonical_check n S m) xs ms) =
+    dbm_list (fold (\<lambda>m S. dbm_minus_canonical n S m) xs ms')"
+    if "dbm_list ms = dbm_list ms'" "\<forall>m \<in> set xs. m 0 0 \<ge> 0" for ms ms' and xs :: "'t DBM list"
+    using that
+  proof (induction xs arbitrary: ms ms')
+    case Nil
+    then show ?case
+      by simp
+  next
+    case (Cons a xs)
+    from Cons.prems show ?case
+      by - (simp, rule Cons.IH,
+          auto intro!: dbm_minus_canonical_fed_equiv simp add: dbm_minus_canonical_check_fed_equiv
+          )
+  qed
+  define xs' where "xs' = filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs"
+  have *: "dbm_list (fold (\<lambda>m S. dbm_minus_canonical_check n S m) xs' [M]) =
+        dbm_list (fold (\<lambda>m S. dbm_minus_canonical n S m) xs' [M])"
+    by (auto intro!: * dest: not_check_diag_nonnegD simp: xs'_def)
+  have **: "list_all (\<lambda> M. check_diag n (uncurry M)) xs \<longleftrightarrow> dbm_list xs = {}"
+    if "\<forall>M \<in> set xs. canonical' n M" for xs :: "'t DBM list"
+    using that by (metis (mono_tags, lifting) Ball_set SUP_bot_conv(2) check_diag_iff_empty)
+  show ?thesis
+    unfolding 
+      dbm_subset_fed_correct[OF \<open>canonical' _ _\<close>] dbm_subset_fed_def
+      xs'_def[symmetric] Let_def
+    apply (subst **)
+     defer
+     apply (subst **)
+    using assms by (auto intro!: dbm_minus_canonical_fold_canonical' canonical simp: *)
+qed
+
+lemma subset_if_pointwise_le:
+  "\<lbrakk>M\<rbrakk> \<subseteq> \<lbrakk>M'\<rbrakk>" if "pointwise_cmp (\<le>) n M M'"
+  using that by (simp add: DBM.less_eq DBM_le_subset pointwise_cmp_def subsetI)
+
+theorem dbm_subset_fed_check_correct:
+  fixes xs :: "(nat \<Rightarrow> nat \<Rightarrow> ('t ::time) DBMEntry) list"
+    and S :: "(nat \<Rightarrow> nat \<Rightarrow> 't DBMEntry) list"
+  assumes "canonical' n M"
+  shows "\<lbrakk>M\<rbrakk> \<subseteq> (\<Union>m\<in>set xs. \<lbrakk>m\<rbrakk>) \<longleftrightarrow> dbm_subset_fed_check n M xs"
+proof -
+  define xs' where "xs' = filter (\<lambda>M. \<not> check_diag n (uncurry M)) xs"
+  define is_direct_subset where
+    "is_direct_subset = list_ex (\<lambda> M'. dbm_subset' n (uncurry M) (uncurry M')) xs'"
+  have "\<lbrakk>M\<rbrakk> \<subseteq> (\<Union>m\<in>set xs. \<lbrakk>m\<rbrakk>)" if is_direct_subset
+  proof -
+    from that have "\<exists>m \<in> set xs'. \<lbrakk>M\<rbrakk> \<subseteq> \<lbrakk>m\<rbrakk>"
+      unfolding is_direct_subset_def list_ex_iff dbm_subset'_def
+      by (auto intro!: subset_if_pointwise_le)
+    then show ?thesis
+      unfolding xs'_def by auto
+  qed
+  then show ?thesis
+    apply (cases is_direct_subset; simp add:
+        dbm_subset_fed_check_def is_direct_subset_def dbm_subset_fed_def xs'_def[symmetric]
+        )
+    unfolding dbm_subset_fed_correct[
+        OF \<open>canonical' n M\<close>, of xs, unfolded Let_def dbm_subset_fed_def, folded xs'_def, symmetric
+        ]
+    unfolding dbm_subset_fed_correct'[
+        OF \<open>canonical' n M\<close>, of xs, folded xs'_def, unfolded Let_def, symmetric
+        ]
+    ..
+qed
+
+end (* Default Clock Numbering *)
+
+
+subsection \<open>Refined DBM Operations\<close>
+
+(* This is the odd one out *)
+definition
+  "V_dbm = (\<lambda> i j. if i = j then 0 else if i = 0 \<and> j > 0 then 0 else \<infinity>)"
+
+definition and_entry_upd ::
+  "nat \<Rightarrow> nat \<Rightarrow> int DBMEntry \<Rightarrow> int DBM' \<Rightarrow> int DBM'" where
+  "and_entry_upd a b e M = M((a,b) := min (M (a, b)) e)"
+
+definition
+  "and_entry_repair_upd n a b e M \<equiv>
+   Normalized_Zone_Semantics_Impl_Semantic_Refinement.repair_pair n (and_entry_upd a b e M) a b"
+
+definition
+  "dbm_minus_canonical_upd n xs m =
+  concat (map (\<lambda>(i, j). map (\<lambda> M. and_entry_repair_upd n j i (neg_dbm_entry (m (i, j))) M) xs)
+    [(i, j). i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m (i, j) \<noteq> \<infinity>])"
+
+definition
+  "dbm_minus_canonical_check_upd n xs M =
+  filter (\<lambda>M. \<not> check_diag n M) (dbm_minus_canonical_upd n xs M)"
+
+definition
+  "dbm_subset_fed_upd n M xs \<equiv>
+   let xs = filter (\<lambda>M. \<not> check_diag n M) xs;
+       is_direct_subset = list_ex (\<lambda>M'. dbm_subset' n M M') xs
+   in is_direct_subset \<or>
+     list_all (\<lambda> M. check_diag n M) (fold (\<lambda>m S. dbm_minus_canonical_check_upd n S m) xs [M])"
+
+lemma list_all_filter_neg:
+  "list_all P (filter (\<lambda>x. \<not> P x) xs) \<longleftrightarrow> (filter (\<lambda>x. \<not> P x) xs) = []"
+  by (metis Cons_eq_filterD list_all_simps(2) list_pred_cases)
+
+lemma dbm_subset_fed_upd_alt_def:
+  "dbm_subset_fed_upd n M xs \<equiv>
+   let xs = filter (\<lambda>M. \<not> check_diag n M) xs
+   in if xs = [] then check_diag n M
+      else if list_ex (\<lambda>M'. dbm_subset' n M M') xs then True
+      else fold (\<lambda>m S. dbm_minus_canonical_check_upd n S m) xs [M] = []"
+  unfolding dbm_subset_fed_upd_def short_circuit_conv using list_last
+  by (intro eq_reflection;force simp: list_all_filter_neg dbm_minus_canonical_check_upd_def Let_def)
+
+definition
+  "V_dbm' n = (\<lambda>(i, j). (if i = j \<or> i = 0 \<and> j > 0 \<or> i > n \<or> j > n then 0 else \<infinity>))"
+
+definition
+  down_upd :: "nat \<Rightarrow> _ DBM' \<Rightarrow> _ DBM'"
+where
+  "down_upd n M \<equiv> \<lambda>(i, j).
+  if i = 0 \<and> j > 0 \<and> i \<le> n \<and> j \<le> n then Min ({Le 0} \<union> {M (k, j) | k. 1 \<le> k \<and> k \<le> n}) else M (i, j)"
+
+definition
+  "restrict_zero_upd n M x \<equiv>
+    let
+      M1 = and_entry_upd x 0 (Le 0) M;
+      M2 = and_entry_upd 0 x (Le 0) M1
+    in Normalized_Zone_Semantics_Impl_Semantic_Refinement.repair_pair n M2 x 0"
+
+definition
+  free_upd :: "nat \<Rightarrow> _ DBM' \<Rightarrow> nat \<Rightarrow> _ DBM'"
+where
+  "free_upd n M x \<equiv>
+   \<lambda> (i, j). if i = x \<and> j \<noteq> x \<and> i \<le> n \<and> j \<le> n
+    then \<infinity> else if i \<noteq> x \<and> j = x \<and> i \<le> n \<and> j \<le> n then M (i, 0) else M (i, j)"
+
+definition
+  "pre_reset_upd n M x \<equiv> free_upd n (restrict_zero_upd n M x) x"
+
+definition
+  "pre_reset_list_upd n M r \<equiv> fold (\<lambda> x M. pre_reset_upd n M x) r M"
+
+
+subsection \<open>Transferring Properties\<close>
+
+context
+  includes lifting_syntax
+begin
+
+lemma neg_dbm_entry_transfer[transfer_rule]:
+  "(rel_DBMEntry ri ===> rel_DBMEntry ri) neg_dbm_entry neg_dbm_entry"
+  by (auto elim!: DBMEntry.rel_cases intro!: rel_funI)
+
+lemma fold_min_transfer:
+  "((list_all2 (rel_DBMEntry ri)) ===> rel_DBMEntry ri ===> rel_DBMEntry ri) (fold min) (fold min)"
+  by transfer_prover
+
+context
+  fixes n :: nat
+begin
+
+definition
+  "RI2 M D \<equiv> RI n (uncurry M) D"
+
+lemma and_entry_transfer[transfer_rule]:
+  "((=) ===> (=) ===> rel_DBMEntry ri ===> RI2 ===> RI2) and_entry and_entry_upd"
+  unfolding and_entry_def and_entry_upd_def
+  unfolding rel_fun_def eq_onp_def RI2_def
+  by (auto intro: min_ri_transfer[unfolded rel_fun_def, rule_format])
+
+lemma FWI_transfer[transfer_rule]:
+  "(RI2 ===> eq_onp (\<lambda>x. x = n) ===> eq_onp (\<lambda>x. x < Suc n) ===> RI2) FWI FWI'" (is "?A")
+and FW_transfer[transfer_rule]:
+  "(RI2 ===> eq_onp (\<lambda>x. x = n) ===> RI2) FW FW'" (is "?B")
+proof -
+  define RI' where
+    "RI' = (eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)"
+  have RI_iff: "RI' M M' \<longleftrightarrow> RI n (uncurry M) (uncurry M')" for M M'
+    unfolding RI'_def rel_fun_def by auto
+  { fix M D k assume "RI n (uncurry M) D" "k < Suc n"
+    then have "RI' M (curry D)"
+      by (simp add: RI_iff)
+    with \<open>k < Suc n\<close> have "RI' (FWI M n k) (FWI (curry D) n k)"
+      unfolding FWI_def
+      by (intro fwi_RI_transfer[of n, folded RI'_def, unfolded rel_fun_def, rule_format])
+        (auto simp: eq_onp_def)
+    then have "RI n (uncurry (FWI M n k)) (uncurry (FWI (curry D) n k))"
+      by (simp add: RI_iff)
+  } note * = this
+  show ?A
+    unfolding FWI'_def
+    unfolding rel_fun_def
+    unfolding RI2_def
+    apply clarsimp
+    apply (subst (asm) (3) eq_onp_def)
+    apply (subst (asm) (3) eq_onp_def)
+    apply clarsimp
+    by (intro *)
+  { fix M D assume "RI n (uncurry M) D"
+    then have "RI' M (curry D)"
+      by (simp add: RI_iff)
+    then have "RI' (FW M n) (FW (curry D) n)"
+      by (intro FW_RI_transfer[of n, folded RI'_def, unfolded rel_fun_def, rule_format])
+         (auto simp: eq_onp_def)
+    then have "RI n (uncurry (FW M n)) (FW' D n)"
+      by (simp add: RI_iff FW'_def)
+  } note * = this
+  show ?B
+    unfolding rel_fun_def
+    unfolding RI2_def
+    apply clarsimp
+    apply (subst (asm) (3) eq_onp_def)
+    apply clarsimp
+    by (intro *)
+qed
+
+lemma repair_pair_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> RI2)
+    repair_pair Normalized_Zone_Semantics_Impl_Semantic_Refinement.repair_pair"
+  unfolding repair_pair_def Normalized_Zone_Semantics_Impl_Semantic_Refinement.repair_pair_def
+  by transfer_prover
+
+lemma and_entry_transfer_weak:
+  "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> rel_DBMEntry ri ===> RI2 ===> RI2)
+    and_entry and_entry_upd"
+  using and_entry_transfer unfolding rel_fun_def eq_onp_def by auto
+
+lemma and_entry_repair_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> rel_DBMEntry ri
+    ===> RI2 ===> RI2) and_entry_repair and_entry_repair_upd
+  "
+  supply [transfer_rule] = and_entry_transfer_weak
+  unfolding and_entry_repair_def and_entry_repair_upd_def by transfer_prover
+
+lemma dbm_minus_canonical_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> list_all2 RI2 ===> RI2 ===> list_all2 RI2)
+    dbm_minus_canonical dbm_minus_canonical_upd
+  "
+  unfolding dbm_minus_canonical_def dbm_minus_canonical_upd_def
+  apply (intro rel_funI)
+  apply (rule concat_transfer[unfolded rel_fun_def, rule_format])
+  apply (rule list.map_transfer[unfolded rel_fun_def, rule_format,
+        where Rb = "rel_prod (eq_onp (\<lambda>x. x < Suc n)) (eq_onp (\<lambda>x. x < Suc n))"])
+   apply clarsimp
+  subgoal
+    apply (rule list_all2_mapI)
+    apply (erule list_all2_mono)
+    apply (rule and_entry_repair_transfer[unfolded rel_fun_def, rule_format])
+        apply assumption+
+    subgoal
+      unfolding RI2_def rel_fun_def
+      by (auto intro: neg_dbm_entry_transfer[unfolded rel_fun_def, rule_format])
+    apply assumption
+    done
+  subgoal premises prems for n1 n2 xs ys M D
+  proof -
+    have [simp]: "D (i, j) \<noteq> \<infinity> \<longleftrightarrow> M i j \<noteq> \<infinity>" if "i < Suc n" "j < Suc n" for i j
+      using prems that by (auto 4 3 simp: eq_onp_def rel_fun_def RI2_def elim!: DBMEntry.rel_cases)
+    from prems have [simp]: "n1 = n" "n2 = n"
+      by (auto simp: eq_onp_def)
+    let ?a = "(concat
+       (map (\<lambda>i. concat
+                   (map (\<lambda>j. if (0 < i \<or> 0 < j) \<and>
+                                 i \<le> n \<and> j \<le> n \<and> M i j \<noteq> \<infinity>
+                              then [(i, j)] else [])
+                     [0..<Suc n]))
+         [0..<Suc n]))"
+    let ?b = "(concat
+       (map (\<lambda>i. concat
+                   (map (\<lambda>j. if (0 < i \<or> 0 < j) \<and>
+                                 i \<le> n \<and> j \<le> n \<and> D (i, j) \<noteq> \<infinity>
+                              then [(i, j)] else [])
+                     [0..<Suc n]))
+         [0..<Suc n]))"
+    have "?b = ?a"
+      by (auto intro!: arg_cong[where f = concat] simp del: upt_Suc)
+    then show ?thesis
+      by (simp del: upt_Suc add: list_all2_same eq_onp_def)
+  qed
+  done
+
+lemma check_diag_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> (=)) (\<lambda>n M. check_diag n (uncurry M)) check_diag"
+  unfolding RI2_def rel_fun_def check_diag_def
+  by (auto 0 5 dest: neutral_RI simp: eq_onp_def less_Suc_eq_le neutral[symmetric])
+
+lemma dbm_minus_canonical_check_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> list_all2 RI2 ===> RI2 ===> list_all2 RI2)
+    dbm_minus_canonical_check dbm_minus_canonical_check_upd
+  "
+  unfolding dbm_minus_canonical_check_def dbm_minus_canonical_check_upd_def by transfer_prover
+
+lemma le_rel_DBMEntry_iff:
+  "a \<le> b \<longleftrightarrow> x \<le> y" if "rel_DBMEntry ri a x" "rel_DBMEntry ri b y"
+  using that by (auto elim!: DBMEntry.rel_cases simp: dbm_entry_simps ri_def)
+
+lemma dbm_subset'_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> RI2 ===> (=))
+    (\<lambda> n M M'. dbm_subset' n (uncurry M) (uncurry M')) dbm_subset'"
+  unfolding RI2_def rel_fun_def dbm_subset'_def
+  using le_rel_DBMEntry_iff by (clarsimp simp: eq_onp_def pointwise_cmp_def less_Suc_eq_le) meson
+
+lemma dbm_subset_fed_transfer:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> list_all2 RI2 ===> (=)) dbm_subset_fed_check dbm_subset_fed_upd"
+  unfolding dbm_subset_fed_check_def dbm_subset_fed_upd_def by transfer_prover
+
+lemma V_dbm_transfer[transfer_rule]:
+  "RI2 V_dbm (V_dbm' n)"
+  unfolding V_dbm_def V_dbm'_def RI2_def by (auto simp: rel_fun_def neutral eq_onp_def zero_RI)
+
+lemma down_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> RI2) down down_upd"
+  unfolding down_def down_upd_def
+  apply (clarsimp simp: rel_fun_def RI2_def eq_onp_def, goal_cases)
+  subgoal premises prems for M D x
+  proof -
+    have A: "(insert (Le 0) {M k x |k. Suc 0 \<le> k \<and> k \<le> n})
+        = (set (Le 0 # [M k x. k \<leftarrow> [1..<Suc n]]))"
+      by auto
+    have B: "insert (Le 0) {D (k, x) |k. Suc 0 \<le> k \<and> k \<le> n} 
+        = (set (Le 0 # [D (k, x). k \<leftarrow> [1..<Suc n]]))"
+      by auto
+    show ?thesis
+      unfolding A B Min.set_eq_fold
+      apply (rule fold_min_transfer[unfolded rel_fun_def, rule_format])
+      unfolding list.rel_map list_all2_same using prems by (auto simp: zero_RI)
+  qed
+  done
+
+lemma free_upd[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> (=) ===> RI2) free free_upd"
+  unfolding free_def free_upd_def by (auto simp: RI2_def rel_fun_def eq_onp_def)
+
+lemma pre_reset_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> eq_onp (\<lambda>x. x < Suc n) ===> RI2) pre_reset pre_reset_upd"
+proof -
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n"
+    by (simp add: eq_onp_def)
+  note [transfer_rule] = and_entry_transfer_weak
+  have [transfer_rule]:
+    "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> eq_onp (\<lambda>x. x < Suc n) ===> RI2) free free_upd"
+    using free_upd unfolding rel_fun_def eq_onp_def by blast
+  show ?thesis
+    unfolding pre_reset_def pre_reset_upd_def
+    unfolding restrict_zero_def restrict_zero_upd_def
+    by transfer_prover
+qed
+
+lemma pre_reset_list_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 ===> list_all2 (eq_onp (\<lambda>x. x < Suc n)) ===> RI2)
+    pre_reset_list pre_reset_list_upd"
+  unfolding pre_reset_list_def pre_reset_list_upd_def by transfer_prover
+
+end (* Fixed n *)
+end (* Lifting Syntax *)
+
+(* Unused *)
+definition unbounded_dbm where
+  "unbounded_dbm \<equiv> \<lambda> i j. if i = j then 0 else \<infinity>"
+
+lemma canonical_unbounded_dbm:
+  "canonical unbounded_dbm n"
+  by (auto simp: unbounded_dbm_def any_le_inf)
+
+lemma diag_unbounded_dbm:
+  "unbounded_dbm i i = 0"
+  unfolding unbounded_dbm_def by simp
+
+lemma down_diag:
+  "down n M i i = M i i"
+  unfolding down_def by auto
+
+definition
+  "abstr_FW n cc M v \<equiv> FW (abstr cc M v) n"
+
+definition
+  "abstr_FW_upd n cc M \<equiv> FW' (abstr_upd cc M) n"
+
+lemma abstr_mono:
+  "abstr cc M v i j \<le> M i j"
+  by (subst abstr.simps, induction cc arbitrary: M) (auto intro: order.trans abstra_mono)
+
+lemma abstr_FW_mono:
+  "abstr_FW n cc M v i j \<le> M i j" if "i \<le> n" "j \<le> n"
+  unfolding abstr_FW_def by (blast intro: that abstr_mono fw_mono order.trans)
+
+lemma abstr_FW_diag_preservation:
+  "\<forall>k\<le>n. abstr_FW n cc M v k k \<le> 0" if "\<forall>k\<le>n. M k k \<le> 0"
+  using that by (blast intro: abstr_FW_mono order.trans)
+
+lemma FW_canonical:
+  "canonical' n (FW M n)"
+  unfolding canonical'_def using FW_canonical[of n M] by (simp add: check_diag_def neutral)
+
+lemma abstr_FW_canonical:
+  "canonical' n (abstr_FW n cc M v)"
+  unfolding abstr_FW_def by (rule FW_canonical)
+
+lemma down_check_diag:
+  "check_diag n (uncurry (down n M))" if "check_diag n (uncurry M)"
+  using that unfolding check_diag_def down_def by force
+
+context Default_Nat_Clock_Numbering
+begin
+
+lemma clock_numbering:
+  "\<forall> c. v c > 0 \<and> (\<forall>x. \<forall>y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+  by (metis neq0_conv v_0 v_id)
+
+lemma V_dbm_correct:
+  "\<lbrakk>V_dbm\<rbrakk> = V"
+  unfolding V_def DBM_zone_repr_def DBM_val_bounded_alt_def2 
+  by (auto simp: dbm_entry_val'_iff_bounded V_dbm_def dbm_entry_simps)
+
+lemma abstr_correct:
+  "\<lbrakk>abstr cc M v\<rbrakk> = \<lbrakk>M\<rbrakk> \<inter> {u. u \<turnstile> cc}" if "\<forall>c\<in>collect_clks cc. 0 < c \<and> c \<le> n"
+  apply (rule dbm_abstr_zone_eq2)
+  subgoal
+    by (rule clock_numbering)
+  subgoal
+    using that by (auto simp: v_is_id)
+  done
+
+(* Unused *)
+lemma unbounded_dbm_correct:
+  "\<lbrakk>unbounded_dbm\<rbrakk> = UNIV"
+  unfolding DBM_zone_repr_def DBM_val_bounded_alt_def2 unbounded_dbm_def neutral
+  by (simp add: dbm_entry_val'_iff_bounded any_le_inf)
+
+(* Unused *)
+lemma abstr_correct':
+  "\<lbrakk>abstr cc unbounded_dbm v\<rbrakk> = {u. u \<turnstile> cc}" if "\<forall>c\<in>collect_clks cc. 0 < c \<and> c \<le> n"
+  by (simp add: unbounded_dbm_correct abstr_correct[OF that] del: abstr.simps)
+
+lemma abstr_FW_correct:
+  "\<lbrakk>abstr_FW n cc M v\<rbrakk> = \<lbrakk>M\<rbrakk> \<inter> {u. u \<turnstile> cc}" if "\<forall>c\<in>collect_clks cc. 0 < c \<and> c \<le> n"
+  unfolding abstr_FW_def by (subst FW_zone_equiv[symmetric]; intro surj_on abstr_correct that)
+
+lemma abstr_FW_correct':
+  "\<lbrakk>abstr_FW n cc unbounded_dbm v\<rbrakk> = {u. u \<turnstile> cc}" if "\<forall>c\<in>collect_clks cc. 0 < c \<and> c \<le> n"
+  by (simp add: unbounded_dbm_correct abstr_FW_correct[OF that])
+
+lemma down_V:
+  "\<lbrakk>down n M\<rbrakk> \<subseteq> V"
+  by (rule V_structuralI) (auto simp: down_def neutral intro: Min_le)
+
+lemma down_correct':
+  "\<lbrakk>down n M\<rbrakk> = \<lbrakk>M\<rbrakk>\<down> \<inter> V" if "canonical M n"
+  apply safe
+  subgoal for u
+    by (erule down_sound, rule that)
+  subgoal for u
+    using down_V by blast
+  subgoal for u
+    unfolding V_def by (erule down_complete) simp
+  done
+
+lemma down_correct:
+  "\<lbrakk>down n M\<rbrakk> = \<lbrakk>M\<rbrakk>\<down> \<inter> V" if "canonical' n M"
+  using that unfolding canonical'_def
+  apply standard
+  subgoal
+    by (rule down_correct')
+  by (frule down_check_diag) (simp add: check_diag_empty zone_time_pre_def)
+
+lemma pre_reset_diag_preservation:
+  "pre_reset n M x i i \<le> M i i" if "i \<le> n"
+  unfolding pre_reset_def by (auto simp add: free_diag restrict_zero_mono that)
+
+lemma pre_reset_list_diag:
+  "pre_reset_list n M r i i \<le> M i i" if "i \<le> n"
+  apply (induction r arbitrary: M)
+   apply (simp add: pre_reset_list_def; fail)
+  apply (simp add: pre_reset_list_Cons, blast intro: pre_reset_diag_preservation that order.trans)
+  done
+
+context
+  includes lifting_syntax
+begin
+
+lemma abstra_upd_abstra:
+  "abstra_upd ac M (i, j) = abstra ac (curry M) v i j" if "0 < constraint_clk ac" "i \<le> n" "j \<le> n"
+  using that by (cases ac) (auto simp: le_n_iff v_is_id v_0)
+
+lemma abstra_transfer[transfer_rule]:
+  "(rel_acconstraint (eq_onp (\<lambda> x. 0 < x \<and> x < Suc n)) ri ===> RI2 n ===> RI2 n)
+    (\<lambda> cc M. abstra cc M v) abstra_upd"
+  apply (intro rel_funI)
+  apply (subst RI2_def)
+  apply (intro rel_funI)
+  apply (elim rel_prod.cases)
+  apply (simp only:)
+  apply (subst abstra_upd_abstra)
+  by (auto 4 3 simp: eq_onp_def RI2_def rel_fun_def
+       intro!: min_ri_transfer[unfolded rel_fun_def, rule_format]
+       elim!: acconstraint.rel_cases
+     )
+
+lemma abstr_transfer[transfer_rule]:
+  "(list_all2 (rel_acconstraint (eq_onp (\<lambda> x. 0 < x \<and> x < Suc n)) ri) ===> RI2 n ===> RI2 n)
+    (\<lambda> cc M. abstr cc M v) abstr_upd"
+  unfolding abstr.simps abstr_upd_def by transfer_prover
+
+lemma abstr_FW_transfer[transfer_rule]:
+  "(list_all2 (rel_acconstraint (eq_onp (\<lambda> x. 0 < x \<and> x < Suc n)) ri) ===> RI2 n ===> RI2 n)
+    (\<lambda> cc M. abstr_FW n cc M v) (abstr_FW_upd n)"
+proof -
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n"
+    by (simp add: eq_onp_def)
+  show ?thesis
+    unfolding abstr_FW_def abstr_FW_upd_def by transfer_prover
+qed
+
+end
+
+end
+
+lemma RI2_trivial_transfer[transfer_rule]: "(RI2 n) (curry (conv_M M)) M"
+  unfolding RI2_def rel_fun_def by (auto simp: eq_onp_def)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Deadlock/Deadlock_Checking.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Deadlock/Deadlock_Checking.thy
@@ -0,0 +1,184 @@
+theory Deadlock_Checking
+  imports Deadlock_Impl Munta_Model_Checker.UPPAAL_Model_Checking
+begin
+
+paragraph \<open>Standard Deadlock Checker Implementation\<close>
+context Reachability_Problem_Impl
+begin
+
+definition deadlock_checker where
+  "deadlock_checker \<equiv>
+    let
+      key = return \<circ> fst;
+      sub = subsumes_impl;
+      copy = state_copy_impl;
+      start = a0_impl;
+      final = (\<lambda>_. return False);
+      succs = succs_impl;
+      empty = emptiness_check_impl;
+      P = check_deadlock_neg_impl;
+      trace = tracei
+    in do {
+      r1 \<leftarrow> is_start_in_states_impl;
+      if r1 then do {
+        r2 \<leftarrow> check_passed_impl succs start final sub empty key copy trace P;
+        return r2
+      }
+      else return True
+    }
+    "
+
+interpretation ta_bisim: Bisimulation_Invariant
+  "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  "(\<lambda>(l, u) (l', u'). l' = l \<and> (\<forall> c. c \<in> clk_set (conv_A A) \<longrightarrow> u c = u' c))"
+  "(\<lambda>_. True)" "(\<lambda>_. True)"
+  by (rule ta_bisimulation[of "conv_A A"])
+
+lemma deadlock_zero_clock_val_iff:
+  "(\<exists>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<and> deadlock (l0, u0)) \<longleftrightarrow> deadlock (l0, \<lambda>_. 0)"
+  apply safe
+  subgoal for u
+    using clocks_I[of "\<lambda>_. 0" u] by (subst ta_bisim.deadlock_iff; simp)
+  by auto
+
+context
+  assumes F_fun_False: "\<And>x. F_fun x = False" and F_False: "F = (\<lambda>_. False)"
+begin
+lemma F_impl_is_False:
+  "F_impl = (\<lambda>_. return False)"
+  unfolding F_impl_def F_fun_False by simp
+
+lemma deadlock_checker_correct:
+  "(uncurry0 deadlock_checker, uncurry0 (Refine_Basic.RETURN (deadlock (l0, \<lambda>_. 0))))
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+  unfolding
+    deadlock_checker_def Let_def F_impl_is_False[symmetric]
+    check_passed_impl_start_def[symmetric] deadlock_zero_clock_val_iff[symmetric]
+  using check_passed_impl_start_hnr[OF F_False] .
+
+lemmas deadlock_checker_hoare = deadlock_checker_correct[to_hnr, unfolded hn_refine_def, simplified]
+
+end
+
+end
+
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+text \<open>
+  item @{term equiv.defs.states'}
+  item @{term EA} is @{term equiv.state_ta},
+    the state automaton constructed from UPPAAL network (\<open>equiv\<close>/\<open>N\<close>)
+  item @{term A} is @{term equiv.defs.prod_ta},
+    the product automaton constructed from UPPAAL network (\<open>equiv\<close>/\<open>N\<close>)
+\<close>
+
+context
+  fixes \<phi>
+  assumes formula_is_false: "formula = formula.EX (and \<phi> (not \<phi>))"
+begin
+
+lemma F_is_FalseI:
+  shows "PR_CONST (\<lambda>(x, y). F x y) = (\<lambda>_. False)"
+  unfolding F_def by (subst formula_is_false) auto
+
+lemma final_fun_is_False:
+  "final_fun a = False"
+  unfolding final_fun_def by (subst formula_is_false) auto
+
+lemmas deadlock_checker_hoare = impl.deadlock_checker_hoare[
+    OF final_fun_is_False F_is_FalseI, folded deadlock_start_iff has_deadlock_def]
+
+end
+
+schematic_goal deadlock_checker_alt_def:
+  "impl.deadlock_checker \<equiv> ?impl"
+  unfolding impl.deadlock_checker_def impl.succs_impl_def
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding k_impl_alt_def
+  unfolding impl.check_deadlock_neg_impl_def impl.check_deadlock_impl_def
+  unfolding impl.is_start_in_states_impl_def
+  apply (abstract_let k_i ceiling)
+  apply (abstract_let "inv_fun :: (_ \<times> int list \<Rightarrow> _)" inv)
+  apply (abstract_let "trans_fun" trans)
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+schematic_goal deadlock_checker_alt_def_refined:
+  "impl.deadlock_checker \<equiv> ?impl"
+  unfolding deadlock_checker_alt_def
+  unfolding fw_impl'_int
+  unfolding inv_fun_def trans_fun_def trans_s_fun_def trans_i_fun_def
+  unfolding trans_i_from_impl
+  unfolding runf_impl runt_impl check_g_impl pairs_by_action_impl check_pred_impl
+  apply (abstract_let "IArray (map IArray inv)" inv_array)
+  apply (abstract_let "IArray (map IArray trans_out_map)" trans_out)
+  apply (abstract_let "IArray (map IArray trans_in_map)" trans_in)
+  apply (abstract_let "IArray (map IArray trans_i_map)" trans_internal)
+  apply (abstract_let "IArray bounds" bounds_array)
+  apply (abstract_let PF PF)
+  apply (abstract_let PT PT)
+  unfolding PF_alt_def PT_alt_def
+  apply (abstract_let PROG' PROG')
+  unfolding PROG'_def
+  apply (abstract_let "length prog" len_prog)
+  apply (abstract_let "IArray (map (map_option stripf) prog)" prog_f)
+  apply (abstract_let "IArray (map (map_option stript) prog)" prog_t)
+  apply (abstract_let "IArray prog" prog_array)
+  unfolding all_actions_by_state_impl
+  apply (abstract_let "[0..<p]" p_ran)
+  apply (abstract_let "[0..<na]" num_actions_ran)
+  apply (abstract_let "{0..<p}" num_processes_ran)
+  apply (abstract_let "[0..<m+1]" num_clocks_ran)
+  by (rule Pure.reflexive)
+
+end
+
+concrete_definition deadlock_checker uses
+  UPPAAL_Reachability_Problem_precompiled'.deadlock_checker_alt_def_refined
+
+definition
+  "precond_dc
+    num_processes num_clocks clock_ceiling max_steps I T prog bounds program s0 num_actions
+  \<equiv>
+    if UPPAAL_Reachability_Problem_precompiled'
+      num_processes num_clocks max_steps I T prog bounds program s0 num_actions clock_ceiling
+    then
+      deadlock_checker
+        num_processes num_clocks max_steps I T prog bounds program s0 num_actions clock_ceiling
+      \<bind> (\<lambda>x. return (Some x))
+    else return None"
+
+theorem deadlock_check:
+  "<emp>
+      precond_dc p m k max_steps I T prog bounds P s0 na
+   <\<lambda> Some r \<Rightarrow> \<up>(
+       UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k \<and>
+       r = has_deadlock (conv (N p I P T prog bounds)) max_steps (repeat 0 p, s0, \<lambda>_ . 0))
+    | None \<Rightarrow> \<up>(\<not> UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k)
+   >t"
+proof -
+  define A where "A \<equiv> conv (N p I P T prog bounds)"
+  note [sep_heap_rules] = UPPAAL_Reachability_Problem_precompiled'.deadlock_checker_hoare[
+      OF _ HOL.refl,
+      of p m max_steps I T prog bounds P s0 na k,
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      folded A_def
+      ]
+  show ?thesis
+    unfolding A_def[symmetric] precond_dc_def
+    by (sep_auto simp:
+        deadlock_checker.refine[symmetric] UPPAAL_Reachability_Problem_precompiled_defs.init_def 
+        mod_star_conv has_deadlock_def)
+qed
+
+export_code precond_dc checking SML
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Deadlock/Deadlock_Impl.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Deadlock/Deadlock_Impl.thy
@@ -0,0 +1,1403 @@
+theory Deadlock_Impl
+  imports Deadlock Munta_Base.Abstract_Term
+begin
+
+paragraph \<open>Misc\<close>
+
+lemma constraint_clk_conv_ac:
+  "constraint_clk (conv_ac ac) = constraint_clk ac"
+  by (cases ac; auto)
+
+lemma constraint_clk_conv_cc:
+  "collect_clks (conv_cc cc) = collect_clks cc"
+  by (auto simp: collect_clks_def constraint_clk_conv_ac image_def)
+
+lemma atLeastLessThan_alt_def:
+  "{a..<b} = {k. a \<le> k \<and> k < b}"
+  by auto
+
+lemma atLeastLessThan_Suc_alt_def:
+  "{a..<Suc b} = {k. a \<le> k \<and> k \<le> b}"
+  by auto
+
+lemma (in Graph_Defs) deadlock_if_deadlocked:
+  "deadlock y" if "deadlocked y"
+  using that unfolding deadlock_def by auto
+
+
+
+subsection \<open>Functional Refinement\<close>
+
+paragraph \<open>Elementary list operations\<close>
+
+lemma map_conv_rev_fold:
+  "map f xs = rev (fold (\<lambda> a xs. f a # xs) xs [])"
+proof -
+  have "fold (\<lambda> a xs. f a # xs) xs ys = rev (map f xs) @ ys" for ys
+    by (induction xs arbitrary: ys) auto
+  then show ?thesis
+    by simp
+qed
+
+lemma concat_map_conv_rev_fold:
+  "concat (map f xs) = rev (fold (\<lambda> xs ys. rev (f xs) @ ys) xs [])"
+proof -
+  have "rev (fold (\<lambda> xs ys. rev (f xs) @ ys) xs ys) = rev ys @ List.maps f xs" for ys
+    by (induction xs arbitrary: ys) (auto simp: maps_simps)
+  then show ?thesis
+    by (simp add: concat_map_maps)
+qed
+
+lemma concat_conv_fold_rev:
+  "concat xss = fold (@) (rev xss) []"
+  using fold_append_concat_rev[of "rev xss"] by simp
+
+lemma filter_conv_rev_fold:
+  "filter P xs = rev (fold (\<lambda>x xs. if P x then x # xs else xs) xs [])"
+proof -
+  have "rev ys @ filter P xs = rev (fold (\<lambda>x xs. if P x then x # xs else xs) xs ys)" for ys
+    by (induction xs arbitrary: ys) (auto, metis revg.simps(2) revg_fun)
+  from this[symmetric] show ?thesis
+    by simp
+qed
+
+paragraph \<open>DBM operations\<close>
+
+definition
+  "free_upd1 n M c =
+  (let
+    M1 = fold (\<lambda>i M. if i \<noteq> c then (M((i, c) := op_mtx_get M (i, 0))) else M) [0..<Suc n] M;
+    M2 = fold (\<lambda>i M. if i \<noteq> c then M((c,i) := \<infinity>)         else M) [0..<Suc n] M1
+  in
+    M2
+  )
+  "
+
+definition
+  "pre_reset_upd1 n M x \<equiv> free_upd1 n (restrict_zero_upd n M x) x"
+
+definition
+  "pre_reset_list_upd1 n M r \<equiv> fold (\<lambda> x M. pre_reset_upd1 n M x) r M"
+
+definition
+  "upd_pairs xs = fold (\<lambda>(p,q) f. f(p:=q)) xs"
+
+lemma upd_pairs_Nil:
+  "upd_pairs [] f = f"
+  unfolding upd_pairs_def by simp
+
+lemma upd_pairs_Cons:
+  "upd_pairs ((p, q) # xs) f = upd_pairs xs (f(p:=q))"
+  unfolding upd_pairs_def by simp
+
+lemma upd_pairs_no_upd:
+  assumes "p \<notin> fst ` set xs"
+  shows "upd_pairs xs f p = f p"
+  using assms by (induction xs arbitrary: f) (auto simp: upd_pairs_Nil upd_pairs_Cons)
+
+lemma upd_pairs_upd:
+  assumes "(p, y) \<in> set xs" "distinct (map fst xs)"
+  shows "upd_pairs xs f p = y"
+  using assms proof (induction xs arbitrary: f)
+  case Nil
+  then show ?case
+    by (simp add: upd_pairs_Nil)
+next
+  case (Cons x xs)
+  then show ?case
+    by (cases x) (auto simp add: upd_pairs_Cons upd_pairs_no_upd)
+qed
+
+lemma upd_pairs_append:
+  "upd_pairs (xs @ ys) = upd_pairs ys o upd_pairs xs"
+  unfolding upd_pairs_def fold_append ..
+
+lemma upd_pairs_commute_single:
+  "upd_pairs xs (f(a := b)) = (upd_pairs xs f)(a := b)" if "a \<notin> fst ` set xs"
+  using that by (induction xs arbitrary: f) (auto simp: upd_pairs_Nil upd_pairs_Cons fun_upd_twist)
+
+lemma upd_pairs_append':
+  "upd_pairs (xs @ ys) = upd_pairs xs o upd_pairs ys" if "fst ` set xs \<inter> fst ` set ys = {}"
+  using that
+proof (induction xs)
+  case Nil
+  then show ?case
+    by (simp add: upd_pairs_Nil)
+next
+  case (Cons a xs)
+  then show ?case
+    by (cases a) (auto simp add: upd_pairs_Nil upd_pairs_Cons upd_pairs_commute_single)
+qed
+
+definition
+  "upd_pairs' xs = fold (\<lambda>(p,q) f. f(p:=f q)) xs"
+
+lemma upd_pairs'_Nil:
+  "upd_pairs' [] f = f"
+  unfolding upd_pairs'_def by simp
+
+lemma upd_pairs'_Cons:
+  "upd_pairs' ((p, q) # xs) f = upd_pairs' xs (f(p:=f q))"
+  unfolding upd_pairs'_def by simp
+
+lemma upd_pairs'_upd_pairs:
+  assumes "fst ` set xs \<inter> snd ` set xs = {}"
+  shows   "upd_pairs' xs f = upd_pairs (map (\<lambda>(p, q). (p, f q)) xs) f"
+  using assms
+proof (induction xs arbitrary: f)
+  case Nil
+  then show ?case
+    by (simp add: upd_pairs_Nil upd_pairs'_Nil)
+next
+  case (Cons x xs)
+  obtain a b where [simp]: "x = (a, b)"
+    by (cases x)
+  from Cons.prems have *:
+    "map (\<lambda>(p, q). (p, if q = a then f b else f q)) xs = map (\<lambda>(p, q). (p, f q)) xs"
+    by auto
+  from Cons show ?case
+    by (auto simp add: upd_pairs_Cons upd_pairs'_Cons *)
+qed
+
+lemma upd_pairs_map:
+  "upd_pairs (map f xs) = fold (\<lambda>pq g. let (p,q) = f pq in g(p:=q)) xs"
+  unfolding upd_pairs_def fold_map
+  by (intro arg_cong2[where f = fold] ext) (auto simp: comp_def split: prod.split)
+
+lemma down_upd_alt_def:
+  "down_upd n M =
+  upd_pairs ([((0, j), fold min [M (k, j). k \<leftarrow> [1..<Suc n]] (Le 0)). j \<leftarrow> [1..<Suc n]]) M"
+proof -
+  define xs where
+    "xs = [((0::nat, j), Min ({Le 0} \<union> {M (k, j) | k. 1 \<le> k \<and> k \<le> n})). j \<leftarrow> [1..<Suc n]]"
+  have "distinct (map fst xs)"
+    unfolding xs_def by (auto simp add: map_concat comp_def distinct_map)
+  have *: "fold min [M (k, j). k\<leftarrow>[1..<Suc n]] (Le 0) = Min ({Le 0} \<union> {M (k, j) |k. 1 \<le> k \<and> k \<le> n})"
+    for j
+    by (subst Min.set_eq_fold[symmetric])
+       (auto simp: atLeastLessThan_Suc_alt_def simp del: upt_Suc intro: arg_cong[where f = Min])
+  show ?thesis
+    unfolding * xs_def[symmetric]
+  by (intro ext, auto simp add: down_upd_def)
+     (solves \<open>subst upd_pairs_upd[OF _ \<open>distinct _\<close>] upd_pairs_no_upd,
+      auto simp: image_def xs_def\<close>)+
+qed
+
+lemma down_upd_alt_def1:
+  "down_upd n M =
+  fold (\<lambda>j M. let (p,q) = ((0, j), fold min [M (k, j). k \<leftarrow> [1..<Suc n]] (Le 0)) in M(p:=q))
+  [1..<Suc n] M"
+proof -
+  have *: "
+    fold (\<lambda>j M.  let (p,q) = ((0,j), fold min [M (k, j). k \<leftarrow> [1..<Suc n]] (Le 0)) in M(p:=q))  xs M
+  = fold (\<lambda>j M'. let (p,q) = ((0,j), fold min [M (k, j). k \<leftarrow> [1..<Suc n]] (Le 0)) in M'(p:=q)) xs M
+  " for xs
+  proof (induction xs arbitrary: M)
+    case Nil
+    then show ?case
+      by simp
+  next
+    case (Cons x xs)
+    then show ?case
+      by (auto 4 7 intro: arg_cong2[where f = min] fold_cong)
+  qed
+  then show ?thesis
+    unfolding down_upd_alt_def upd_pairs_map ..
+qed
+
+lemma
+  "free_upd n M c =
+  upd_pairs ([((c,i), \<infinity>). i \<leftarrow> [0..<Suc n], i \<noteq> c] @ [((i,c), M(i,0)). i \<leftarrow> [0..<Suc n], i \<noteq> c]) M"
+  if "c \<le> n"
+proof -
+  let ?xs1 = "\<lambda>n. [((c,i), \<infinity>). i \<leftarrow> [0..<Suc n], i \<noteq> c]"
+  let ?xs2 = "\<lambda>n. [((i,c), M(i,0)). i \<leftarrow> [0..<Suc n], i \<noteq> c]"
+  define xs where "xs = ?xs1 n @ ?xs2 n"
+  have "distinct (map fst xs)"
+    unfolding xs_def
+    apply (clarsimp simp del: upt_Suc simp add: map_concat comp_def if_distrib split: if_split)
+     apply (auto split: if_split_asm simp: distinct_map inj_on_def intro!: distinct_concat)
+    done
+  from \<open>c \<le> n\<close> show ?thesis
+    unfolding xs_def[symmetric]
+  by (intro ext, auto simp: free_upd_def)
+     (solves \<open>subst upd_pairs_upd[OF _ \<open>distinct (map fst xs)\<close>] upd_pairs_no_upd,
+      auto simp: xs_def\<close>)+
+qed
+
+lemma free_upd_alt_def1:
+  "free_upd n M c = (let
+    M1 = upd_pairs' ([((i,c), (i,0)). i \<leftarrow> [0..<Suc n], i \<noteq> c]) M;
+    M2 = upd_pairs ([((c,i), \<infinity>). i \<leftarrow> [0..<Suc n], i \<noteq> c]) M1
+  in
+    M2
+  )" (is "_ = ?r")
+  if "0 < c" "c \<le> n"
+proof -
+  let ?xs1 = "\<lambda>n. [((c,i), \<infinity>). i \<leftarrow> [0..<Suc n], i \<noteq> c]"
+  let ?xs2 = "\<lambda>n. [((i,c), M(i,0)). i \<leftarrow> [0..<Suc n], i \<noteq> c]"
+  define xs where "xs = ?xs1 n @ ?xs2 n"
+  let ?t = "upd_pairs xs M"
+  have "distinct (map fst xs)"
+    unfolding xs_def
+    apply (clarsimp simp del: upt_Suc simp add: map_concat comp_def if_distrib split: if_split)
+     apply (auto split: if_split_asm simp: distinct_map inj_on_def intro!: distinct_concat)
+    done
+  from \<open>c \<le> n\<close> have "free_upd n M c = ?t"
+    by (intro ext, auto simp: free_upd_def)
+      (solves \<open>subst upd_pairs_upd[OF _ \<open>distinct _\<close>] upd_pairs_no_upd, auto simp: xs_def\<close>)+
+  also have "\<dots> = ?r"
+    unfolding xs_def
+    apply (subst upd_pairs_append')
+    subgoal
+      by auto
+    apply (subst upd_pairs'_upd_pairs)
+    subgoal
+      using \<open>c > 0\<close> by auto
+    apply (simp del: upt_Suc add: map_concat)
+    apply (intro arg_cong2[where f = upd_pairs] arg_cong[where f = concat])
+    by (simp del: upt_Suc)+
+  finally show ?thesis .
+qed
+
+lemma free_upd_free_upd1:
+  "free_upd n M c = free_upd1 n M c" if "c > 0" "c \<le> n"
+proof -
+  let ?x1 = "\<lambda>xs. upd_pairs' ([((i,c), (i,0)). i \<leftarrow> xs, i \<noteq> c]) M"
+  let ?x2 = "\<lambda>xs. fold (\<lambda>i M. if i \<noteq> c then (M((i, c) := M(i, 0))) else M) xs M"
+  have *: "?x1 xs = ?x2 xs" for xs
+    by (induction xs arbitrary: M) (auto simp: upd_pairs'_Nil upd_pairs'_Cons)
+  let ?y1 = "\<lambda>xs. upd_pairs ([((c,i), \<infinity>). i \<leftarrow> xs, i \<noteq> c])"
+  let ?y2 = "fold (\<lambda>i M. if i \<noteq> c then M((c,i) := \<infinity>) else M)"
+  have **: "?y1 xs M = ?y2 xs M" for xs M
+    by (induction xs arbitrary: M) (auto simp: upd_pairs_Nil upd_pairs_Cons)
+  show ?thesis
+    unfolding free_upd_alt_def1[OF that] free_upd1_def op_mtx_get_def * ** ..
+qed
+
+lemma free_upd_alt_def:
+  "free_upd n M c =
+    fold (\<lambda>i M. if i \<noteq> c then (M((c,i) := \<infinity>, (i, c) := M(i, 0))) else M) [0..<Suc n] M"
+  if "c \<le> n"
+  oops
+
+lemma pre_reset_upd_pre_reset_upd1:
+  "pre_reset_upd n M c = pre_reset_upd1 n M c" if "c > 0" "c \<le> n"
+  unfolding pre_reset_upd_def pre_reset_upd1_def free_upd_free_upd1[OF that] ..
+
+lemma pre_reset_list_upd_pre_reset_list_upd1:
+  "pre_reset_list_upd n M cs = pre_reset_list_upd1 n M cs" if "\<forall>c \<in> set cs. 0 < c \<and> c \<le> n"
+  unfolding pre_reset_list_upd_def pre_reset_list_upd1_def
+  using that by (intro fold_cong; simp add: pre_reset_upd_pre_reset_upd1)
+
+lemma pre_reset_list_transfer'[transfer_rule]:
+  includes lifting_syntax shows
+  "(eq_onp (\<lambda>x. x = n) ===> RI2 n ===> list_all2 (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ===> RI2 n)
+    pre_reset_list pre_reset_list_upd1"
+  unfolding rel_fun_def
+  apply safe
+  apply (subst pre_reset_list_upd_pre_reset_list_upd1[symmetric])
+  subgoal
+    unfolding list.rel_eq_onp by (auto simp: eq_onp_def list_all_iff)
+  by (rule pre_reset_list_transfer[unfolded rel_fun_def, rule_format])
+     (auto simp: eq_onp_def list_all2_iff)
+
+
+
+subsection \<open>Imperative Refinement\<close>
+
+context
+  fixes n :: nat
+  notes [id_rules] = itypeI[of n "TYPE (nat)"]
+    and [sepref_import_param] = IdI[of n]
+begin
+
+interpretation DBM_Impl n .
+
+sepref_definition down_impl is
+  "RETURN o PR_CONST (down_upd n)" :: "mtx_assnd \<rightarrow>a mtx_assn"
+  unfolding down_upd_alt_def1 upd_pairs_map PR_CONST_def
+  unfolding Let_def prod.case
+  unfolding fold_map comp_def
+  unfolding neutral[symmetric]
+  by sepref
+
+context
+  notes [map_type_eqs] = map_type_eqI[of "TYPE(nat * nat \<Rightarrow> 'e)" "TYPE('e i_mtx)"]
+begin
+
+sepref_register
+  abstr_upd "FW'' n" "Normalized_Zone_Semantics_Impl_Semantic_Refinement.repair_pair n"
+  and_entry_upd "free_upd1 n" "restrict_zero_upd n" "pre_reset_upd1 n"
+
+end
+
+lemmas [sepref_fr_rules] = abstr_upd_impl.refine fw_impl_refine_FW''
+
+sepref_definition abstr_FW_impl is
+  "uncurry (RETURN oo PR_CONST (abstr_FW_upd n))" ::
+  "(list_assn (acconstraint_assn clock_assn id_assn))k *a mtx_assnd \<rightarrow>a mtx_assn"
+  unfolding abstr_FW_upd_def FW''_def[symmetric] PR_CONST_def by sepref
+
+sepref_definition free_impl is
+  "uncurry (RETURN oo PR_CONST (free_upd1 n))" ::
+  "[\<lambda>(_, i). i\<le>n]a mtx_assnd *a nat_assnk \<rightarrow> mtx_assn"
+  unfolding free_upd1_def op_mtx_set_def[symmetric] PR_CONST_def by sepref
+
+sepref_definition and_entry_impl is
+  "uncurry2 (uncurry (\<lambda>x. RETURN ooo and_entry_upd x))" ::
+  "[\<lambda>(((i, j),_),_). i\<le>n \<and> j \<le> n]a nat_assnk *a nat_assnk *a id_assnk *a mtx_assnd \<rightarrow> mtx_assn"
+  unfolding and_entry_upd_def by sepref
+
+lemmas [sepref_fr_rules] = and_entry_impl.refine
+
+sepref_definition restrict_zero_impl is
+  "uncurry (RETURN oo PR_CONST (restrict_zero_upd n))" ::
+  "[\<lambda>(_, i). i\<le>n]a mtx_assnd *a nat_assnk \<rightarrow> mtx_assn"
+  unfolding restrict_zero_upd_def PR_CONST_def by sepref
+
+lemmas [sepref_fr_rules] = free_impl.refine restrict_zero_impl.refine
+
+sepref_definition pre_reset_impl is
+  "uncurry (RETURN oo PR_CONST (pre_reset_upd1 n))" ::
+  "[\<lambda>(_, i). i\<le>n]a mtx_assnd *a (clock_assn)k \<rightarrow> mtx_assn"
+  unfolding pre_reset_upd1_def PR_CONST_def by sepref
+
+lemmas [sepref_fr_rules] = pre_reset_impl.refine
+
+sepref_definition pre_reset_list_impl is
+  "uncurry (RETURN oo PR_CONST (pre_reset_list_upd1 n))" ::
+  "mtx_assnd *a (list_assn (clock_assn))k \<rightarrow>a mtx_assn"
+  unfolding pre_reset_list_upd1_def PR_CONST_def by sepref
+
+sepref_definition and_entry_repair_impl is
+  "uncurry2 (uncurry (\<lambda>x. RETURN ooo PR_CONST (and_entry_repair_upd n) x))" ::
+  "[\<lambda>(((i, j),_),_). i\<le>n \<and> j \<le> n]a nat_assnk *a nat_assnk *a id_assnk *a mtx_assnd \<rightarrow> mtx_assn"
+  unfolding and_entry_repair_upd_def PR_CONST_def by sepref
+
+private definition
+  "V_dbm'' = V_dbm' n"
+
+text \<open>We use @{term V_dbm''} because we cannot register @{term V_dbm'} twice with the refinement
+  framework: we view it as a function first, and as a DBM later.\<close>
+
+lemma V_dbm'_alt_def:
+  "V_dbm' n = op_amtx_new (Suc n) (Suc n) (V_dbm'')"
+  unfolding V_dbm''_def by simp
+
+notation fun_rel_syn (infixr "\<rightarrow>" 60)
+
+text \<open>We need the custom rule here because @{term V_dbm'} is a higher-order constant\<close>
+lemma [sepref_fr_rules]:
+  "(uncurry0 (return V_dbm''), uncurry0 (RETURN (PR_CONST (V_dbm''))))
+  \<in> unit_assnk \<rightarrow>a pure (nat_rel \<times>r nat_rel \<rightarrow> Id)"
+  by sepref_to_hoare sep_auto
+
+sepref_register "V_dbm'' :: nat \<times> nat \<Rightarrow> _ DBMEntry"
+
+text \<open>Necessary to solve side conditions of @{term op_amtx_new}\<close>
+lemma V_dbm''_bounded:
+  "mtx_nonzero V_dbm'' \<subseteq> {0..<Suc n} \<times> {0..<Suc n}"
+  unfolding mtx_nonzero_def V_dbm''_def V_dbm'_def neutral by auto
+
+text \<open>We need to pre-process the lemmas due to a failure of \<open>TRADE\<close>\<close>
+lemma V_dbm''_bounded_1:
+  "(a, b) \<in> mtx_nonzero V_dbm'' \<Longrightarrow> a < Suc n"
+  using V_dbm''_bounded by auto
+
+lemma V_dbm''_bounded_2:
+  "(a, b) \<in> mtx_nonzero V_dbm'' \<Longrightarrow> b < Suc n"
+  using V_dbm''_bounded by auto
+
+lemmas [sepref_opt_simps] = V_dbm''_def
+
+sepref_definition V_dbm_impl is
+  "uncurry0 (RETURN (PR_CONST (V_dbm' n)))" :: "unit_assnk \<rightarrow>a mtx_assn"
+  supply V_dbm''_bounded_1[simp] V_dbm''_bounded_2[simp]
+  using V_dbm''_bounded
+  apply (subst V_dbm'_alt_def)
+  unfolding PR_CONST_def by sepref
+
+text \<open>This form of 'type casting' is used to assert a bound on the natural number.\<close>
+private definition make_clock :: "nat \<Rightarrow> nat" where [sepref_opt_simps]:
+  "make_clock x = x"
+
+lemma make_clock[sepref_import_param]:
+  "(make_clock, make_clock) \<in> [\<lambda>i. i \<le> n]f nat_rel \<rightarrow> nbn_rel (Suc n)"
+  unfolding make_clock_def by (rule frefI) simp
+
+private definition "get_entries m =
+  [(make_clock i, make_clock j).
+    i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> i \<le> n \<and> j \<le> n \<and> m (i, j) \<noteq> \<infinity>]"
+
+private lemma get_entries_alt_def:
+  "get_entries m = [(make_clock i, make_clock j).
+    i\<leftarrow>[0..<Suc n], j\<leftarrow>[0..<Suc n], (i > 0 \<or> j > 0) \<and> m (i, j) \<noteq> \<infinity>]"
+  unfolding get_entries_def by (intro arg_cong[where f = concat] map_cong) auto
+
+private definition
+  "upd_entry i j M m = and_entry_repair_upd n j i (neg_dbm_entry (m (i, j))) (op_mtx_copy M)"
+
+private definition
+  "upd_entries i j m = map (\<lambda> M. upd_entry i j M m)"
+
+context
+  notes [map_type_eqs] = map_type_eqI[of "TYPE(nat * nat \<Rightarrow> 'e)" "TYPE('e i_mtx)"]
+begin
+
+sepref_register
+  "dbm_minus_canonical_check_upd n"
+  "dbm_subset_fed_upd n" "abstr_FW_upd n" "pre_reset_list_upd1 n" "V_dbm' n" "down_upd n"
+  upd_entry upd_entries get_entries "and_entry_repair_upd n"
+
+end
+
+lemma [sepref_import_param]: "(neg_dbm_entry,neg_dbm_entry) \<in> Id\<rightarrow>Id" by simp
+
+lemmas [sepref_fr_rules] = and_entry_repair_impl.refine
+
+sepref_definition upd_entry_impl is
+  "uncurry2 (uncurry (\<lambda>x. RETURN ooo PR_CONST upd_entry x))" ::
+  "[\<lambda>(((i, j),_),_). i\<le>n \<and> j \<le> n]a
+    nat_assnk *a nat_assnk *a mtx_assnk *a mtx_assnk \<rightarrow> mtx_assn"
+  unfolding upd_entry_def PR_CONST_def by sepref
+
+text \<open>This is to ensure that the refinement initially infers the right 'type' for list arguments.\<close>
+lemma [intf_of_assn]:
+  "intf_of_assn AA TYPE('aa) \<Longrightarrow> intf_of_assn (list_assn(AA)) TYPE('aa list)"
+  by (rule intf_of_assnI)
+
+lemmas [sepref_fr_rules] = upd_entry_impl.refine
+
+sepref_definition upd_entries_impl is
+  "uncurry2 (uncurry (\<lambda>x. RETURN ooo PR_CONST upd_entries x))" ::
+  "[\<lambda>(((i, j),_),_). i\<le>n \<and> j \<le> n]a
+    nat_assnk *a nat_assnk *a mtx_assnk *a (list_assn mtx_assn)k \<rightarrow> list_assn mtx_assn"
+  unfolding upd_entries_def PR_CONST_def
+  unfolding map_conv_rev_fold rev_conv_fold
+  supply [sepref_fr_rules] = HOL_list_empty_hnr_aux
+  by sepref
+
+lemma [sepref_import_param]: "((=), (=)) \<in> Id\<rightarrow>Id\<rightarrow>Id" by simp
+
+sepref_definition get_entries_impl is
+  "RETURN o PR_CONST get_entries" ::
+  "mtx_assnk \<rightarrow>a list_assn ((clock_assn) \<times>a (clock_assn))"
+  unfolding get_entries_alt_def PR_CONST_def
+  unfolding map_conv_rev_fold
+  unfolding concat_conv_fold_rev
+  supply [sepref_fr_rules] = HOL_list_empty_hnr_aux
+  by sepref
+
+lemmas [sepref_fr_rules] = upd_entries_impl.refine get_entries_impl.refine
+
+private lemma dbm_minus_canonical_upd_alt_def:
+  "dbm_minus_canonical_upd n xs m =
+  concat (map (\<lambda>ij. map (\<lambda> M.
+      and_entry_repair_upd n (snd ij) (fst ij) (neg_dbm_entry (m (fst ij, snd ij))) (op_mtx_copy M))
+    xs) (get_entries m))"
+  unfolding dbm_minus_canonical_upd_def op_mtx_copy_def get_entries_def split_beta make_clock_def
+  by simp
+
+sepref_definition dbm_minus_canonical_impl is
+  "uncurry (RETURN oo PR_CONST (dbm_minus_canonical_check_upd n))" ::
+  "(list_assn mtx_assn)k *a mtx_assnk \<rightarrow>a list_assn mtx_assn"
+  unfolding dbm_minus_canonical_check_upd_def dbm_minus_canonical_upd_alt_def PR_CONST_def
+  unfolding upd_entry_def[symmetric] upd_entries_def[symmetric]
+  unfolding filter_conv_rev_fold concat_map_conv_rev_fold rev_conv_fold
+  supply [sepref_fr_rules] = HOL_list_empty_hnr_aux
+  by sepref
+
+lemmas [sepref_fr_rules] = dbm_minus_canonical_impl.refine
+
+sepref_definition dbm_subset_fed_impl is
+  "uncurry (RETURN oo PR_CONST (dbm_subset_fed_upd n))" ::
+  "mtx_assnd *a (list_assn mtx_assn)d \<rightarrow>a bool_assn"
+  unfolding dbm_subset_fed_upd_alt_def PR_CONST_def
+  unfolding list_ex_foldli filter_conv_rev_fold
+  unfolding short_circuit_conv
+  supply [sepref_fr_rules] = HOL_list_empty_hnr_aux
+  by sepref
+
+end (* Fixed n *)
+
+
+subsection \<open>Implementation for a Concrete Automaton\<close>
+
+context TA_Impl
+begin
+
+sublocale TA_Impl_Op
+  where loc_rel = loc_rel and f = "PR_CONST E_op''" and op_impl = E_op''_impl
+  unfolding PR_CONST_def by standard (rule E_op''_impl.refine)
+
+end
+
+context TA_Impl
+begin
+
+definition
+  "check_deadlock_dbm l M = dbm_subset_fed_upd n M (
+   [down_upd n (abstr_FW_upd n (inv_of_A l) (abstr_FW_upd n g
+      (pre_reset_list_upd1 n (abstr_FW_upd n (inv_of_A l') (V_dbm' n)) r)
+    )). (g, a, r, l') \<leftarrow> trans_fun l
+   ]
+  )"
+
+abbreviation zone_of ("\<lbrakk>_\<rbrakk>") where "zone_of M \<equiv> [curry (conv_M M)]v,n"
+
+theorem check_deadlock_dbm_correct:
+  "TA.check_deadlock l \<lbrakk>M\<rbrakk> = check_deadlock_dbm l M" if
+  "\<lbrakk>M\<rbrakk> \<subseteq> V" "l \<in> states'" "Deadlock.canonical' n (curry (conv_M M))"
+proof -
+  text \<open>0. Setup \<open>&\<close> auxiliary facts\<close>
+  include lifting_syntax
+  note [simp del] = And.simps abstr.simps (* this should better be definitions *)
+  have inv_of_simp: "inv_of (conv_A A) l = conv_cc (inv_of A l)" for l
+    using inv_fun unfolding inv_rel_def b_rel_def fun_rel_def conv_A_def
+    by (force split: prod.split simp: inv_of_def)
+
+  have trans_funD: "l' \<in> states"
+    "collect_clks (inv_of A l) \<subseteq> clk_set A" "collect_clks (inv_of A l') \<subseteq> clk_set A"
+    "collect_clks g \<subseteq> clk_set A" "set r \<subseteq> clk_set A"
+    if "(g, a, r, l') \<in> set (trans_fun l)" for g a r l'
+    subgoal
+      using \<open>l \<in> states'\<close> that trans_impl_states by auto
+    subgoal
+      by (metis collect_clks_inv_clk_set)
+    subgoal
+      by (metis collect_clks_inv_clk_set)
+    subgoal
+      using that \<open>l \<in> _\<close> by (intro collect_clocks_clk_set trans_impl_trans_of)
+    subgoal
+      using that \<open>l \<in> _\<close> by (intro reset_clk_set trans_impl_trans_of)
+    done
+
+  text \<open>1. Transfer to most abstract DBM operations\<close>
+  have [transfer_rule]:
+    "(eq_onp (\<lambda> (g, a, r, l'). (g, a, r, l') \<in> set (trans_fun l)) ===> RI2 n)
+      (\<lambda>(g, a, r, l'). down n
+          (abstr_FW n (conv_cc (inv_of A l))
+            (abstr_FW n (conv_cc g)
+              (pre_reset_list n (abstr_FW n (conv_cc (inv_of A l')) V_dbm v) r) v) v))
+      (\<lambda>(g, a, r, l'). down_upd n
+          (abstr_FW_upd n (inv_of A l)
+            (abstr_FW_upd n g (pre_reset_list_upd1 n (abstr_FW_upd n (inv_of A l') (V_dbm' n)) r))
+          ))
+    " (is "(_ ===> RI2 n) (\<lambda> (g, a, r, l'). ?f g a r l') (\<lambda> (g, a, r, l'). ?g g a r l')")
+  proof -
+    {
+    fix g a r l' assume "(g, a, r, l') \<in> set (trans_fun l)"
+    have *: "rel_acconstraint (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ri (conv_ac ac) ac"
+      if "constraint_clk ac > 0" "constraint_clk ac \<le> n" for ac
+      using that by (cases ac) (auto simp: eq_onp_def ri_def)
+    have *: "list_all2 (rel_acconstraint (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ri) (conv_cc g) g"
+      if "collect_clks g \<subseteq> clk_set A" for g
+      unfolding list_all2_map1 using that clock_range
+      by (auto simp: collect_clks_def intro!: * list.rel_refl_strong)
+    have [transfer_rule]:
+      "list_all2 (rel_acconstraint (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ri) (conv_cc g) g"
+      "list_all2 (rel_acconstraint (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ri)
+        (conv_cc (inv_of A l)) (inv_of A l)"
+      "list_all2 (rel_acconstraint (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) ri)
+        (conv_cc (inv_of A l')) (inv_of A l')"
+      by (intro * trans_funD[OF \<open>(g, a, r, l') \<in> set (trans_fun l)\<close>])+
+    have [transfer_rule]:
+      "list_all2 (eq_onp (\<lambda>x. 0 < x \<and> x < Suc n)) r r"
+      using trans_funD(5)[OF \<open>(g, a, r, l') \<in> set (trans_fun l)\<close>] clock_range
+      by (auto 4 3 dest: bspec subsetD simp: eq_onp_def list_all2_same)
+    have "RI2 n (?f g a r l') (?g g a r l')"
+      by transfer_prover
+    } then show ?thesis
+    by (intro rel_funI, clarsimp simp: eq_onp_def)
+  qed
+  have [transfer_rule]:
+    "(eq_onp (\<lambda>l'. l' = l) ===> list_all2 (eq_onp (\<lambda>(g,a,r,l'). (g,a,r,l') \<in> set (trans_fun l))))
+      trans_fun trans_fun
+    "
+    unfolding rel_fun_def eq_onp_def by (auto intro: list.rel_refl_strong)
+  note [transfer_rule] = dbm_subset_fed_transfer
+  have [transfer_rule]: "eq_onp (\<lambda>l'. l' = l) l l" "(eq_onp (\<lambda>x. x < Suc n)) n n"
+    by (auto simp: eq_onp_def)
+  have *: "
+    (dbm_subset_fed_check n (curry (conv_M M))
+         (map (\<lambda>(g, a, r, l').
+                  down n
+          (abstr_FW n (conv_cc (inv_of A l))
+            (abstr_FW n (conv_cc g)
+              (pre_reset_list n (abstr_FW n (conv_cc (inv_of A l')) V_dbm v) r) v) v))
+           (trans_fun l))) =
+    (dbm_subset_fed_upd n M
+         (map (\<lambda>(g, a, r, l').
+                  down_upd n
+          (abstr_FW_upd n (inv_of A l)
+            (abstr_FW_upd n g (pre_reset_list_upd1 n (abstr_FW_upd n (inv_of A l') (V_dbm' n)) r))
+          ))
+         (trans_fun l)))
+    "
+    by transfer_prover
+
+  text \<open>2. Semantic argument establishing equivalences between zones and DBMs\<close>
+  have **:
+    "[down n (abstr_FW n (conv_cc (inv_of A l)) (abstr_FW n (conv_cc g)
+          (pre_reset_list n (abstr_FW n (conv_cc (inv_of A l')) V_dbm v) r) v) v)]v,n
+  = (zone_set_pre ({u. u \<turnstile> inv_of (conv_A A) l'} \<inter> V) r \<inter> {u. \<forall> x \<in> set r. u x \<ge> 0}
+       \<inter> {u. u \<turnstile> conv_cc g} \<inter> {u. u \<turnstile> inv_of (conv_A A) l})\<down> \<inter> V"
+    if "(g, a, r, l') \<in> set(trans_fun l)" for g a r l'
+  proof -
+    from trans_funD[OF \<open>(g, a, r, l') \<in> set (trans_fun l)\<close>] have clock_conditions:
+      "\<forall>c\<in>collect_clks (conv_cc (inv_of A l)). 0 < c \<and> c \<le> n"
+      "\<forall>c\<in>collect_clks (conv_cc (inv_of A l')). 0 < c \<and> c \<le> n"
+      "\<forall>c\<in>collect_clks (conv_cc g). 0 < c \<and> c \<le> n"
+      "\<forall>c\<in>set r. 0 < c \<and> c \<le> n"
+      using \<open>l \<in> states'\<close> clock_range by (auto simp: constraint_clk_conv_cc)
+    have structural_conditions:
+      "abstr_FW n (conv_cc (inv_of A l')) V_dbm v 0 0 \<le> 0"
+      "\<forall>x\<in>set r. abstr_FW n (map conv_ac (inv_of A l')) V_dbm v 0 x \<le> 0"
+      subgoal
+        using abstr_FW_diag_preservation[of n V_dbm "conv_cc (inv_of A l')" v]
+        by (simp add: V_dbm_def)
+      subgoal
+        using \<open>\<forall>c\<in>set r. 0 < c \<and> c \<le> n\<close>
+        by (auto 4 3 simp: V_dbm_def intro: abstr_FW_mono order.trans)
+      done
+    note side_conditions = clock_conditions structural_conditions
+      abstr_FW_canonical[unfolded canonical'_def] abstr_FW_canonical
+    show ?thesis
+      apply (subst dbm.down_correct, rule side_conditions)
+      apply (subst dbm.abstr_FW_correct, rule side_conditions)
+      apply (subst dbm.abstr_FW_correct, rule side_conditions)
+      apply (subst dbm.pre_reset_list_correct, (rule side_conditions)+)
+      apply (subst dbm.abstr_FW_correct, (rule side_conditions)+)
+      apply (simp add: inv_of_simp dbm.V_dbm_correct atLeastLessThanSuc_atLeastAtMost Int_commute)
+      done
+  qed
+  have **:
+    "(\<Union>x\<in>set (trans_fun l).
+              dbm.zone_of
+               (case x of (g, a, r, l') \<Rightarrow>
+                  down n
+        (abstr_FW n (conv_cc (inv_of A l))
+          (abstr_FW n (conv_cc g)
+            (pre_reset_list n (abstr_FW n (conv_cc (inv_of A l')) V_dbm v) r) v) v)))
+    = (\<Union>(g,a,r,l')\<in>set (trans_fun l).
+    ((zone_set_pre ({u. u \<turnstile> inv_of (conv_A A) l'} \<inter> V) r \<inter> {u. \<forall> x \<in> set r. u x \<ge> 0}
+           \<inter> {u. u \<turnstile> conv_cc g} \<inter> {u. u \<turnstile> inv_of (conv_A A) l})\<down> \<inter> V)
+    )
+    "
+    by (auto simp: **)
+
+  text \<open>3. Putting it all together\<close>
+  have transD: "\<exists> g'. (g', a, r, l') \<in> set (trans_fun l) \<and> g = conv_cc g'"
+    if "conv_A A \<turnstile> l \<longrightarrow>g,a,r l'" for g a r l'
+    using trans_of_trans_impl[OF _ \<open>l \<in> states'\<close>] that
+    unfolding trans_of_def conv_A_def by (auto 5 0 split: prod.split_asm)
+  have transD2:
+    "conv_A A \<turnstile> l \<longrightarrow>conv_cc g,a,r l'" if "(g, a, r, l') \<in> set (trans_fun l)" for g a r l'
+    using trans_impl_trans_of[OF that \<open>l \<in> states'\<close>]
+    unfolding trans_of_def conv_A_def by (auto 4 3 split: prod.split)
+  show ?thesis
+    unfolding TA.check_deadlock_alt_def[OF \<open>_ \<subseteq> V\<close>] check_deadlock_dbm_def inv_of_A_def *[symmetric]
+    apply (subst dbm.dbm_subset_fed_check_correct[symmetric, OF that(3)])
+    apply (simp add: **)
+    apply safe
+    subgoal for x
+      by (auto 4 5 elim!: transD[elim_format] dest: subsetD)
+    subgoal for x
+      apply (drule subsetD, assumption)
+      apply clarsimp
+      subgoal for g a r l'
+        apply (inst_existentials "
+          (zone_set_pre ({u. u \<turnstile> inv_of (conv_A A) l'} \<inter> V) r \<inter> {u. \<forall>x\<in>set r. 0 \<le> u x}
+           \<inter> {u. u \<turnstile> conv_cc g} \<inter> {u. u \<turnstile> inv_of (conv_A A) l})\<down> \<inter> V" "conv_cc g" a r l')
+        apply (auto dest: transD2)
+        done
+      done
+    done
+qed
+
+lemmas [sepref_fr_rules] =
+  V_dbm_impl.refine abstr_FW_impl.refine pre_reset_list_impl.refine
+  down_impl.refine dbm_subset_fed_impl.refine
+
+sepref_definition check_deadlock_impl is
+  "uncurry (RETURN oo PR_CONST check_deadlock_dbm)" ::
+  "location_assnk *a (mtx_assn n)d \<rightarrow>a bool_assn"
+  unfolding check_deadlock_dbm_def PR_CONST_def
+  unfolding case_prod_beta
+  unfolding map_conv_rev_fold
+  apply (rewrite "HOL_list.fold_custom_empty")
+  by sepref
+
+end (* Reachability Problem Impl *)
+
+
+subsection \<open>Checking a Property on the Reachable Set\<close>
+
+locale Worklist_Map2_Impl_check =
+  Worklist_Map2_Impl_finite +
+  fixes Q :: "'a \<Rightarrow> bool"
+  fixes Qi
+  assumes Q_refine: "(Qi,RETURN o PR_CONST Q) \<in> Ad \<rightarrow>a bool_assn"
+  and F_False: "F = (\<lambda> _. False)"
+  and Q_mono: "\<And> a b. a \<preceq> b \<Longrightarrow> \<not> empty a \<Longrightarrow> reachable a \<Longrightarrow> reachable b \<Longrightarrow> Q a \<Longrightarrow> Q b"
+begin
+
+definition check_passed :: "bool nres" where
+  "check_passed = do {
+    (r, passed) \<leftarrow> pw_algo_map2;
+    ASSERT (finite (dom passed));
+    passed \<leftarrow> ran_of_map passed;
+    r \<leftarrow> nfoldli passed (\<lambda>b. \<not>b)
+      (\<lambda> passed' _.
+        do {
+          passed' \<leftarrow> SPEC (\<lambda>l. set l = passed');
+          nfoldli passed' (\<lambda>b. \<not>b)
+            (\<lambda>v' _.
+              if Q v' then RETURN True else RETURN False
+            )
+            False
+        }
+      )
+      False;
+    RETURN r
+  }"
+
+lemma check_passed_correct:
+  "check_passed \<le> SPEC (\<lambda> r. (r \<longleftrightarrow> (\<exists> a. reachable a \<and> \<not> empty a \<and> Q a)))"
+proof -
+  have [simp]: "F_reachable = False"
+    unfolding F_reachable_def F_False Search_Space_Defs.F_reachable_def by simp
+  define outer_inv where "outer_inv passed done todo r \<equiv>
+    set done \<union> set todo = ran passed \<and>
+    (\<not> r \<longrightarrow> (\<forall> S \<in> set done. \<forall> x \<in> S. \<not> Q x)) \<and>
+    (r \<longrightarrow> (\<exists> a. reachable a \<and> \<not> empty a \<and> Q a))
+  " for passed :: "'c \<Rightarrow> 'a set option" and "done" todo and r :: bool
+  define inner_inv where "inner_inv passed done todo r \<equiv>
+    set done \<union> set todo = passed \<and>
+    (\<not> r \<longrightarrow> (\<forall> x \<in> set done. \<not> Q x)) \<and>
+    (r \<longrightarrow> (\<exists> a. reachable a \<and> \<not> empty a \<and> Q a))
+  " for passed :: "'a set" and "done" todo and r :: bool
+  show ?thesis
+    supply [refine_vcg] = pw_algo_map2_correct
+    unfolding check_passed_def
+    apply (refine_rcg refine_vcg)
+    subgoal
+      by auto
+    subgoal for _ brk_false passed range_list
+      apply (rule nfoldli_rule[where I = "outer_inv passed"])
+
+(* Init: I [] range_list False *)
+      subgoal
+        unfolding outer_inv_def
+        by auto
+
+(* Inner loop *)
+        apply (refine_rcg refine_vcg)
+      subgoal for current "done" todo r xs
+        apply clarsimp
+        apply (rule nfoldli_rule[where I = "inner_inv current"])
+
+        subgoal (* Init: I [] xs False *)
+          unfolding inner_inv_def
+          by auto
+
+(* inner inv *)
+        subgoal for p x l1 l2 r'
+          by (clarsimp simp: inner_inv_def outer_inv_def map_set_rel_def)
+             (metis Sup_insert Un_insert_left insert_iff subset_Collect_conv)
+
+(* break inner \<longrightarrow> break outer *)
+        subgoal for p l1 l2 r'
+          unfolding inner_inv_def outer_inv_def
+          by (metis append.assoc append_Cons append_Nil set_append)
+
+(* finished inner \<longrightarrow> outer inv *)
+        subgoal for p r'
+          unfolding inner_inv_def outer_inv_def
+          by clarsimp
+        done
+
+(* break outer *)
+      subgoal for l1 l2 r
+        unfolding outer_inv_def
+        by auto
+
+(* outer finished *)
+      subgoal for r
+        unfolding outer_inv_def
+        apply clarsimp
+        apply (elim allE impE)
+        apply (intro conjI; assumption)
+        apply safe
+        apply (drule Q_mono, assumption+)
+        unfolding map_set_rel_def
+        by auto
+      done
+    done
+qed
+
+schematic_goal pw_algo_map2_refine:
+  "(?x, uncurry0 (PR_CONST pw_algo_map2)) \<in>
+  unit_assnk \<rightarrow>a bool_assn \<times>a hm.hms_assn' K (lso_assn A)"
+  unfolding PR_CONST_def hm.hms_assn'_id_hms_assn[symmetric] by (rule pw_algo_map2_impl.refine_raw)
+
+sepref_register pw_algo_map2
+
+sepref_register "PR_CONST Q"
+
+sepref_thm check_passed_impl is
+  "uncurry0 check_passed" :: "unit_assnk \<rightarrow>a bool_assn"
+  supply [sepref_fr_rules] = pw_algo_map2_refine ran_of_map_impl.refine lso_id_hnr Q_refine
+  using pure_K left_unique_K right_unique_K
+  unfolding check_passed_def
+  apply (rewrite in Q PR_CONST_def[symmetric])
+  unfolding hm.hms_fold_custom_empty
+  unfolding list_of_set_def[symmetric]
+  by sepref
+
+concrete_definition (in -) check_passed_impl
+  uses Worklist_Map2_Impl_check.check_passed_impl.refine_raw is "(uncurry0 ?f,_)\<in>_"
+
+lemma check_passed_impl_hnr:
+    "(uncurry0 (check_passed_impl succsi a0i Fi Lei emptyi keyi copyi tracei Qi),
+      uncurry0 (RETURN (\<exists>a. reachable a \<and> \<not> empty a \<and> Q a)))
+    \<in> unit_assnk \<rightarrow>a bool_assn"
+    using
+      check_passed_impl.refine[
+        OF Worklist_Map2_Impl_check_axioms,
+        FCOMP check_passed_correct[THEN Id_SPEC_refine, THEN nres_relI]
+        ]
+    by (simp add: RETURN_def)
+
+end
+
+
+subsection \<open>Complete Deadlock Checker\<close>
+
+paragraph \<open>Miscellaneous Properties\<close>
+
+context E_From_Op_Bisim
+begin
+
+text \<open>More general variant of @{thm E_from_op_reachability_check} *)\<close>
+theorem E_from_op_reachability_check:
+  assumes "\<And>a b. P a \<Longrightarrow> a \<sim> b \<Longrightarrow> wf_state a \<Longrightarrow> wf_state b \<Longrightarrow> P b"
+  shows
+  "(\<exists>l' D'. E** a0 (l', D') \<and> P (l', D')) \<longleftrightarrow> (\<exists>l' D'. E_from_op** a0 (l', D') \<and> P (l', D'))"
+  by (rule E_E1_steps_equiv[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state];
+     (assumption | rule assms))
+
+end
+
+context Regions_TA
+begin
+
+lemma check_deadlock_anti_mono:
+  "check_deadlock l Z" if "check_deadlock l Z'" "Z \<subseteq> Z'"
+  using that unfolding check_deadlock_def by auto
+
+lemma global_clock_numbering:
+  "global_clock_numbering A v n"
+  using valid_abstraction clock_numbering_le clock_numbering by blast
+
+text \<open>Variant of @{thm bisim} without emptiness.\<close>
+lemma bisim:
+  "Bisimulation_Invariant
+  (\<lambda> (l, Z) (l', Z'). A \<turnstile> \<langle>l, Z\<rangle> \<leadsto>\<beta> \<langle>l', Z'\<rangle>)
+  (\<lambda> (l, D) (l', D'). \<exists> a. A \<turnstile>' \<langle>l, D\<rangle> \<leadsto>\<N>(a) \<langle>l', D'\<rangle>)
+  (\<lambda> (l, Z) (l', D). l = l' \<and> Z = [D]v,n)
+  (\<lambda> _. True) (\<lambda> (l, D). valid_dbm D)"
+proof (standard, goal_cases)
+  \<comment> \<open>\<open>\<beta> \<Rightarrow> \<N>\<close>\<close>
+  case (1 a b a')
+  then show ?case
+    by (blast elim: norm_beta_complete1[OF global_clock_numbering valid_abstraction])
+next
+  \<comment> \<open>\<open>\<N> \<Rightarrow> \<beta>\<close>\<close>
+  case (2 a a' b')
+  then show ?case
+    by (blast intro: norm_beta_sound''[OF global_clock_numbering valid_abstraction])
+next
+  \<comment> \<open>\<open>\<beta>\<close> invariant\<close>
+  case (3 a b)
+  then show ?case
+    by simp
+next
+  \<comment> \<open>\<open>\<N>\<close> invariant\<close>
+  case (4 a b)
+  then show ?case
+    by (auto intro: valid_dbm_step_z_norm''[OF global_clock_numbering valid_abstraction])
+qed
+
+lemma steps_z_beta_reaches:
+  "reaches (l, Z) (l', Z')" if "A \<turnstile> \<langle>l, Z\<rangle> \<leadsto>\<beta>* \<langle>l', Z'\<rangle>" "Z' \<noteq> {}"
+  using that
+proof (induction "(l', Z')" arbitrary: l' Z' rule: rtranclp_induct)
+  case base
+  then show ?case
+    by blast
+next
+  case (step y l'' Z'')
+  obtain l' Z' where [simp]: "y = (l', Z')"
+    by force
+  from step.prems step.hyps(2) have "Z' \<noteq> {}"
+    by (clarsimp simp: step_z_beta'_empty)
+  from step.prems step.hyps(3)[OF \<open>y = _\<close> this] step.hyps(2) show ?case
+    including graph_automation_aggressive by auto
+qed
+
+lemma reaches_steps_z_beta_iff:
+  "reaches (l, Z) (l', Z') \<longleftrightarrow> A \<turnstile> \<langle>l, Z\<rangle> \<leadsto>\<beta>* \<langle>l', Z'\<rangle> \<and> Z' \<noteq> {}" if "Z \<noteq> {}"
+  apply safe
+  subgoal
+    including graph_automation_aggressive by (induction rule: rtranclp_induct) auto
+  using that by (auto dest: steps_z_beta_reaches elim: rtranclp.cases)
+
+end
+
+lemma dbm_int_init_dbm:
+  "dbm_int (curry init_dbm) n"
+  unfolding init_dbm_def by auto
+
+context TA_Start
+begin
+
+lemma wf_dbm_canonical'D:
+  "Deadlock.canonical' n (curry (conv_M D))" if "wf_dbm D"
+  using that unfolding wf_dbm_def Deadlock.canonical'_def check_diag_conv_M_iff by simp
+
+lemma subsumes_dbm_subsetD:
+  "dbm_subset n M M'" if "subsumes n (l, M) (l', M')" "\<not> check_diag n M"
+  using that by (cases "l = l'") (auto simp: subsumes_simp_1 subsumes_simp_2)
+
+lemma subsumes_loc_eqD:
+  "l = l'" if "subsumes n (l, M) (l', M')" "\<not> check_diag n M"
+  using that by (cases "l = l'") (auto simp: subsumes_simp_1 subsumes_simp_2)
+
+lemma init_dbm_zone:
+  "[curry (init_dbm :: real DBM')]v,n = {u. \<forall>c\<in>{1..n}. u c = 0}"
+  using init_dbm_semantics by blast
+
+lemma not_check_deadlock_mono:
+  "(case a of (l, M) \<Rightarrow> \<not> TA.check_deadlock l (dbm.zone_of (curry (conv_M M)))) \<Longrightarrow>
+  a \<sim> b \<Longrightarrow> wf_state a \<Longrightarrow> wf_state b \<Longrightarrow>
+  case b of (l, M) \<Rightarrow> \<not> TA.check_deadlock l (dbm.zone_of (curry (conv_M M)))"
+  unfolding TA.check_deadlock_def state_equiv_def dbm_equiv_def by auto
+
+lemma not_check_deadlock_non_empty:
+  "Z \<noteq> {}" if "\<not> TA.check_deadlock l Z"
+  using that unfolding TA.check_deadlock_def by auto
+
+end (* TA Start *)
+
+context TA_Impl
+begin
+
+lemma op_E_from_op_iff:
+  "op.E_from_op = E_op''.E_from_op"
+  unfolding PR_CONST_def ..
+
+lemmas reachable_states = reachable_states[unfolded op_E_from_op_iff]
+
+lemma E_op''_states:
+  "l' \<in> states" if "E_op''.E_from_op (l, M) (l', M')" "l \<in> states"
+  using that by (force simp: a0_def state_set_def E_op''.E_from_op_def)
+
+subsubsection \<open>Instantiating the Checker Algorithm\<close>
+
+corollary check_deadlock_dbm_correct':
+  assumes "l \<in> states'" "wf_state (l, M)"
+  shows "TA.check_deadlock l (dbm.zone_of (curry (conv_M M))) = check_deadlock_dbm l M"
+  apply (rule check_deadlock_dbm_correct)
+  using assms
+  unfolding wf_state_def Deadlock.canonical'_def prod.case
+    apply (blast dest: wf_dbm_altD)
+   apply (rule assms)
+  using assms
+  unfolding wf_state_def Deadlock.canonical'_def prod.case
+  apply -
+  apply (drule wf_dbm_altD(1))
+  unfolding canonical'_conv_M_iff check_diag_conv_M_iff by simp
+
+corollary check_deadlock_dbm_correct'':
+  assumes "E_op''.E_from_op** a0 (l, M)"
+  shows "TA.check_deadlock l (dbm.zone_of (curry (conv_M M))) = check_deadlock_dbm l M"
+  using assms
+  apply -
+  apply (rule check_deadlock_dbm_correct')
+   apply (drule reachable_states, use states'_states in auto; fail)
+  unfolding wf_state_def prod.case apply (erule E_op''.reachable_wf_dbm)
+  done
+
+lemma not_check_deadlock_non_empty:
+  "Z \<noteq> {}" if "\<not> TA.check_deadlock l Z"
+  using that unfolding TA.check_deadlock_def by auto
+
+sepref_register check_deadlock_dbm :: "'s \<Rightarrow> int DBMEntry i_mtx \<Rightarrow> bool"
+
+sepref_definition check_deadlock_neg_impl is
+  "RETURN o (\<lambda>(l, M). \<not> check_deadlock_dbm l M)" :: "(location_assn \<times>a (mtx_assn n))d \<rightarrow>a bool_assn"
+  supply [sepref_fr_rules] = check_deadlock_impl.refine
+  by sepref
+
+lemma not_check_deadlock_compatible:
+  assumes
+    "(case a of (l, Z) \<Rightarrow> \<lambda>(l', D'). l' = l \<and> dbm.zone_of (curry (conv_M D')) = Z) b"
+   "case b of (l, M) \<Rightarrow> l \<in> states' \<and> wf_state (l, M)"
+ shows
+   "(case a of (l, Z) \<Rightarrow> \<not> TA.check_deadlock l Z) = (case b of (l, M) \<Rightarrow> \<not> check_deadlock_dbm l M)"
+  using assms by (auto simp: check_deadlock_dbm_correct'[symmetric])
+
+lemma deadlock_check_diag:
+  "\<not> check_diag n M" if "\<not> check_deadlock_dbm l M" "E_op''.E_from_op** a0 (l, M)"
+  using that(1)
+  apply (subst (asm) check_deadlock_dbm_correct'[symmetric])
+  subgoal
+    using reachable_states that(2) states'_states by auto
+  subgoal
+    unfolding wf_state_def using E_op''.reachable_wf_dbm[OF that(2)] by simp
+  using canonical_check_diag_empty_iff by (blast dest: not_check_deadlock_non_empty)
+
+(* Duplication *)
+lemma norm_final_bisim:
+  "Bisimulation_Invariant
+     (\<lambda>(l, D) (l', D'). \<exists>a. step_z_norm'' (conv_A A) l D a l' D')
+     E_op''.E_from_op
+     (\<lambda> (l, M) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = [M]v,n)
+     (\<lambda>(l, y). valid_dbm y) wf_state"
+  by (rule
+    Bisimulation_Invariant_sim_replace[OF
+      Bisimulation_Invariant_composition[OF
+        step_z_norm''_step_impl'_equiv[unfolded step_impl'_E] E_op''.E_from_op_bisim
+      ]
+    ])
+    (auto simp add: state_equiv_def dbm_equiv_def)
+
+interpretation bisim:
+  Bisimulation_Invariant
+  "\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z'"
+  "\<lambda>a b. op.E_from_op a b"
+  "\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z"
+  "\<lambda>_. True" "\<lambda>(l, M). l \<in> states \<and> wf_state (l, M)"
+  unfolding op_E_from_op_iff
+  by (rule Bisimulation_Invariant_strengthen_post',
+      (rule Bisimulation_Invariant_composition[OF TA.bisim norm_final_bisim]
+        Bisimulation_Invariant_sim_replace
+        )+) (auto 4 3 dest: E_op''_states wf_dbm_D(3) simp: wf_state_def)
+
+lemmas beta_final_bisim = bisim.Bisimulation_Invariant_axioms
+
+definition
+  "is_start_in_states = (trans_fun l0 \<noteq> [])"
+
+lemma is_start_in_states:
+  "l0 \<in> Simulation_Graphs_TA.state_set A" if "is_start_in_states"
+proof -
+  from that obtain g a r l' where "(g, a, r, l') \<in> set (trans_fun l0)"
+    by (cases "hd (trans_fun l0)" rule: prod_cases4)
+       (auto dest: hd_in_set simp: is_start_in_states_def)
+  from trans_impl_trans_of[OF this] states'_states have "A \<turnstile> l0 \<longrightarrow>g,a,r l'"
+    by simp
+  then show ?thesis
+    unfolding Simulation_Graphs_TA.state_set_def trans_of_def by auto
+qed
+
+lemma deadlocked_if_not_is_start_in_states:
+  "deadlocked (l0, Z0)" if "\<not> is_start_in_states"
+proof -
+  have *: False if "A \<turnstile> l0 \<longrightarrow>g,a,r l'" for g a r l'
+    using trans_of_trans_impl[OF that] \<open>\<not> _\<close> states'_states unfolding is_start_in_states_def by auto
+  { fix l g2 a2 r2 l' assume A: "conv_A A \<turnstile> l \<longrightarrow>g2,a2,r2 l'"
+    obtain g1 a1 r1 where **: "A \<turnstile> l \<longrightarrow>g1,a1,r1 l'"
+      using A unfolding trans_of_def conv_A_def by (cases A) force
+    then have "\<exists> g1 a1 r1. A \<turnstile> l \<longrightarrow>g1,a1,r1 l'"
+      by auto
+  } note ** = this
+  show ?thesis
+    unfolding deadlocked_def
+    apply (rule ccontr)
+    apply clarsimp
+    apply (cases rule: step'.cases, assumption)
+    apply (cases rule: step_a.cases, assumption)
+    apply (auto 4 3 elim: * dest: ** step_delay_loc)
+    done
+qed
+
+lemma deadlock_if_not_is_start_in_states:
+  "deadlock (l0, Z0)" if "\<not> is_start_in_states"
+  unfolding deadlock_def using deadlocked_if_not_is_start_in_states[OF that] by blast
+
+sepref_definition is_start_in_states_impl is
+  "uncurry0 (RETURN is_start_in_states)" :: "unit_assnk \<rightarrow>a bool_assn"
+  unfolding is_start_in_states_def by sepref
+
+text \<open>Turning this into a Hoare triple:\<close>
+thm is_start_in_states_impl.refine[to_hnr, unfolded hn_refine_def, simplified]
+lemma is_start_in_states_impl_hoare:
+  "<emp> is_start_in_states_impl
+   <\<lambda>r. \<up> ((r \<longrightarrow> l0 \<in> Simulation_Graphs_TA.state_set A)
+         \<and> (\<not>r \<longrightarrow> deadlocked (l0, \<lambda>_ . 0))
+   )>t"
+  by (sep_auto
+      intro: deadlocked_if_not_is_start_in_states
+      simp: mod_star_conv is_start_in_states[simplified]
+      heap: is_start_in_states_impl.refine[to_hnr, unfolded hn_refine_def, simplified]
+      )
+
+lemma deadlock_if_deadlocked:
+  "deadlock y" if "deadlocked y"
+  using that unfolding deadlock_def by (inst_existentials y) auto
+
+lemma is_start_in_states_impl_hoare':
+  "<emp> is_start_in_states_impl
+   <\<lambda>r. \<up> ((r \<longrightarrow> l0 \<in> Simulation_Graphs_TA.state_set A)
+         \<and> (\<not>r \<longrightarrow> (\<exists>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<and> deadlock (l0, u0)))
+   )>t"
+  by (rule cons_post_rule[OF is_start_in_states_impl_hoare]) (auto intro: deadlock_if_deadlocked)
+
+end (* TA Impl *)
+
+context Reachability_Problem_Impl
+begin
+
+context
+  assumes l0_in_A: "l0 \<in> Simulation_Graphs_TA.state_set A"
+begin
+
+interpretation TA:
+  Regions_TA_Start_State v n "Suc n" "{1..<Suc n}" k "conv_A A" l0 "{u. \<forall>c\<in>{1..n}. u c = 0}"
+  apply standard
+  subgoal
+    by (intro l0_state_set l0_in_A)
+  subgoal
+    unfolding V'_def
+    apply safe
+    subgoal for x
+      unfolding V_def by auto
+    apply (inst_existentials "curry init_dbm :: real DBM")
+     apply (simp add: init_dbm_zone)
+    by (rule dbm_int_init_dbm)
+  subgoal
+    by auto
+  done
+
+interpretation bisim:
+  Bisimulation_Invariant
+  "\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z'"
+  "\<lambda>a b. op.E_from_op a b"
+  "\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z"
+  "\<lambda>_. True" "\<lambda>(l, M). l \<in> states \<and> wf_state (l, M)"
+  by (rule beta_final_bisim)
+
+lemma check_deadlock:
+  "(\<exists>a. E_op''.E_from_op** a0 a \<and>
+    \<not> (case a of (l, M) \<Rightarrow> check_diag n M) \<and> (case a of (l, M) \<Rightarrow> \<not> check_deadlock_dbm l M))
+  \<longleftrightarrow> (\<exists>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<and> deadlock (l0, u0))" (is "?l \<longleftrightarrow> ?r")
+proof -
+  text \<open>@{term TA.reaches} corresponds to non-empty steps of @{term step_z_beta'}\<close>
+  text \<open>@{term bisim.A.reaches} corresponds to steps of @{term step_z_beta'}\<close>
+  text \<open>@{term \<open>E_op''.E_from_op** a0\<close>} corresponds to steps of @{term op.E_from_op} (@{term E_op''})\<close>
+  have "?r \<longleftrightarrow> (\<exists>l Z. TA.reaches (l0, {u. \<forall>c\<in>{1..n}. u c = 0}) (l, Z) \<and> \<not>TA.check_deadlock l Z)"
+    using TA.deadlock_check unfolding TA.a0_def from_R_def by simp
+  also have
+    "\<dots> \<longleftrightarrow> (\<exists>l Z. bisim.A.reaches (l0, {u. \<forall>c\<in>{1..n}. u c = 0}) (l, Z) \<and> \<not>TA.check_deadlock l Z)"
+    apply safe
+    subgoal for l Z
+      by (subst (asm) TA.reaches_steps_z_beta_iff) auto
+    subgoal for l Z
+      by (subst TA.reaches_steps_z_beta_iff) (auto dest: not_check_deadlock_non_empty)
+    done
+  also have "\<dots> \<longleftrightarrow> (\<exists>l M. E_op''.E_from_op** a0 (l, M) \<and> \<not> check_deadlock_dbm l M)"
+    using bisim.reaches_ex_iff[where
+        \<phi> = "\<lambda> (l, Z). \<not>TA.check_deadlock l Z" and \<psi> = "\<lambda>(l, M). \<not> check_deadlock_dbm l M",
+        OF not_check_deadlock_compatible, of "(l0, {u. \<forall>c\<in>{1..n}. u c = 0})" a0
+        ]
+    using wf_state_init init_dbm_zone states'_states unfolding a0_def by force
+  also have "\<dots> \<longleftrightarrow> ?l"
+    by (auto 4 4 dest: deadlock_check_diag)
+  finally show ?thesis ..
+qed
+
+lemma check_deadlock':
+  "(\<nexists>a. E_op''.E_from_op** a0 a \<and>
+    \<not> (case a of (l, M) \<Rightarrow> check_diag n M) \<and> (case a of (l, M) \<Rightarrow> \<not> check_deadlock_dbm l M))
+  \<longleftrightarrow> (\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0))" (is "?l \<longleftrightarrow> ?r")
+  using check_deadlock by auto
+
+context
+  assumes "F = (\<lambda> _. False)"
+begin
+
+interpretation Worklist_Map2_Impl_check
+  op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes'
+  "\<lambda> (l, M). F l" state_assn'
+  succs_impl a0_impl F_impl subsumes_impl emptiness_check_impl fst "return o fst" state_copy_impl
+  tracei location_assn "\<lambda>(l, M). \<not> check_deadlock_dbm l M" check_deadlock_neg_impl
+  apply standard
+  subgoal
+    using check_deadlock_neg_impl.refine unfolding PR_CONST_def .
+  subgoal
+    unfolding F_rel_def unfolding \<open>F = _\<close> by auto
+  subgoal for a b
+    apply (clarsimp simp: E_op''.reachable_def)
+    apply (subst (asm) check_deadlock_dbm_correct''[symmetric], assumption)
+    apply (subst (asm) check_deadlock_dbm_correct''[symmetric], assumption)
+    apply (frule subsumes_loc_eqD, assumption)
+    apply (drule subsumes_dbm_subsetD, assumption)
+    apply (drule dbm_subset_conv)
+    apply (subst (asm) dbm_subset_correct''[symmetric])
+    by (auto dest: TA.check_deadlock_anti_mono E_op''.reachable_wf_dbm simp: E_op''.reachable_def)
+  done
+
+lemmas check_passed_impl_hnr =
+  check_passed_impl_hnr[unfolded op.reachable_def, unfolded op_E_from_op_iff check_deadlock]
+
+end (* F is always false *)
+
+end (* l0 belongs to A *)
+
+definition
+  "check_passed_impl_start \<equiv> do {
+    r1 \<leftarrow> is_start_in_states_impl;
+    if r1 then do {
+      r2 \<leftarrow> check_passed_impl succs_impl a0_impl F_impl subsumes_impl emptiness_check_impl
+            (return \<circ> fst) state_copy_impl tracei check_deadlock_neg_impl;
+             return r2
+    }
+    else return True
+  }"
+
+lemma check_passed_impl_start_hnr:
+  "(uncurry0 check_passed_impl_start,
+    uncurry0 (RETURN (\<exists>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<and> deadlock (l0, u0)))
+  ) \<in> unit_assnk \<rightarrow>a bool_assn" if "F = (\<lambda>_. False)"
+proof -
+  define has_deadlock where "has_deadlock \<equiv> \<exists>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<and> deadlock (l0, u0)"
+  note [sep_heap_rules] =
+    check_passed_impl_hnr[
+      OF _ that, to_hnr, unfolded hn_refine_def, folded has_deadlock_def, simplified
+    ]
+    is_start_in_states_impl_hoare'[folded has_deadlock_def, simplified]
+  show ?thesis
+    unfolding has_deadlock_def[symmetric] check_passed_impl_start_def
+    by sepref_to_hoare (sep_auto simp: mod_star_conv)
+qed
+
+end (* Reachability Problem Impl *)
+
+
+context Reachability_Problem_precompiled
+begin
+
+lemma F_is_False_iff:
+  "PR_CONST F = (\<lambda>_. False) \<longleftrightarrow> final = []"
+  unfolding F_def by (cases final; simp; metis)
+
+lemma F_impl_False: "F_impl = (\<lambda>_. return False)" if "final = []"
+  using that unfolding F_impl_def unfolding final_fun_def List.member_def by auto
+
+definition deadlock_checker where
+  "deadlock_checker \<equiv>
+    let
+      key = return \<circ> fst;
+      sub = subsumes_impl;
+      copy = state_copy_impl;
+      start = a0_impl;
+      final = (\<lambda>_. return False);
+      succs = succs_impl;
+      empty = emptiness_check_impl;
+      P = check_deadlock_neg_impl;
+      trace = tracei
+    in do {
+      r1 \<leftarrow> is_start_in_states_impl;
+      if r1 then do {
+        r2 \<leftarrow> check_passed_impl succs start final sub empty key copy trace P;
+        return r2
+      }
+      else return True
+    }"
+
+theorem deadlock_checker_hnr:
+  assumes "final = []"
+  shows
+    "(uncurry0 deadlock_checker, uncurry0 (RETURN (\<exists>u0. (\<forall>c\<in>{1..m}. u0 c = 0) \<and> deadlock (0, u0))))
+     \<in> unit_assnk \<rightarrow>a bool_assn"
+  unfolding deadlock_checker_def Let_def
+  using check_passed_impl_start_hnr[
+      unfolded F_is_False_iff F_impl_False[OF assms] check_passed_impl_start_def,
+      OF assms] .
+
+schematic_goal deadlock_checker_alt_def:
+  "deadlock_checker \<equiv> ?impl"
+  unfolding deadlock_checker_def
+  unfolding succs_impl_def
+  unfolding E_op''_impl_def abstr_repair_impl_def abstra_repair_impl_def
+  unfolding start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding check_deadlock_neg_impl_def check_deadlock_impl_def
+  unfolding is_start_in_states_impl_def
+   apply (abstract_let "IArray.sub (IArray (map (IArray o map int) k))" k)
+   apply (abstract_let "inv_fun" inv_fun)
+   apply (abstract_let "trans_impl" trans)
+   unfolding inv_fun_def[abs_def] trans_impl_def[abs_def]
+   apply (abstract_let "IArray inv" inv_ia)
+   apply (abstract_let "IArray trans_map" trans_map)
+   unfolding trans_map_def label_def
+   unfolding init_dbm_impl_def a0_impl_def
+   unfolding subsumes_impl_def
+   unfolding emptiness_check_impl_def
+   unfolding state_copy_impl_def
+  by (rule Pure.reflexive)
+
+end (* Reachability Problem Precompiled *)
+
+concrete_definition deadlock_checker_impl
+  uses Reachability_Problem_precompiled.deadlock_checker_alt_def
+
+export_code deadlock_checker_impl in SML_imp module_name TA
+
+definition [code]:
+  "check_deadlock n m k I T \<equiv>
+    if Reachability_Problem_precompiled n m k I T
+    then deadlock_checker_impl m k I T \<bind> (\<lambda>x. return (Some x))
+    else return None"
+
+theorem deadlock_check:
+  "(uncurry0 (check_deadlock n m k I T),
+    uncurry0 (
+       RETURN (
+        if (Reachability_Problem_precompiled n m k I T)
+        then Some (
+            if
+              \<exists> u0. (\<forall> c \<in> {1..m}. u0 c = 0) \<and>
+              Graph_Defs.deadlock
+                (\<lambda>(l, u) (l', u').
+                    (conv_A (Reachability_Problem_precompiled_defs.A n I T)) \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+                (0, u0)
+            then True
+            else False
+          )
+        else None
+       )
+    )
+   ) \<in> unit_assnk \<rightarrow>a id_assn"
+proof -
+  define reach_check where
+    "reach_check =
+    (\<exists> u0. (\<forall> c \<in> {1..m}. u0 c = 0) \<and>
+              Graph_Defs.deadlock
+            (\<lambda>(l, u) (l', u').
+                (conv_A (Reachability_Problem_precompiled_defs.A n I T)) \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+            (0, u0))"
+  note [sep_heap_rules] = Reachability_Problem_precompiled.deadlock_checker_hnr[OF _ HOL.refl,
+      to_hnr, unfolded hn_refine_def, rule_format,
+      of n m k I T,
+      unfolded reach_check_def[symmetric]
+      ]
+  show ?thesis
+    unfolding reach_check_def[symmetric]
+    by sepref_to_hoare
+       (sep_auto simp: deadlock_checker_impl.refine[symmetric] mod_star_conv check_deadlock_def)
+qed
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/Mlton_Main.sml
--- /dev/null
+++ thys/Munta_Model_Checker/ML/Mlton_Main.sml
@@ -0,0 +1,1 @@
+val _ = if MLton.isMLton then main() else ()
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/Muntax.sml
--- /dev/null
+++ thys/Munta_Model_Checker/ML/Muntax.sml
@@ -0,0 +1,122 @@
+open Model_Checker;
+open Util;
+
+(*** Wrapping up the checker ***)
+fun run_and_print check_deadlock s =
+  let
+    val debug_level: Int32.int Unsynchronized.ref = ref 0
+    val _ = debug_level := 2
+    val t = Time.now ()
+    val r = parse_convert_run check_deadlock s ()
+    val t = Time.- (Time.now (), t)
+    val _ = println("")
+    val _ = println("Internal time for precondition check + actual checking: " ^ Time.toString t)
+    val _ = println("")
+  in
+    case r of
+      Error es => let
+        val _ = println "Failure:"
+        val _ = map println es
+      in () end
+    | Result r => let
+      val _ = if !debug_level >= 1 then let
+          val _ = println("# explored states: " ^ Int.toString(Tracing.get_count ()))
+          val _ = println("")
+        in () end else ()
+      in println r end
+  end
+
+fun check_and_verify_from_stream stream _ =
+  let
+    val args = CommandLine.arguments()
+    val check_deadlock = List.find (fn x => x = "-deadlock" orelse x = "-dc") args <> NONE
+    val input = read_lines stream
+  in
+    if input = ""
+    then println "Failed to read line from input!"
+      (* We append a space to terminate the input for the parser *)
+    else input ^ " " |> run_and_print check_deadlock
+  end
+
+val to_large_int = fn x => x;
+
+fun print_json () =
+  let
+    val messages = Logging.get_trace ()
+    val jsons = filter (fn (i, s) => i = 2) messages |> map snd
+    val _ = println ""
+    val _ = println "The following JSON was read by the parser:"
+    val _ = if length jsons > 0 then println (hd jsons) else ()
+  in () end
+
+fun print_explored () =
+  let
+    val messages = Logging.get_trace ()
+    val explored = filter (fn (i, s) => i = 5 andalso String.isPrefix "Explored: " s) messages
+    val explored = map (fn (_, s) => String.extract (s, String.size "Explored: ", NONE)) explored
+    val final = filter (fn (i, s) => i = 5 andalso String.isPrefix "Final: " s) messages
+    val final = map (fn (_, s) => String.extract (s, String.size "Final: ", NONE)) final
+    val _ = println ""
+    val _ = println "The search explored the following states:"
+    val _ = map println explored
+    val _ = if length final > 0 then println "The search hit the following final state:" else ()
+    val _ = if length final > 0 then map println final else [()]
+  in () end
+
+val arguments = common_arguments @ [
+  (["explored", "e"], "Report set of explored states (only works for reachability formulas).", Flag),
+  (["show", "s"], "Print back the JSON that was parsed from the input.", Flag)
+]
+
+fun main () =
+  let
+    val _ = read_arguments arguments
+    val check_deadlock = is_present "deadlock"
+    val cpu_time = is_present "cpu-time"
+    val trace_explored = is_present "explored"
+    val model = find_arg "model"
+    val num_threads = find_arg "num-threads"
+    val show_json = is_present "show"
+    val show_help = is_present "help"
+    fun convert f NONE = NONE
+      | convert f (SOME x) = SOME (f x)
+        handle Fail msg => (println ("Argument error: " ^ msg); OS.Process.exit OS.Process.failure)
+    fun int_of_string err_msg s = case Int.fromString s of
+        NONE => raise Fail (err_msg ^ " should be an integer")
+      | SOME x => x
+    fun the_default x NONE = x
+      | the_default _ (SOME x) = x
+    val num_threads = num_threads
+      |> convert (int_of_string "Number of threads")
+      |> the_default 1
+    val _ = println "** Configuration options **"
+    val _ = "* Deadlock: " ^ (if check_deadlock then "True" else "False") |> println
+    val _ = "* Model: " ^ (case model of NONE => "from stdin" | SOME m => m) |> println
+    val _ = "* Report explored states: " ^ (if trace_explored then "True" else "False") |> println
+    (* val _ = "* Num Threads: " ^ Int.toString num_threads |> println *)
+    (* val _ = "* Measuring CPU time: " ^ (if cpu_time then "True" else "False") |> println *)
+    val num_threads = 10000
+    (* val num_threads = num_threads |> to_large_int |> nat_of_integer *)
+    (* val _ = if cpu_time then Timing.set_cpu true else () *)
+    val _ = if show_json then Logging.set_level 2 else ()
+    val _ = if trace_explored then Logging.set_level 5 else ()
+    val input = ref ""
+  in
+    if show_help then
+      print_usage arguments
+    else (case model of
+      NONE => input := read_lines TextIO.stdIn |
+      SOME f =>
+        let
+          val file = TextIO.openIn f
+          val s = read_lines file
+          val _ = TextIO.closeIn file
+          val _ = input := s
+        in () end;
+      run_and_print check_deadlock (!input);
+      (if show_json then print_json () else ());
+      (if trace_explored then print_explored () else ())
+    )
+  end
+
+(* val main = check_and_verify_from_stream TextIO.stdIn *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/Unsynchronized.sml
--- /dev/null
+++ thys/Munta_Model_Checker/ML/Unsynchronized.sml
@@ -0,0 +1,6 @@
+structure Unsynchronized =
+struct
+
+datatype ref = datatype ref;
+
+end;
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/Util.sml
--- /dev/null
+++ thys/Munta_Model_Checker/ML/Util.sml
@@ -0,0 +1,94 @@
+structure Util =
+struct
+
+(** Printing utilites **)
+fun println s = print (s ^ "\n")
+
+fun list_to_string f = (fn x => "[" ^ x ^ "]") o String.concatWith ", " o map f
+
+fun print_result NONE = println("Invalid input\n")
+    | print_result (SOME true) = println("Property is satisfied\n")
+    | print_result (SOME false) = println("Property is not satisfied\n")
+
+(** File reading **)
+fun read_lines stream =
+  let
+    val input = TextIO.inputLine stream
+      handle Size => (println "Input line too long!"; raise Size)
+  in
+    case input of
+      NONE => ""
+    | SOME input => input ^ read_lines stream
+  end
+
+fun read_file f =
+  let
+    val file = TextIO.openIn f
+    val s = read_lines file
+    val _ = TextIO.closeIn file
+  in s end
+
+(** Processing command line arguments **)
+datatype 'a argument = Is_Present | Is_Not_Present | Has_Val of 'a
+
+type arguments = string argument list
+
+exception Unknown_Argument of string
+
+datatype argument_type = Arg | Flag
+
+val common_arguments = [
+  (["deadlock", "dc"], "Ignore formula and run deadlock check.", Flag),
+  (["model", "m"],
+    "Input file for the model & formula. "
+    ^ "If this is not specified, the input is read from stdin.",     
+    Arg),
+  (["help", "h", "?"], "Show this help string.", Flag),
+  (["cpu-time", "cpu"], "Report CPU time.", Flag),
+  (["num-threads", "n"], "Number of threads.", Arg)
+]
+
+val the_args = ref []
+
+fun is_present arg = case List.find (fn (k, v) => k = arg) (!the_args) of
+    NONE => raise Unknown_Argument arg
+  | SOME (_, Is_Not_Present) => false
+  | _ => true
+
+fun find_arg arg = case List.find (fn (k, v) => k = arg) (!the_args) of
+    NONE => raise Unknown_Argument arg
+  | SOME (_, Has_Val v) => SOME v
+  | _ => NONE
+
+val get_arg = the o find_arg
+
+fun find_with_arg P [] = NONE
+  | find_with_arg P [_] = NONE
+  | find_with_arg P (x :: y :: xs) = if P x then SOME y else find_with_arg P (y :: xs)
+
+fun read_arguments arguments =
+  let
+    val args = CommandLine.arguments()
+    fun find_name names name = List.find (fn x => "-" ^ x = name) names <> NONE
+    fun find_it (names, Flag) =
+      (case List.find (find_name names) args of
+        NONE => Is_Not_Present
+      | SOME _ => Is_Present)
+    | find_it (names, Arg) =
+      (case find_with_arg (find_name names) args of
+        NONE => Is_Not_Present
+      | SOME x => Has_Val x)
+    val args = map (fn (names, _, typ) => (hd names, find_it (names, typ))) arguments
+  in
+    the_args := args
+  end
+
+fun print_usage arguments =
+  arguments
+  |> map (fn (names, descr, _) =>
+      (map (fn s => "-" ^ s ^ " ") names|> String.concat) ^ ": " ^ descr ^ "\n")
+  |> String.concat
+  |> (fn s => "Usage:\n" ^ s)
+  |> println
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/Writeln.sml
--- /dev/null
+++ thys/Munta_Model_Checker/ML/Writeln.sml
@@ -0,0 +1,1 @@
+fun writeln s = (print s; print "\n")
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/check_benchmark.sh
--- /dev/null
+++ thys/Munta_Model_Checker/ML/check_benchmark.sh
@@ -0,0 +1,24 @@
+#!/bin/bash
+
+# Arguments
+expected_states=$1  # expected number of states
+expected_property=$2  # e.g. "is satisfied" or "is not satisfied"
+command_to_run="${@:3}"  # command to execute
+
+# Run command and capture output
+command_output=$($command_to_run 2>&1)
+last_three_lines=$(echo "$command_output" | tail -n 3)
+
+expected_states="# explored states: $expected_states"
+
+# Check if output matches expected format
+if echo "$last_three_lines" | grep -Fxq "$expected_states" && echo "$last_three_lines" | grep -Fxq "$expected_property"; then
+    echo "Match found!"
+    exit 0
+else
+    echo "No match!"
+    echo "Actual output:"
+    echo "$last_three_lines"
+    exit 1
+fi
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/library.ML
--- /dev/null
+++ thys/Munta_Model_Checker/ML/library.ML
@@ -0,0 +1,75 @@
+(*
+The content of this file is taken from Isabelle: https://isabelle.in.tum.de.
+It is a small subset of src/Pure/General/basics.ML and src/Pure/library.ML.
+
+Isabelle's original copyright notice:
+
+ISABELLE COPYRIGHT NOTICE, LICENCE AND DISCLAIMER.
+
+Copyright (c) 1986-2025,
+  University of Cambridge,
+  Technische Universitaet Muenchen,
+  and contributors.
+
+  All rights reserved.
+
+Redistribution and use in source and binary forms, with or without 
+modification, are permitted provided that the following conditions are 
+met:
+
+* Redistributions of source code must retain the above copyright 
+notice, this list of conditions and the following disclaimer.
+
+* Redistributions in binary form must reproduce the above copyright 
+notice, this list of conditions and the following disclaimer in the 
+documentation and/or other materials provided with the distribution.
+
+* Neither the name of the University of Cambridge or the Technische
+Universitaet Muenchen nor the names of their contributors may be used
+to endorse or promote products derived from this software without
+specific prior written permission.
+
+THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS 
+IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED 
+TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A 
+PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT 
+OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, 
+SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT 
+LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, 
+DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY 
+THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT 
+(INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE 
+OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+
+*)
+
+infix 1 |>
+
+signature LIBRARY =
+sig
+
+  val |> : 'a * ('a -> 'b) -> 'b
+
+  val the: 'a option -> 'a
+
+  val snd: 'a * 'b -> 'b
+
+  val filter: ('a -> bool) -> 'a list -> 'a list
+
+end;
+
+structure Library: LIBRARY =
+struct
+
+fun x |> f = f x;
+
+fun the (SOME x) = x
+  | the NONE = raise Option.Option;
+
+fun snd (x, y) = y;
+
+val filter = List.filter;
+
+end;
+
+open Library;
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ML/muntax.mlb
--- /dev/null
+++ thys/Munta_Model_Checker/ML/muntax.mlb
@@ -0,0 +1,22 @@
+$(SML_LIB)/basis/basis.mlb
+$(SML_LIB)/basis/mlton.mlb
+
+local
+    Unsynchronized.sml
+
+    Writeln.sml
+
+    ann "nonexhaustiveMatch ignore" in
+      Simple_Model_Checker.ML
+    end
+in
+    structure Tracing
+    structure Logging
+    structure Model_Checker
+end
+
+library.ML
+Unsynchronized.sml
+Util.sml
+Muntax.sml
+Mlton_Main.sml
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Model_Checking/Shortest_SCC_Paths.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Model_Checking/Shortest_SCC_Paths.thy
@@ -0,0 +1,75 @@
+theory Shortest_SCC_Paths
+  imports Gabow_SCC.Gabow_SCC_Code
+begin
+
+definition compute_SCC_tr ::
+  "('a :: hashable \<Rightarrow> bool, 'a \<Rightarrow> 'a list, 'a list, 'b) gen_g_impl_scheme \<Rightarrow> _" where
+  "compute_SCC_tr =
+    Gabow_SCC_Code.compute_SCC_tr (=) bounded_hashcode_nat (def_hashmap_size TYPE('a))"
+
+\<comment> \<open>Example usage:\<close>
+term "compute_SCC_tr \<lparr> gi_V = (\<lambda> x. True), gi_E = (\<lambda> x. [3]), gi_V0 = [1], \<dots> = 3 \<rparr>"
+
+text \<open>
+  Efficiently calculate shortest paths in a graph with non-negative edge weights,
+  and where the only cycles are 0-cycles.
+\<close>
+definition
+"calc_shortest_scc_paths G n \<equiv>
+let
+  sccs = compute_SCC_tr G;
+  d = replicate n None @ [Some 0];
+  d = (
+    fold
+      (\<lambda> vs d.
+        fold
+          (\<lambda> u d.
+            fold
+              (\<lambda> v d.
+                case d ! u of
+                  None \<Rightarrow> d
+                | Some du \<Rightarrow> (
+                  case d ! v of
+                    None \<Rightarrow> d[v := Some (du + more G u v)]
+                  | Some dv \<Rightarrow>
+                      if du + more G u v < dv
+                      then d[v := Some (du + more G u v)]
+                      else d
+                  )
+              )
+              (gi_E G u)
+              d
+          )
+          vs
+          d
+      )
+      sccs
+      d
+  );
+  d = (
+    fold
+      (\<lambda> vs d.
+        let
+          dscc =
+            fold
+            (\<lambda> v dscc.
+              case dscc of
+                None \<Rightarrow> d ! v
+              | Some d' \<Rightarrow> (
+                  case d ! v of
+                    None \<Rightarrow> dscc
+                  | Some dv \<Rightarrow> Some (min dv d')
+                )
+            )
+            vs
+            None
+        in
+          fold (\<lambda> v d. d[v:=dscc]) vs d
+      )
+      sccs
+      d
+  )
+in d
+"
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Model_Checking/Simple_Network_Language_Export_Code.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Model_Checking/Simple_Network_Language_Export_Code.thy
@@ -0,0 +1,1677 @@
+section \<open>Assembling the Model Checker and Generating Code\<close>
+
+text \<open>
+We assemble the model checker:
+
+item A parser is added to parse JSON files
+item The resulting JSON data is validated and converted to the simple networks language
+item The verified model checker is run on this input, and the output is printed
+item The verified part includes a ``renaming'' step to normalize identifiers in the input
+
+Then the code is exported:
+
+item We add some logging utilities (via code printings)
+item Code generation is setup
+item We export this to SML via reflection
+item And test it on a few small examples
+\<close>
+
+theory Simple_Network_Language_Export_Code
+  imports
+    Munta_Model_Checker.JSON_Parsing
+    Munta_Model_Checker.Simple_Network_Language_Renaming
+    Munta_Model_Checker.Simple_Network_Language_Deadlock_Checking
+    Shortest_SCC_Paths
+    Munta_Base.Error_List_Monad
+begin
+
+datatype result =
+  Renaming_Failed | Preconds_Unsat | Sat | Unsat
+
+abbreviation "renum_automaton \<equiv> Simple_Network_Rename_Defs.renum_automaton"
+
+locale Simple_Network_Rename_Formula_String_Defs =
+  Simple_Network_Rename_Defs_int where automata = automata for automata ::
+    "(nat list \<times> nat list \<times>
+     (String.literal act, nat, String.literal, int, String.literal, int) transition list
+      \<times> (nat \<times> (String.literal, int) cconstraint) list) list"
+begin
+
+definition check_renaming where "check_renaming urge \<Phi> L0 s0 \<equiv> combine [
+    assert (\<forall>i<n_ps. \<forall>x\<in>loc_set. \<forall>y\<in>loc_set. renum_states i x = renum_states i y \<longrightarrow> x = y)
+      STR ''Location renamings are injective'',
+    assert (inj_on renum_clocks (insert urge clk_set'))
+      STR ''Clock renaming is injective'',
+    assert (inj_on renum_vars var_set)
+      STR ''Variable renaming is injective'',
+    assert (inj_on renum_acts act_set)
+      STR ''Action renaming is injective'',
+    assert (fst ` set bounds' = var_set)
+      STR ''Bound set is exactly the variable set'',
+    assert (\<Union> ((\<lambda>g. fst ` set g) ` set (map (snd o snd o snd) automata)) \<subseteq> loc_set)
+      STR ''Invariant locations are contained in the location set'',
+    assert (\<Union> ((set o fst) ` set automata) \<subseteq> loc_set)
+      STR ''Broadcast locations are containted in the location set'',
+    assert ((\<Union>x\<in>set automata. set (fst (snd x))) \<subseteq> loc_set)
+      STR ''Urgent locations are containted in the location set'',
+    assert (urge \<notin> clk_set')
+      STR ''Urge not in clock set'',
+    assert (L0 \<in> states)
+      STR ''Initial location is in the state set'',
+    assert (fst ` set s0 = var_set)
+      STR ''Initial state has the correct domain'',
+    assert (distinct (map fst s0))
+      STR ''Initial state is unambiguous'',
+    assert (set2_formula \<Phi> \<subseteq> loc_set)
+      STR ''Formula locations are contained in the location set'',
+    assert (locs_of_formula \<Phi> \<subseteq> {0..<n_ps})
+      STR ''Formula automata are contained in the automata set'',
+    assert (vars_of_formula \<Phi> \<subseteq> var_set)
+      STR ''Variables of the formula are contained in the variable set''
+  ]
+"
+
+end (* Simple_Network_Rename_Formula_String_Defs *)
+
+locale Simple_Network_Rename_Formula_String =
+  Simple_Network_Rename_Formula_String_Defs +
+  fixes urge :: String.literal
+  fixes \<Phi> :: "(nat, nat, String.literal, int) formula"
+    and s0 :: "(String.literal \<times> int) list"
+    and L0 :: "nat list"
+  assumes renum_states_inj:
+    "\<forall>i<n_ps. \<forall>x\<in>loc_set. \<forall>y\<in>loc_set. renum_states i x = renum_states i y \<longrightarrow> x = y"
+  and renum_clocks_inj: "inj_on renum_clocks (insert urge clk_set')"
+  and renum_vars_inj:   "inj_on renum_vars var_set"
+  and renum_acts_inj:   "inj_on renum_acts act_set"
+  and bounds'_var_set:  "fst ` set bounds' = var_set"
+  and loc_set_invs: "\<Union> ((\<lambda>g. fst ` set g) ` set (map (snd o snd o snd) automata)) \<subseteq> loc_set"
+  and loc_set_broadcast: "\<Union> ((set o fst) ` set automata) \<subseteq> loc_set"
+  and loc_set_urgent: "\<Union> ((set o (fst o snd)) ` set automata) \<subseteq> loc_set"
+  and urge_not_in_clk_set: "urge \<notin> clk_set'"
+  assumes L0_states: "L0 \<in> states"
+      and s0_dom: "fst ` set s0 = var_set" and s0_distinct: "distinct (map fst s0)"
+  assumes formula_dom:
+    "set2_formula \<Phi> \<subseteq> loc_set"
+    "locs_of_formula \<Phi> \<subseteq> {0..<n_ps}"
+    "vars_of_formula \<Phi> \<subseteq> var_set"
+begin
+
+interpretation Simple_Network_Rename_Formula
+  by (standard;
+      rule renum_states_inj renum_clocks_inj renum_vars_inj bounds'_var_set renum_acts_inj
+        loc_set_invs loc_set_broadcast loc_set_urgent urge_not_in_clk_set
+        infinite_literal infinite_UNIV_nat L0_states s0_dom s0_distinct formula_dom
+     )+ \<comment> \<open>slow: 40s\<close>
+
+lemmas Simple_Network_Rename_intro = Simple_Network_Rename_Formula_axioms
+
+end
+
+lemma is_result_assert_iff:
+  "is_result (assert b m) \<longleftrightarrow> b"
+  unfolding assert_def by auto
+
+lemma is_result_combine_Cons_iff:
+  "is_result (combine (x # xs)) \<longleftrightarrow> is_result x \<and> is_result (combine xs)"
+  by (cases x; cases "combine xs") auto
+
+lemma is_result_combine_iff:
+  "is_result (a <|> b) \<longleftrightarrow> is_result a \<and> is_result b"
+  by (cases a; cases b) (auto simp: combine2_def)
+
+context Simple_Network_Rename_Formula_String_Defs
+begin
+
+lemma check_renaming:
+  "Simple_Network_Rename_Formula_String
+    broadcast bounds'
+    renum_acts renum_vars renum_clocks renum_states
+    automata urge \<Phi> s0 L0 \<longleftrightarrow>
+  is_result (check_renaming urge \<Phi> L0 s0)
+  "
+  unfolding check_renaming_def Simple_Network_Rename_Formula_String_def
+  by (simp add: is_result_combine_Cons_iff is_result_assert_iff del: combine.simps(2))
+
+end
+
+context Simple_Network_Impl_nat_defs
+begin
+
+definition check_precond1 where
+"check_precond1 =
+  combine [
+    assert (m > 0)
+      (STR ''At least one clock''),
+    assert (0 < length automata)
+      (STR ''At least one automaton''),
+    assert (\<forall>i<n_ps. let (_, _, trans, _) = (automata ! i) in \<forall> (l, _, _, _, _, _, l') \<in> set trans.
+      l < num_states i \<and> l' < num_states i)
+      (STR ''Number of states is correct (transitions)''),
+    assert (\<forall>i < n_ps. let (_, _, _, inv) = (automata ! i) in \<forall> (x, _) \<in> set inv. x < num_states i)
+      (STR ''Number of states is correct (invariants)''),
+    assert (\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, _, _, f, _, _) \<in> set trans.
+      \<forall>(x, upd) \<in> set f. x < n_vs \<and> (\<forall>i \<in> vars_of_exp upd. i < n_vs))
+      (STR ''Variable set bounded (updates)''),
+    assert (\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, b, _, _, _, _, _) \<in> set trans.
+      \<forall>i \<in> vars_of_bexp b. i < n_vs)
+      (STR ''Variable set bounded (guards)''),
+    assert (\<forall> i < n_vs. fst (bounds' ! i) = i)
+      (STR ''Bounds first index''),
+    assert (\<forall>a \<in> set broadcast. a < num_actions)
+      (STR ''Broadcast actions bounded''),
+    assert (\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, _, a, _, _, _) \<in> set trans.
+        pred_act (\<lambda>a. a < num_actions) a)
+      (STR ''Actions bounded (transitions)''),
+    assert (\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, g, _, _, r, _) \<in> set trans.
+      (\<forall>c \<in> set r. 0 < c \<and> c \<le> m) \<and>
+      (\<forall> (c, x) \<in> collect_clock_pairs g. 0 < c \<and> c \<le> m \<and> x \<in> \<nat>))
+      (STR ''Clock set bounded (transitions)''),
+    assert (\<forall>(_, _, _, inv) \<in> set automata. \<forall>(l, g) \<in> set inv.
+      (\<forall> (c, x) \<in> collect_clock_pairs g. 0 < c \<and> c \<le> m \<and> x \<in> \<nat>))
+      (STR ''Clock set bounded (invariants)''),
+    assert (\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, g, a, _, _, _) \<in> set trans.
+      case a of In a \<Rightarrow> a \<in> set broadcast \<longrightarrow> g = [] | _ \<Rightarrow> True)
+      (STR ''Broadcast receivers are unguarded''),
+    assert (\<forall>(_, U, _, _)\<in>set automata. U = [])
+      STR ''Urgency was removed''
+  ]
+"
+
+lemma check_precond1:
+  "is_result check_precond1
+  \<longleftrightarrow> Simple_Network_Impl_nat_urge broadcast bounds' automata m num_states num_actions"
+  unfolding check_precond1_def Simple_Network_Impl_nat_def Simple_Network_Impl_nat_urge_def
+    Simple_Network_Impl_nat_urge_axioms_def
+  by (simp add: is_result_combine_Cons_iff is_result_assert_iff del: combine.simps(2))
+
+context
+  fixes k :: "nat list list list"
+    and L0 :: "nat list"
+    and s0 :: "(nat \<times> int) list"
+    and formula :: "(nat, nat, nat, int) formula"
+begin
+
+definition check_precond2 where
+  "check_precond2 \<equiv>
+  combine [
+    assert (\<forall>i < n_ps. \<forall>(l, g) \<in> set ((snd o snd o snd) (automata ! i)).
+      \<forall>(x, m) \<in> collect_clock_pairs g. m \<le> int (k ! i ! l ! x))
+      (STR ''Ceiling invariants''),
+    assert (\<forall>i < n_ps. \<forall>(l, _, g, _) \<in> set ((fst o snd o snd) (automata ! i)).
+      (\<forall>(x, m) \<in> collect_clock_pairs g. m \<le> int (k ! i ! l ! x)))
+      (STR ''Ceiling transitions''),
+    assert (\<forall>i < n_ps. \<forall> (l, b, g, a, upd, r, l') \<in> set ((fst o snd o snd) (automata ! i)).
+       \<forall>c \<in> {0..<m+1} - set r. k ! i ! l' ! c \<le> k ! i ! l ! c)
+      (STR ''Ceiling resets''),
+    assert (length k = n_ps)
+      (STR ''Ceiling length''),
+    assert (\<forall> i < n_ps. length (k ! i) = num_states i)
+      (STR ''Ceiling length automata)''),
+    assert (\<forall> xs \<in> set k. \<forall> xxs \<in> set xs. length xxs = m + 1)
+      (STR ''Ceiling length clocks''),
+    assert (\<forall>i < n_ps. \<forall>l < num_states i. k ! i ! l ! 0 = 0)
+      (STR ''Ceiling zero clock''),
+    assert (\<forall>(_, _, _, inv) \<in> set automata. distinct (map fst inv))
+      (STR ''Unambiguous invariants''),
+    assert (bounded bounds (map_of s0))
+      (STR ''Initial state bounded''),
+    assert (length L0 = n_ps)
+      (STR ''Length of initial state''),
+    assert (\<forall>i < n_ps. L0 ! i \<in> fst ` set ((fst o snd o snd) (automata ! i)))
+      (STR ''Initial state has outgoing transitions''),
+    assert (vars_of_formula formula \<subseteq> {0..<n_vs})
+      (STR ''Variable set of formula'')
+]"
+
+lemma check_precond2:
+  "is_result check_precond2 \<longleftrightarrow>
+  Simple_Network_Impl_nat_ceiling_start_state_axioms broadcast bounds' automata m num_states
+      k L0 s0 formula"
+  unfolding check_precond2_def Simple_Network_Impl_nat_ceiling_start_state_axioms_def
+  by (simp add: is_result_combine_Cons_iff is_result_assert_iff del: combine.simps(2))
+
+end
+
+definition
+  "check_precond k L0 s0 formula \<equiv> check_precond1 <|> check_precond2 k L0 s0 formula"
+
+lemma check_precond:
+  "Simple_Network_Impl_nat_ceiling_start_state broadcast bounds' automata m num_states
+      num_actions k L0 s0 formula \<longleftrightarrow> is_result (check_precond k L0 s0 formula)"
+  unfolding check_precond_def is_result_combine_iff check_precond1 check_precond2
+    Simple_Network_Impl_nat_ceiling_start_state_def
+  ..
+
+end
+
+derive "show" acconstraint act sexp formula
+
+fun shows_exp and shows_bexp where
+  "shows_exp (const c) = show c" |
+  "shows_exp (var v) = show v" |
+  "shows_exp (if_then_else b e1 e2) =
+    shows_bexp b @ '' ? '' @ shows_exp e1 @ '' : '' @ shows_exp e2" |
+  "shows_exp (binop _ e1 e2) = ''binop '' @ shows_exp e1 @ '' '' @ shows_exp e2" |
+  "shows_exp (unop _ e) = ''unop '' @ shows_exp e" |
+  "shows_bexp (bexp.lt a b) = shows_exp a @ '' < '' @ shows_exp b" |
+  "shows_bexp (bexp.le a b) = shows_exp a @ '' <= '' @ shows_exp b" |
+  "shows_bexp (bexp.eq a b) = shows_exp a @ '' = '' @ shows_exp b" |
+  "shows_bexp (bexp.ge a b) = shows_exp a @ '' >= '' @ shows_exp b" |
+  "shows_bexp (bexp.gt a b) = shows_exp a @ '' > '' @ shows_exp b" |
+  "shows_bexp bexp.true = ''true''" |
+  "shows_bexp (bexp.not b) = ''! '' @ shows_bexp b" |
+  "shows_bexp (bexp.and a b) = shows_bexp a @ '' && '' @ shows_bexp b" |
+  "shows_bexp (bexp.or a b) = shows_bexp a @ '' || '' @ shows_bexp b" |
+  "shows_bexp (bexp.imply a b) = shows_bexp a @ '' -> '' @ shows_bexp b"
+
+instantiation bexp :: ("show", "show") "show"
+begin
+
+definition "shows_prec p (e :: (_, _) bexp) rest = shows_bexp e @ rest" for p
+
+definition "shows_list (es) s =
+  map shows_bexp es |> intersperse '', '' |> (\<lambda>xs. ''['' @ concat xs @ '']'' @ s)"
+instance
+  by standard (simp_all add: shows_prec_bexp_def shows_list_bexp_def show_law_simps)
+
+end
+
+instantiation exp :: ("show", "show") "show"
+begin
+
+definition "shows_prec p (e :: (_, _) exp) rest = shows_exp e @ rest" for p
+
+definition "shows_list (es) s =
+  map shows_exp es |> intersperse '', '' |> (\<lambda>xs. ''['' @ concat xs @ '']'' @ s)"
+instance
+  by standard (simp_all add: shows_prec_exp_def shows_list_exp_def show_law_simps)
+
+end
+
+
+definition
+  "pad m s = replicate m CHR '' '' @ s"
+
+definition "shows_rat r \<equiv> case r of fract.Rat s f b \<Rightarrow>
+  (if s then '''' else ''-'') @ show f @ (if b \<noteq> 0 then ''.'' @ show b else '''')"
+
+fun shows_json :: "nat \<Rightarrow> JSON \<Rightarrow> string" where
+  "shows_json n (Nat m) = show m |> pad n"
+| "shows_json n (Rat r) = shows_rat r |> pad n"
+| "shows_json n (JSON.Int r) = show r |> pad n"
+| "shows_json n (Boolean b) = (if b then ''true'' else ''false'') |> pad n"
+| "shows_json n Null = ''null'' |> pad n"
+| "shows_json n (String s) = pad n (''\"'' @ s @ ''\"'')"
+| "shows_json n (JSON.Array xs) = (
+  if xs = Nil then
+    pad n ''[]''
+  else
+    pad n ''[\<newline>''
+    @ concat (map (shows_json (n + 2)) xs |> intersperse '',\<newline>'')
+    @ ''\<newline>''
+    @ pad n '']''
+  )"
+| "shows_json n (JSON.Object xs) = (
+  if xs = Nil then
+    pad n ''{}''
+  else
+    pad n ''{\<newline>''
+    @ concat (
+      map (\<lambda>(k, v). pad (n + 2) (''\"'' @ k @ ''\"'') @ '':\<newline>'' @ shows_json (n + 4) v) xs
+      |> intersperse '',\<newline>''
+    )
+    @ ''\<newline>''
+    @ pad n ''}''
+  )"
+
+instantiation JSON :: "show"
+begin
+
+definition "shows_prec p (x :: JSON) rest = shows_json 0 x @ rest" for p
+
+definition "shows_list jsons s =
+  map (shows_json 0) jsons |> intersperse '', '' |> (\<lambda>xs. ''['' @ concat xs @ '']'' @ s)"
+instance
+  by standard (simp_all add: shows_prec_JSON_def shows_list_JSON_def show_law_simps)
+
+end
+
+
+definition rename_network where
+  "rename_network broadcast bounds' automata renum_acts renum_vars renum_clocks renum_states \<equiv>
+  let
+    automata = map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states) automata;
+    broadcast = map renum_acts broadcast;
+    bounds' = map (\<lambda>(a,b,c). (renum_vars a, b, c)) bounds'
+  in
+    (broadcast, automata, bounds')
+"
+
+definition
+  "show_clock inv_renum_clocks = show o inv_renum_clocks"
+
+definition
+  "show_locs inv_renum_states = show o map_index inv_renum_states"
+
+definition
+  "show_vars inv_renum_vars = show o map_index (\<lambda>i v. show (inv_renum_vars i) @ ''='' @ show v)"
+
+definition
+  "show_state inv_renum_states inv_renum_vars \<equiv> \<lambda>(L, vs).
+  let L = show_locs inv_renum_states L; vs = show_vars inv_renum_vars vs in
+  ''<'' @ L @ ''>, <'' @ vs @ ''>''"
+
+definition do_rename_mc where
+  "do_rename_mc f dc broadcast bounds' automata k urge L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks
+\<equiv>
+let
+   _ = println (STR ''Checking renaming'');
+   formula = (if dc then formula.EX (not sexp.true) else formula);
+   renaming_valid = Simple_Network_Rename_Formula_String_Defs.check_renaming
+      broadcast bounds' renum_acts renum_vars renum_clocks renum_states automata urge formula L0 s0;
+   _ = println (STR ''Renaming network'');
+   (broadcast, automata, bounds') = rename_network
+      broadcast bounds' (map (conv_urge urge) automata)
+      renum_acts renum_vars renum_clocks renum_states;
+   _ = trace_level 4 (\<lambda>_. return (STR ''Automata after renaming''));
+   _ = map (\<lambda>a. show a |> String.implode |> (\<lambda>s. trace_level 4 (\<lambda>_. return s))) automata;
+   _ = println (STR ''Renaming formula'');
+   formula =
+    (if dc then formula.EX (not sexp.true) else map_formula renum_states renum_vars id formula);
+    _ = println (STR ''Renaming state'');
+   L0 = map_index renum_states L0;
+   s0 = map (\<lambda>(x, v). (renum_vars x, v)) s0;
+   show_clock = show o inv_renum_clocks;
+   show_state = show_state inv_renum_states inv_renum_vars
+in
+  if is_result renaming_valid then do {
+    let _ = println (STR ''Checking preconditions'');
+    let r = Simple_Network_Impl_nat_defs.check_precond
+      broadcast bounds' automata m num_states num_actions k L0 s0 formula;
+    let _ = (case r of Result _ \<Rightarrow> ()
+      | Error es \<Rightarrow>
+        let
+          _ = println STR '''';
+          _ = println STR ''The following pre-conditions were not satisified:''
+        in
+          let _ = map println es in println STR '''');
+    let _ = println (STR ''Running precond_mc'');
+    let r = f show_clock show_state
+        broadcast bounds' automata m num_states num_actions k L0 s0 formula;
+    Some r
+  }
+  else do {
+    let _ = println (STR ''The following conditions on the renaming were not satisfied:'');
+    let _ = the_errors renaming_valid |> map println;
+    None}
+"
+
+definition rename_mc where
+  "rename_mc dc broadcast bounds' automata k urge L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks
+\<equiv>
+do {
+  let r = do_rename_mc (if dc then precond_dc else precond_mc)
+    dc broadcast bounds' automata k urge L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks;
+  case r of Some r \<Rightarrow> do {
+    r \<leftarrow> r;
+    case r of
+      None \<Rightarrow> return Preconds_Unsat
+    | Some False \<Rightarrow> return Unsat
+    | Some True \<Rightarrow> return Sat
+  }
+  | None \<Rightarrow> return Renaming_Failed
+}
+"
+
+(*
+definition rename_mc where
+  "rename_mc dc broadcast bounds' automata k L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks
+\<equiv>
+let
+   _ = println (STR ''Checking renaming'');
+  formula = (if dc then formula.EX (not sexp.true) else formula);
+   renaming_valid = Simple_Network_Rename_Formula_String_Defs.check_renaming
+      broadcast bounds' renum_vars renum_clocks renum_states automata formula L0 s0;
+   _ = println (STR ''Renaming network'');
+   (broadcast, automata, bounds') = rename_network
+      broadcast bounds' automata renum_acts renum_vars renum_clocks renum_states;
+   _ = println (STR ''Automata after renaming'');
+   _ = map (\<lambda>a. show a |> String.implode |> println) automata;
+   _ = println (STR ''Renaming formula'');
+   formula =
+    (if dc then formula.EX (not sexp.true) else map_formula renum_states renum_vars id formula);
+    _ = println (STR ''Renaming state'');
+   L0 = map_index renum_states L0;
+   s0 = map (\<lambda>(x, v). (renum_vars x, v)) s0;
+   show_clock = show o inv_renum_clocks;
+   show_state = show_state inv_renum_states inv_renum_vars
+in
+  if is_result renaming_valid then do {
+    let _ = println (STR ''Checking preconditions'');
+    let r = Simple_Network_Impl_nat_defs.check_precond
+      broadcast bounds' automata m num_states num_actions k L0 s0 formula;
+    let _ = (case r of Result _ \<Rightarrow> [()]
+      | Error es \<Rightarrow> let _ = println (STR ''The following pre-conditions were not satisified'') in
+          map println es);
+    let _ = println (STR ''Running precond_mc'');
+    r \<leftarrow> (if dc
+      then precond_dc show_clock show_state
+        broadcast bounds' automata m num_states num_actions k L0 s0 formula
+      else precond_mc show_clock show_state
+        broadcast bounds' automata m num_states num_actions k L0 s0 formula);
+    case r of
+      None \<Rightarrow> return Preconds_Unsat
+    | Some False \<Rightarrow> return Unsat
+    | Some True \<Rightarrow> return Sat
+  } 
+  else do {
+    let _ = println (STR ''The following conditions on the renaming were not satisfied:'');
+    let _ = the_errors renaming_valid |> map println;
+    return Renaming_Failed}
+"
+*)
+
+theorem model_check_rename:
+  "<emp> rename_mc False broadcast bounds automata k urge L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks
+    <\<lambda> Sat \<Rightarrow> \<up>(
+        (\<not> has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_ . 0) \<longrightarrow>
+          N broadcast automata bounds,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> formula
+        ))
+     | Unsat \<Rightarrow> \<up>(
+        (\<not> has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_ . 0) \<longrightarrow>
+          \<not> N broadcast automata bounds,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> formula
+        ))
+     | Renaming_Failed \<Rightarrow> \<up>(\<not> Simple_Network_Rename_Formula
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata formula)
+     | Preconds_Unsat \<Rightarrow> \<up>(\<not> Simple_Network_Impl_nat_ceiling_start_state
+        (map renum_acts broadcast)
+        (map (\<lambda>(a,p). (renum_vars a, p)) bounds)
+        (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+          (map (conv_urge urge) automata))
+        m num_states num_actions k
+        (map_index renum_states L0) (map (\<lambda>(x, v). (renum_vars x, v)) s0)
+        (map_formula renum_states renum_vars id formula))
+    >t"
+proof -
+  have *: "
+    Simple_Network_Rename_Formula_String
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states automata urge formula s0 L0
+  = Simple_Network_Rename_Formula
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata formula
+  "
+    unfolding
+      Simple_Network_Rename_Formula_String_def Simple_Network_Rename_Formula_def
+      Simple_Network_Rename_Start_def Simple_Network_Rename_Start_axioms_def
+      Simple_Network_Rename_def Simple_Network_Rename_Formula_axioms_def
+    using infinite_literal by auto
+  define A where "A \<equiv> N broadcast automata bounds"
+  define check where "check \<equiv> A,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> formula"
+  define A' where "A' \<equiv> N
+    (map renum_acts broadcast)
+    (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+      (map (conv_urge urge) automata))
+    (map (\<lambda>(a,p). (renum_vars a, p)) bounds)"
+  define check' where "check' \<equiv>
+    A',(map_index renum_states L0, map_of (map (\<lambda>(x, v). (renum_vars x, v)) s0), \<lambda>_ . 0) \<Turnstile>
+    map_formula renum_states renum_vars id formula"
+  define preconds_sat where "preconds_sat \<equiv>
+    Simple_Network_Impl_nat_ceiling_start_state
+      (map renum_acts broadcast)
+      (map (\<lambda>(a,p). (renum_vars a, p)) bounds)
+      (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+        (map (conv_urge urge) automata))
+      m num_states num_actions k
+      (map_index renum_states L0) (map (\<lambda>(x, v). (renum_vars x, v)) s0)
+      (map_formula renum_states renum_vars id formula)"
+  define renaming_valid where "renaming_valid \<equiv>
+    Simple_Network_Rename_Formula
+      broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata formula
+  "
+  have [simp]: "check \<longleftrightarrow> check'" 
+    if renaming_valid
+    using that unfolding check_def check'_def A_def A'_def renaming_valid_def
+    by (rule Simple_Network_Rename_Formula.models_iff'[symmetric])
+  have test[symmetric, simp]:
+    "Simple_Network_Language_Model_Checking.has_deadlock A (L0, map_of s0, \<lambda>_. 0)
+  \<longleftrightarrow>Simple_Network_Language_Model_Checking.has_deadlock A'
+     (map_index renum_states L0, map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0), \<lambda>_. 0)
+  " if renaming_valid
+    using that unfolding check_def check'_def A_def A'_def renaming_valid_def
+    unfolding Simple_Network_Rename_Formula_def
+    by (elim conjE) (rule Simple_Network_Rename_Start.has_deadlock_iff'[symmetric])
+  note [sep_heap_rules] =
+    model_check[
+    of _ _
+    "map renum_acts broadcast" "map (\<lambda>(a,p). (renum_vars a, p)) bounds"
+    "map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+      (map (conv_urge urge) automata)"
+    m num_states num_actions k "map_index renum_states L0" "map (\<lambda>(x, v). (renum_vars x, v)) s0"
+    "map_formula renum_states renum_vars id formula",
+    folded A'_def preconds_sat_def renaming_valid_def, folded check'_def, simplified
+    ]
+  show ?thesis
+    unfolding rename_mc_def do_rename_mc_def rename_network_def
+    unfolding if_False
+    unfolding Simple_Network_Rename_Formula_String_Defs.check_renaming[symmetric] * Let_def
+    unfolding
+      A_def[symmetric] check_def[symmetric]
+      preconds_sat_def[symmetric] renaming_valid_def[symmetric]
+    by (sep_auto simp: model_checker.refine[symmetric] split: bool.splits)
+qed
+
+theorem deadlock_check_rename:
+  "<emp> rename_mc True broadcast bounds automata k urge L0 s0 formula
+    m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+    inv_renum_states inv_renum_vars inv_renum_clocks
+    <\<lambda> Sat   \<Rightarrow> \<up>(  has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_.  0))
+     | Unsat \<Rightarrow> \<up>(\<not> has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_. 0))
+     | Renaming_Failed \<Rightarrow> \<up>(\<not> Simple_Network_Rename_Formula
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata
+        (formula.EX (not sexp.true)))
+     | Preconds_Unsat \<Rightarrow> \<up>(\<not> Simple_Network_Impl_nat_ceiling_start_state
+        (map renum_acts broadcast)
+        (map (\<lambda>(a,p). (renum_vars a, p)) bounds)
+        (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+          (map (conv_urge urge) automata))
+        m num_states num_actions k
+        (map_index renum_states L0) (map (\<lambda>(x, v). (renum_vars x, v)) s0)
+        (formula.EX (not sexp.true)))
+    >t"
+proof -
+  have *: "
+    Simple_Network_Rename_Formula_String
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states automata urge
+        (formula.EX (not sexp.true)) s0 L0
+  = Simple_Network_Rename_Formula
+        broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata
+        (formula.EX (not sexp.true))
+  "
+    unfolding
+      Simple_Network_Rename_Formula_String_def Simple_Network_Rename_Formula_def
+      Simple_Network_Rename_Start_def Simple_Network_Rename_Start_axioms_def
+      Simple_Network_Rename_def Simple_Network_Rename_Formula_axioms_def
+    using infinite_literal by auto
+  define A where "A \<equiv> N broadcast automata bounds"
+  define A' where "A' \<equiv> N
+    (map renum_acts broadcast)
+    (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+      (map (conv_urge urge) automata))
+    (map (\<lambda>(a,p). (renum_vars a, p)) bounds)"
+  define preconds_sat where "preconds_sat \<equiv>
+    Simple_Network_Impl_nat_ceiling_start_state
+      (map renum_acts broadcast)
+      (map (\<lambda>(a,p). (renum_vars a, p)) bounds)
+      (map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+        (map (conv_urge urge) automata))
+      m num_states num_actions k
+      (map_index renum_states L0) (map (\<lambda>(x, v). (renum_vars x, v)) s0)
+      (formula.EX (not sexp.true))"
+  define renaming_valid where "renaming_valid \<equiv>
+    Simple_Network_Rename_Formula
+      broadcast bounds renum_acts renum_vars renum_clocks renum_states urge s0 L0 automata
+      (formula.EX (not sexp.true))"
+ have test[symmetric, simp]:
+    "Simple_Network_Language_Model_Checking.has_deadlock A (L0, map_of s0, \<lambda>_. 0)
+  \<longleftrightarrow>Simple_Network_Language_Model_Checking.has_deadlock A'
+     (map_index renum_states L0, map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0), \<lambda>_. 0)
+  " if renaming_valid
+    using that unfolding check_def A_def A'_def renaming_valid_def Simple_Network_Rename_Formula_def
+    by (elim conjE) (rule Simple_Network_Rename_Start.has_deadlock_iff'[symmetric])
+  note [sep_heap_rules] =
+    deadlock_check[
+    of _ _
+    "map renum_acts broadcast" "map (\<lambda>(a,p). (renum_vars a, p)) bounds"
+    "map_index (renum_automaton renum_acts renum_vars renum_clocks renum_states)
+      (map (conv_urge urge) automata)"
+    m num_states num_actions k "map_index renum_states L0" "map (\<lambda>(x, v). (renum_vars x, v)) s0",
+    folded preconds_sat_def A'_def renaming_valid_def,
+    simplified
+    ]
+  show ?thesis
+    unfolding rename_mc_def do_rename_mc_def  rename_network_def
+    unfolding if_True
+    unfolding Simple_Network_Rename_Formula_String_Defs.check_renaming[symmetric] * Let_def
+    unfolding A_def[symmetric] preconds_sat_def[symmetric] renaming_valid_def[symmetric]
+    by (sep_auto simp: deadlock_checker.refine[symmetric] split: bool.splits)
+qed
+
+
+paragraph \<open>Code Setup for the Model Checker\<close>
+
+lemmas [code] =
+  reachability_checker_def
+  Alw_ev_checker_def
+  leadsto_checker_def
+  model_checker_def[unfolded PR_CONST_def]
+
+lemmas [code] =
+  Prod_TA_Defs.n_ps_def
+  Simple_Network_Impl_Defs.n_vs_def
+  automaton_of_def
+  Simple_Network_Impl_nat_defs.pairs_by_action_impl_def
+  Simple_Network_Impl_nat_defs.all_actions_from_vec_def
+  Simple_Network_Impl_nat_defs.all_actions_by_state_def
+  Simple_Network_Impl_nat_defs.compute_upds_impl_def
+  Simple_Network_Impl_nat_defs.actions_by_state_def
+  Simple_Network_Impl_nat_defs.check_boundedi_def
+  Simple_Network_Impl_nat_defs.get_committed_def
+  Simple_Network_Impl_nat_defs.make_combs_def
+  Simple_Network_Impl_nat_defs.trans_map_def
+  Simple_Network_Impl_nat_defs.actions_by_state'_def
+  Simple_Network_Impl_nat_defs.bounds_map_def
+  Simple_Network_Impl_nat_defs.bin_actions_def
+  mk_updsi_def
+
+lemma (in Simple_Network_Impl_nat_defs) bounded_s0_iff:
+  "bounded bounds (map_of s0) \<longleftrightarrow> bounded (map_of bounds') (map_of s0)"
+  unfolding bounds_def snd_conv ..
+
+lemma int_Nat_range_iff:
+  "(n :: int) \<in> \<nat> \<longleftrightarrow> n \<ge> 0" for n
+  using zero_le_imp_eq_int unfolding Nats_def by auto
+
+lemmas [code] =
+  Simple_Network_Impl_nat_ceiling_start_state_def
+  Simple_Network_Impl_nat_ceiling_start_state_axioms_def[
+    unfolded Simple_Network_Impl_nat_defs.bounded_s0_iff]
+  Simple_Network_Impl_nat_def[unfolded int_Nat_range_iff]
+  Simple_Network_Impl_nat_urge_def
+  Simple_Network_Impl_nat_urge_axioms_def
+
+lemmas [code_unfold] = bounded_def dom_map_of_conv_image_fst
+
+export_code Simple_Network_Impl_nat_ceiling_start_state_axioms
+
+export_code precond_mc in SML module_name Test
+
+export_code precond_dc checking SML
+
+
+paragraph \<open>Code Setup for Renaming\<close>
+
+lemmas [code] =
+  Simple_Network_Rename_Formula_String_def
+  Simple_Network_Impl.clk_set'_def
+  Simple_Network_Impl.clkp_set'_def
+
+lemmas [code] =
+  Simple_Network_Rename_Defs.renum_automaton_def
+  Simple_Network_Rename_Defs.renum_cconstraint_def
+  Simple_Network_Rename_Defs.map_cconstraint_def
+  Simple_Network_Rename_Defs.renum_reset_def
+  Simple_Network_Rename_Defs.renum_upd_def
+  Simple_Network_Rename_Defs.renum_act_def
+  Simple_Network_Rename_Defs.renum_exp_def
+  Simple_Network_Rename_Defs.renum_bexp_def
+  Simple_Network_Rename_Formula_String_Defs.check_renaming_def
+  Simple_Network_Impl_nat_defs.check_precond_def
+  Simple_Network_Impl_nat_defs.check_precond1_def[unfolded int_Nat_range_iff]
+  Simple_Network_Impl_nat_defs.check_precond2_def[
+    unfolded Simple_Network_Impl_nat_defs.bounded_s0_iff]
+
+lemma (in Prod_TA_Defs) states_mem_iff:
+  "L \<in> states \<longleftrightarrow> length L = n_ps \<and> (\<forall> i. i < n_ps \<longrightarrow>
+    (\<exists> (l, b, g, a, r, u, l') \<in> fst (snd (snd (fst (snd A) ! i))). L ! i = l \<or> L ! i = l'))"
+  unfolding states_def trans_def N_def by (auto split: prod.split)
+
+lemmas [code_unfold] =
+  Prod_TA_Defs.states_mem_iff
+  Simple_Network_Impl.act_set_compute
+  Simple_Network_Impl_Defs.var_set_compute
+  Simple_Network_Impl_Defs.loc_set_compute
+  setcompr_eq_image
+  Simple_Network_Impl.length_automata_eq_n_ps[symmetric]
+
+export_code rename_mc in SML module_name Test
+
+
+
+paragraph \<open>Calculating the Clock Ceiling\<close>
+
+context Simple_Network_Impl_nat_defs
+begin
+
+definition "clkp_inv i l \<equiv>
+  \<Union>g \<in> set (filter (\<lambda>(a, b). a = l) (snd (snd (snd (automata ! i))))). collect_clock_pairs (snd g)"
+
+definition "clkp_set'' i l \<equiv>
+    clkp_inv i l \<union> (\<Union> (l', b, g, _) \<in> set (fst (snd (snd (automata ! i)))).
+      if l' = l then collect_clock_pairs g else {})"
+
+definition
+  "collect_resets i l = (\<Union> (l', b, g, a, f, r, _) \<in> set (fst (snd (snd (automata ! i)))).
+    if l' = l then set r else {})"
+
+context
+  fixes q c :: nat
+begin
+
+  definition "n \<equiv> num_states q"
+
+  definition "V \<equiv> \<lambda> v. v \<le> n"
+
+  definition "
+    bound_g l \<equiv>
+      Max ({0} \<union> \<Union> ((\<lambda> (x, d). if x = c then {d} else {}) ` clkp_set'' q l))
+    "
+
+  definition "
+    bound_inv l \<equiv>
+      Max ({0} \<union> \<Union> ((\<lambda> (x, d). if x = c then {d} else {}) ` clkp_inv q l))
+  "
+
+  definition "
+    bound l \<equiv> max (bound_g l) (bound_inv l)
+  "
+
+definition "
+  resets l \<equiv>
+    fold
+    (\<lambda> (l1, b, g, a, f, r, l') xs. if l1 \<noteq> l \<or> l' \<in> set xs \<or> c \<in> set r then xs else (l' # xs))
+    (fst (snd (snd (automata ! q))))
+    []
+"
+
+text \<open>
+  Edges in the direction nodes to single sink.
+\<close>
+definition "
+  E' l \<equiv> resets l
+"
+
+text \<open>
+  Turning around the edges to obtain a single source shortest paths problem.
+\<close>
+definition "
+  E l \<equiv> if l = n then [0..<n] else filter (\<lambda> l'. l \<in> set (E' l')) [0..<n]
+"
+
+text \<open>
+  Weights already turned around.
+\<close>
+definition "
+  W l l' \<equiv> if l = n then - bound l' else 0
+"
+
+definition G where "
+  G \<equiv> \<lparr> gi_V = V, gi_E = E, gi_V0 = [n], \<dots> = W \<rparr>
+"
+
+definition "
+  local_ceiling_single \<equiv>
+  let
+    w = calc_shortest_scc_paths G n
+  in
+    map (\<lambda> x. case x of None \<Rightarrow> 0 | Some x \<Rightarrow> nat(-x)) w
+"
+
+end
+
+definition "
+  local_ceiling \<equiv>
+    rev $
+    fold
+      (\<lambda> q xs.
+        (\<lambda> x. rev x # xs) $
+        fold
+          (\<lambda> l xs.
+            (\<lambda> x. (0 # rev x) # xs) $
+            fold
+              (\<lambda> c xs. local_ceiling_single q c ! l # xs)
+              [1..<Suc m]
+              []
+          )
+          [0..<n q]
+          []
+      )
+      [0..<n_ps]
+      []
+"
+
+end
+
+
+lemmas [code] =
+  Simple_Network_Impl_nat_defs.local_ceiling_def
+  Simple_Network_Impl_nat_defs.local_ceiling_single_def
+  Simple_Network_Impl_nat_defs.n_def
+  Simple_Network_Impl_nat_defs.G_def
+  Simple_Network_Impl_nat_defs.W_def
+  Simple_Network_Impl_nat_defs.V_def
+  Simple_Network_Impl_nat_defs.E'_def
+  Simple_Network_Impl_nat_defs.E_def
+  Simple_Network_Impl_nat_defs.resets_def
+  Simple_Network_Impl_nat_defs.bound_def
+  Simple_Network_Impl_nat_defs.bound_inv_def
+  Simple_Network_Impl_nat_defs.bound_g_def
+  Simple_Network_Impl_nat_defs.collect_resets_def
+  Simple_Network_Impl_nat_defs.clkp_set''_def
+  Simple_Network_Impl_nat_defs.clkp_inv_def
+
+export_code Simple_Network_Impl_nat_defs.local_ceiling checking SML_imp
+
+
+
+paragraph \<open>Calculating the Renaming\<close>
+
+definition "mem_assoc x = list_ex (\<lambda>(y, _). x = y)"
+
+definition "mk_renaming str xs \<equiv>
+do {
+  mapping \<leftarrow> fold_error
+    (\<lambda>x m.
+      if mem_assoc x m then Error [STR ''Duplicate name: '' + str x] else (x,length m) # m |> Result
+    ) xs [];
+  Result (let
+    m = map_of mapping;
+    f = (\<lambda>x.
+      case m x of
+        None \<Rightarrow> let _ = println (STR ''Key error: '' + str x) in undefined
+      | Some v \<Rightarrow> v);
+    m = map_of (map prod.swap mapping);
+    f_inv = (\<lambda>x.
+      case m x of
+        None \<Rightarrow> let _ = println (STR ''Key error: '' + String.implode (show x)) in undefined
+      | Some v \<Rightarrow> v)
+  in (f, f_inv)
+  )
+}"
+
+definition
+  "extend_domain m d n \<equiv>
+    let
+      (i, xs) = fold
+        (\<lambda>x (i, xs). if x \<in> set d then (i + 1, (x, i + 1) # xs) else (i, xs)) d (n, []);
+      m' = map_of xs
+    in
+      (\<lambda>x. if x \<in> set d then the (m' x) else m x)"
+
+(* Unused *)
+lemma is_result_make_err:
+  "is_result (make_err m e) \<longleftrightarrow> False"
+  by (cases e) auto
+
+(* Unused *)
+lemma is_result_assert_msg_iff:
+  "is_result (assert_msg b m r) \<longleftrightarrow> is_result r \<and> b"
+  unfolding assert_msg_def by (simp add: is_result_make_err)
+
+context Simple_Network_Impl
+begin
+
+definition action_set where
+  "action_set \<equiv>
+    (\<Union>(_, _, trans, _) \<in> set automata. \<Union>(_, _, _, a, _, _, _) \<in> set trans. set_act a)
+    \<union> set broadcast"
+
+definition loc_set' where
+  "loc_set' p = (\<Union>(l, _, _, _, _, _, l')\<in>set (fst (snd (snd (automata ! p)))). {l, l'})" for p
+
+end
+
+
+definition
+  "concat_str = String.implode o concat o map String.explode"
+
+
+paragraph \<open>Unsafe Glue Code\<close>
+definition list_of_set' :: "'a set \<Rightarrow> 'a list" where
+  "list_of_set' xs = undefined"
+
+definition list_of_set :: "'a set \<Rightarrow> 'a list" where
+  "list_of_set xs = list_of_set' xs |> remdups"
+
+code_printing
+  constant list_of_set' \<rightharpoonup> (SML) "(fn Set xs => xs) _"
+       and              (OCaml) "(fun x -> match x with Set xs -> xs) _"
+
+
+
+definition
+  "mk_renaming' xs \<equiv> mk_renaming (String.implode o show) xs"
+
+definition "make_renaming \<equiv> \<lambda> broadcast automata bounds.
+  let
+    action_set = Simple_Network_Impl.action_set automata broadcast |> list_of_set;
+    clk_set = Simple_Network_Impl.clk_set' automata |> list_of_set;
+    clk_set = clk_set @ [STR ''_urge''];
+    loc_set' = (\<lambda>i. Simple_Network_Impl.loc_set' automata i |> list_of_set);
+    loc_set = Prod_TA_Defs.loc_set
+      (set broadcast, map automaton_of automata, map_of bounds);
+    loc_set_diff = (\<lambda>i. loc_set - Simple_Network_Impl.loc_set' automata i |> list_of_set);
+    loc_set = list_of_set loc_set;
+    var_set = Prod_TA_Defs.var_set
+      (set broadcast, map automaton_of automata, map_of bounds) |> list_of_set;
+    n_ps = length automata;
+    num_actions = length action_set;
+    m = length (remdups clk_set);
+    num_states_list = map (\<lambda>i. loc_set' i |> remdups |> length) [0..<n_ps];
+    num_states = (\<lambda>i. num_states_list ! i);
+    mk_renaming = mk_renaming (\<lambda>x. x)
+  in do {
+    ((renum_acts, _), (renum_clocks, inv_renum_clocks), (renum_vars, inv_renum_vars)) \<leftarrow>
+      mk_renaming action_set <|> mk_renaming clk_set <|> mk_renaming var_set;
+    let renum_clocks = Suc o renum_clocks;
+    let inv_renum_clocks = (\<lambda>c. if c = 0 then STR ''0'' else inv_renum_clocks (c - 1));
+    renum_states_list' \<leftarrow> combine_map (\<lambda>i. mk_renaming' (loc_set' i)) [0..<n_ps];
+    let renum_states_list = map fst renum_states_list';
+    let renum_states_list = map_index
+      (\<lambda>i m. extend_domain m (loc_set_diff i) (length (loc_set' i))) renum_states_list;
+    let renum_states = (\<lambda>i. renum_states_list ! i);
+    let inv_renum_states = (\<lambda>i. map snd renum_states_list' ! i);
+    assert (fst ` set bounds \<subseteq> set var_set)
+      STR ''State variables are declared but do not appear in model'';
+    Result (m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
+      inv_renum_states, inv_renum_vars, inv_renum_clocks)
+  }"
+
+definition "preproc_mc \<equiv> \<lambda>dc ids_to_names (broadcast, automata, bounds) L0 s0 formula.
+  let _ = println (STR ''Make renaming'') in
+  case make_renaming broadcast automata bounds of
+    Error e \<Rightarrow> return (Error e)
+  | Result (m, num_states, num_actions, renum_acts, renum_vars, renum_clocks, renum_states,
+      inv_renum_states, inv_renum_vars, inv_renum_clocks) \<Rightarrow> do {
+    let _ = println (STR ''Renaming'');
+    let (broadcast', automata', bounds') = rename_network
+      broadcast bounds automata renum_acts renum_vars renum_clocks renum_states;
+    let _ = println (STR ''Calculating ceiling'');
+    let k = Simple_Network_Impl_nat_defs.local_ceiling broadcast' bounds' automata' m num_states;
+    let _ = println (STR ''Running model checker'');
+    let inv_renum_states = (\<lambda>i. ids_to_names i o inv_renum_states i);
+    r \<leftarrow> rename_mc dc broadcast bounds automata k STR ''_urge'' L0 s0 formula
+      m num_states num_actions renum_acts renum_vars renum_clocks renum_states
+      inv_renum_states inv_renum_vars inv_renum_clocks;
+    return (Result r)
+  }
+"
+
+definition
+  "err s = Error [s]"
+
+definition
+"do_preproc_mc \<equiv> \<lambda>dc ids_to_names (broadcast, automata, bounds) L0 s0 formula. do {
+  r \<leftarrow> preproc_mc dc ids_to_names (broadcast, automata, bounds) L0 s0 formula;
+  return (case r of
+    Error es \<Rightarrow>
+      intersperse (STR ''\<newline>'') es
+      |> concat_str
+      |> (\<lambda>e. STR ''Error during preprocessing:\<newline>'' + e)
+      |> err
+  | Result Renaming_Failed \<Rightarrow> STR ''Renaming failed'' |> err
+  | Result Preconds_Unsat \<Rightarrow> STR ''Input invalid'' |> err
+  | Result Unsat \<Rightarrow>
+      (if dc then STR ''Model has no deadlock!'' else STR ''Property is not satisfied!'') |> Result
+  | Result Sat \<Rightarrow>
+      (if dc then STR ''Model has a deadlock!''  else STR ''Property is satisfied!'') |> Result
+  )
+}"
+
+lemmas [code] =
+  Simple_Network_Impl.action_set_def
+  Simple_Network_Impl.loc_set'_def
+
+export_code do_preproc_mc in SML module_name Main
+
+definition parse where
+  "parse parser s \<equiv> case parse_all lx_ws parser s of
+    Inl e \<Rightarrow> Error [e () ''Parser: '' |> String.implode]
+  | Inr r \<Rightarrow> Result r"
+
+definition get_nat :: "string \<Rightarrow> JSON \<Rightarrow> nat Error_List_Monad.result" where
+  "get_nat s json \<equiv> case json of
+    Object as \<Rightarrow> Error []
+  | _ \<Rightarrow> Error [STR ''JSON Get: expected object'']
+" for json
+
+definition of_object :: "JSON \<Rightarrow> (string \<rightharpoonup> JSON) Error_List_Monad.result" where
+  "of_object json \<equiv> case json of
+    Object as \<Rightarrow> map_of as |> Result
+  | _ \<Rightarrow> Error [STR ''json_to_map: expected object'']
+" for json
+
+definition get where
+  "get m x \<equiv> case m x of
+    None \<Rightarrow> Error [STR ''(Get) key not found: '' + String.implode (show x)]
+  | Some a \<Rightarrow> Result a"
+
+definition
+  "get_default def m x \<equiv> case m x of None \<Rightarrow> def | Some a \<Rightarrow> a"
+
+definition default where
+  "default def x \<equiv> case x of Result s \<Rightarrow> s | Error e \<Rightarrow> def"
+
+definition of_array where
+  "of_array json \<equiv> case json of
+    JSON.Array s \<Rightarrow> Result s
+  | _ \<Rightarrow> Error [STR ''of_array: expected sequence'']
+" for json
+
+definition of_string where
+  "of_string json \<equiv> case json of
+    JSON.String s \<Rightarrow> Result (String.implode s)
+  | _ \<Rightarrow> Error [STR ''of_array: expected sequence'']
+" for json
+
+definition of_nat where
+  "of_nat json \<equiv> case json of
+    JSON.Nat n \<Rightarrow> Result n
+  | _ \<Rightarrow> Error [STR ''of_nat: expected natural number'']
+" for json
+
+definition of_int where
+  "of_int json \<equiv> case json of
+    JSON.Int n \<Rightarrow> Result n
+  | _ \<Rightarrow> Error [STR ''of_int: expected integral number'']
+" for json
+
+definition [consuming]:
+  "lx_underscore = exactly ''_'' with (\<lambda>_. CHR ''_'')"
+
+definition [consuming]:
+  "lx_hyphen = exactly ''-'' with (\<lambda>_. CHR ''-'')"
+
+definition [consuming]:
+  "ta_var_ident \<equiv>
+    (lx_alphanum \<parallel> lx_underscore)
+    -- Parser_Combinator.repeat (lx_alphanum \<parallel> lx_underscore \<parallel> lx_hyphen)
+    with uncurry (#)
+  "
+
+definition [consuming]:
+  "parse_bound \<equiv> ta_var_ident --
+    exactly ''['' *-- lx_int -- exactly '':'' *-- lx_int --* exactly '']''"
+
+definition "parse_bounds \<equiv> parse_list' (lx_ws *-- parse_bound with (\<lambda>(s,p). (String.implode s, p)))"
+
+lemma
+  "parse parse_bounds (STR ''id[-1:2], id[-1:0]'')
+ = Result [(STR ''id'', -1, 2), (STR ''id'', -1, 0)]"
+  by eval
+
+lemma "parse parse_bounds (STR '''') = Result []"
+  by eval
+
+definition [consuming]:
+  "scan_var = ta_var_ident"
+
+abbreviation seq_ignore_left_ws (infixr "**--" 60)
+  where "p **-- q \<equiv> token p *-- q" for p q
+
+abbreviation seq_ignore_right_ws (infixr "--**" 60)
+  where "p --** q \<equiv> token p --* q" for p q
+
+abbreviation seq_ws (infixr "---" 60)
+  where "seq_ws p q \<equiv> token p -- q" for p q
+
+definition scan_acconstraint where [unfolded Let_def, consuming]:
+  "scan_acconstraint \<equiv>
+    let scan =
+      (\<lambda>s c. scan_var --- exactly s **-- token lx_int with (\<lambda>(x, y). c (String.implode x) y)) in
+  (
+    scan ''<''  lt \<parallel>
+    scan ''<='' le \<parallel>
+    scan ''=='' eq \<parallel>
+    scan ''=''  eq \<parallel>
+    scan ''>='' ge \<parallel>
+    scan ''>''  gt
+  )
+"
+
+definition [consuming]:
+  "scan_parens lparen rparen inner \<equiv> exactly lparen **-- inner --** token (exactly rparen)"
+
+definition [consuming]: "scan_loc \<equiv>
+  (scan_var --- (exactly ''.'' *-- scan_var))
+  with (\<lambda> (p, s). loc (String.implode p) (String.implode s))"
+
+definition [consuming]: "scan_bexp_elem \<equiv> scan_acconstraint \<parallel> scan_loc"
+
+abbreviation "scan_parens' \<equiv> scan_parens ''('' '')''"
+
+definition [consuming]: "scan_infix_pair a b s \<equiv> a --- exactly s **-- token b"
+
+lemma [fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> A l2 = A' l2"
+  assumes "\<And>l2. ll_fuel l2 + (if length s > 0 then 1 else 0) \<le> ll_fuel l' \<Longrightarrow> B l2 = B' l2"
+  assumes "s = s'" "l=l'"
+  shows "scan_infix_pair A B s l = scan_infix_pair A' B' s' l'"
+  using assms unfolding scan_infix_pair_def gen_token_def
+  by (cases s; intro Parser_Combinator.bind_cong repeat_cong assms) auto
+
+lemma [fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 < ll_fuel l \<Longrightarrow> A l2 = A' l2" "l = l'"
+  shows "scan_parens' A l = scan_parens' A' l'"
+  using assms unfolding scan_parens_def gen_token_def
+  by (intro Parser_Combinator.bind_cong repeat_cong assms) auto
+
+lemma token_cong[fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l \<Longrightarrow> A l2 = A' l2" "l = l'"
+  shows "token A l = token A' l'"
+  using assms unfolding scan_parens_def gen_token_def
+  by (intro Parser_Combinator.bind_cong repeat_cong assms) auto
+
+lemma is_cparser_scan_parens'[parser_rules]:
+  "is_cparser (scan_parens' a)"
+  unfolding scan_parens_def by simp
+
+fun aexp and mexp and scan_exp and scan_7 and scan_6 and scan_0 where \<comment> \<open>slow: 90s\<close>
+  "aexp ::=
+  token lx_int with exp.const \<parallel> token scan_var with exp.var o String.implode \<parallel>
+  scan_parens' (scan_exp --- exactly ''?'' **-- scan_7 --- exactly '':'' **-- scan_exp)
+  with (\<lambda> (e1, b, e2). exp.if_then_else b e1 e2) \<parallel>
+  tk_lparen **-- scan_exp --** tk_rparen"
+| "mexp ::= chainL1 aexp (multiplicative_op with (\<lambda>f a b. exp.binop f a b))"
+| "scan_exp ::= chainL1 mexp (additive_op with (\<lambda>f a b. exp.binop f a b))"
+| "scan_7 ::=
+    scan_infix_pair scan_6 scan_7 ''->'' with uncurry bexp.imply \<parallel>
+    scan_infix_pair scan_6 scan_7 ''||'' with uncurry bexp.or \<parallel>
+    scan_6" |
+  "scan_6 ::=
+    scan_infix_pair scan_0 scan_6 ''&&'' with uncurry bexp.and \<parallel>
+    scan_0" |
+  "scan_0 ::=
+    (exactly ''~'' \<parallel> exactly ''!'') **-- scan_parens' scan_7 with bexp.not \<parallel>
+    token (exactly ''true'') with (\<lambda>_. bexp.true) \<parallel>
+    scan_infix_pair aexp aexp ''<='' with uncurry bexp.le \<parallel>
+    scan_infix_pair aexp aexp ''<''  with uncurry bexp.lt \<parallel>
+    scan_infix_pair aexp aexp ''=='' with uncurry bexp.eq \<parallel>
+    scan_infix_pair aexp aexp ''>''  with uncurry bexp.gt \<parallel>
+    scan_infix_pair aexp aexp ''>='' with uncurry bexp.ge \<parallel>
+    scan_parens' scan_7"
+
+context
+  fixes elem :: "(char, 'bexp) parser"
+    and Imply Or And :: "'bexp \<Rightarrow> 'bexp \<Rightarrow> 'bexp"
+    and Not :: "'bexp \<Rightarrow> 'bexp"
+begin
+
+fun scan_7' and scan_6' and scan_0' where
+  "scan_7' ::=
+    scan_infix_pair scan_6' scan_7' ''->'' with uncurry Imply \<parallel>
+    scan_infix_pair scan_6' scan_7' ''||'' with uncurry Or \<parallel>
+    scan_6'" |
+  "scan_6' ::=
+    scan_infix_pair scan_0' scan_6' ''&&'' with uncurry And \<parallel>
+    scan_0'" |
+  "scan_0' ::=
+    (exactly ''~'' \<parallel> exactly ''!'') **-- scan_parens' scan_7' with Not \<parallel>
+    elem \<parallel>
+    scan_parens' scan_7'"
+
+context
+  assumes [parser_rules]: "is_cparser elem"
+begin
+
+lemma [parser_rules]:
+  "is_cparser scan_0'"
+  by (simp add: scan_0'.simps[abs_def])
+
+lemma [parser_rules]:
+  "is_cparser scan_6'"
+  by (subst scan_6'.simps[abs_def]) simp
+
+end
+end
+
+abbreviation "scan_bexp \<equiv> scan_7' scan_bexp_elem sexp.imply sexp.or sexp.and sexp.not"
+
+lemma [parser_rules]:
+  "is_cparser scan_bexp"
+  by (subst scan_7'.simps[abs_def]) simp
+
+lemma "parse scan_bexp (STR ''a < 3 && b>=2 || ~ (c <= 4)'')
+= Result (sexp.or (and (lt STR ''a'' 3) (ge STR ''b'' 2)) (not (sexp.le STR ''c'' 4)))"
+  by eval
+
+definition [consuming]: "scan_prefix p head = exactly head **-- p" for p
+
+definition [consuming]: "scan_formula \<equiv>
+  scan_prefix scan_bexp ''E<>'' with formula.EX \<parallel>
+  scan_prefix scan_bexp ''E[]'' with EG \<parallel>
+  scan_prefix scan_bexp ''A<>'' with AX \<parallel>
+  scan_prefix scan_bexp ''A[]'' with AG \<parallel>
+  scan_infix_pair scan_bexp scan_bexp ''-->'' with uncurry Leadsto"
+
+(* unused *)
+lemma is_cparser_token[parser_rules]:
+  "is_cparser (token a)" if "is_cparser a"
+  using that unfolding gen_token_def by simp
+
+definition [consuming]: "scan_action \<equiv>
+  (scan_var --* token (exactly ''?'')) with In o String.implode \<parallel>
+  (scan_var --* token (exactly ''!'')) with Out o String.implode \<parallel>
+  scan_var with Sil o String.implode"
+
+abbreviation orelse (infix "orelse" 58) where
+  "a orelse b \<equiv> default b a"
+
+definition
+  "parse_action s \<equiv> parse scan_action s orelse Sil (STR '''')"
+
+fun chop_sexp where
+  "chop_sexp clocks (and a b) (cs, es) =
+    chop_sexp clocks a (cs, es) |> chop_sexp clocks b" |
+  "chop_sexp clocks (eq a b) (cs, es) =
+    (if a \<in> set clocks then (eq a b # cs, es) else (cs, eq a b # es))" |
+  "chop_sexp clocks (le a b) (cs, es) =
+    (if a \<in> set clocks then (le a b # cs, es) else (cs, le a b # es))" |
+  "chop_sexp clocks (lt a b) (cs, es) =
+    (if a \<in> set clocks then (lt a b # cs, es) else (cs, lt a b # es))" |
+  "chop_sexp clocks (ge a b) (cs, es) =
+    (if a \<in> set clocks then (ge a b # cs, es) else (cs, ge a b # es))" |
+  "chop_sexp clocks (gt a b) (cs, es) =
+    (if a \<in> set clocks then (gt a b # cs, es) else (cs, gt a b # es))" |
+  "chop_sexp clocks a (cs, es) = (cs, a # es)"
+
+fun sexp_to_acconstraint :: "(String.literal, String.literal, String.literal, int) sexp \<Rightarrow> _" where
+  "sexp_to_acconstraint (lt a (b :: int)) = acconstraint.LT a b" |
+  "sexp_to_acconstraint (le a b) = acconstraint.LE a b" |
+  "sexp_to_acconstraint (eq a b) = acconstraint.EQ a b" |
+  "sexp_to_acconstraint (ge a b) = acconstraint.GE a b" |
+  "sexp_to_acconstraint (gt a b) = acconstraint.GT a b"
+
+no_notation top_assn ("true")
+
+fun sexp_to_bexp :: "(String.literal, String.literal, String.literal, int) sexp \<Rightarrow> _" where
+  "sexp_to_bexp (lt a (b :: int)) = bexp.lt (exp.var a) (exp.const b) |> Result" |
+  "sexp_to_bexp (le a b) = bexp.le (exp.var a) (exp.const b) |> Result" |
+  "sexp_to_bexp (eq a b) = bexp.eq (exp.var a) (exp.const b) |> Result" |
+  "sexp_to_bexp (ge a b) = bexp.ge (exp.var a) (exp.const b) |> Result" |
+  "sexp_to_bexp (gt a b) = bexp.gt (exp.var a) (exp.const b) |> Result" |
+  "sexp_to_bexp (and a b) =
+    do {a \<leftarrow> sexp_to_bexp a; b \<leftarrow> sexp_to_bexp b; bexp.and a b |> Result}" |
+  "sexp_to_bexp (sexp.or a b) =
+    do {a \<leftarrow> sexp_to_bexp a; b \<leftarrow> sexp_to_bexp b; bexp.or a b |> Result}" |
+  "sexp_to_bexp (imply a b) =
+    do {a \<leftarrow> sexp_to_bexp a; b \<leftarrow> sexp_to_bexp b; bexp.imply a b |> Result}" |
+  "sexp_to_bexp x        = Error [STR ''Illegal construct in binary operation'']"
+
+(*
+definition [consuming]: "scan_bexp_elem' \<equiv>
+  token (exactly ''true'') with (\<lambda>_. bexp.true) \<parallel>
+  scan_acconstraint with (\<lambda>b. case sexp_to_bexp b of Result b \<Rightarrow> b)"
+
+abbreviation "scan_bexp' \<equiv> scan_7 scan_bexp_elem' bexp.imply bexp.or bexp.and bexp.not"
+
+lemma [parser_rules]:
+  "is_cparser scan_bexp'"
+  by (subst scan_7.simps[abs_def]) simp
+
+lemma token_cong[fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l \<Longrightarrow> A l2 = A' l2" "l = l'"
+  shows "token A l = token A' l'"
+  using assms unfolding scan_parens_def gen_token_def
+  by (intro Parser_Combinator.bind_cong repeat_cong assms) auto
+*)
+
+(*
+abbreviation additive_op where "additive_op \<equiv> 
+  tk_plus \<then> Parser_Combinator.return (+)
+\<parallel> tk_minus \<then> Parser_Combinator.return (-)"
+
+  abbreviation "multiplicative_op \<equiv> 
+    tk_times \<then> return ( * )
+  \<parallel> tk_div \<then> return (div)"  
+  abbreviation "power_op \<equiv> 
+    tk_power \<then> return (\<lambda>a b. a^nat b)" \<comment> \<open>Note: Negative powers are treated as \<open>x0\<close>\<close>
+*)
+
+definition [consuming]:
+  "scan_update \<equiv>
+   scan_var --- (exactly ''='' \<parallel> exactly '':='') **-- scan_exp
+   with (\<lambda>(s, x). (String.implode s, x))"
+
+abbreviation "scan_updates \<equiv> parse_list scan_update"
+
+value "parse scan_updates (STR '' y2  := 0'')"
+(* = Result [(STR ''y2'', exp.const (0 :: int))] *)
+
+value "parse scan_updates (STR ''y2 := 0, x2 := 0'')"
+(* = Result [(STR ''y2'', exp.const 0), (STR ''x2'', exp.const 0)]" *)
+value "parse scan_exp (STR ''( 1 ? L == 0 : 0 )'')"
+(*  = Result (if_then_else (bexp.eq STR ''L'' 0) (exp.const 1) (exp.const 0)) *)
+
+definition compile_invariant where
+  "compile_invariant clocks vars inv \<equiv>
+    let
+      (cs, es) = chop_sexp clocks inv ([], []);
+      g = map sexp_to_acconstraint cs
+    in
+      if es = []
+      then Result (g, bexp.true)
+      else do {
+        let e = fold (and) (tl es) (hd es);
+        b \<leftarrow> sexp_to_bexp e;
+        assert (set_bexp b \<subseteq> set vars) (String.implode (''Unknown variable in bexp: '' @ show b));
+        Result (g, b)
+      }" for inv
+
+definition compile_invariant' where
+  "compile_invariant' clocks vars inv \<equiv>
+  if inv = STR '''' then
+    Result ([], bexp.true)
+  else do {
+    inv \<leftarrow> parse scan_bexp inv |> err_msg (STR ''Failed to parse guard in '' + inv);
+    compile_invariant clocks vars inv
+  }
+" for inv
+
+definition convert_node where
+  "convert_node clocks vars n \<equiv> do {
+    n \<leftarrow>  of_object n;
+    ID \<leftarrow> get n ''id'' \<bind> of_nat;
+    name \<leftarrow> get n ''name'' \<bind> of_string;
+    inv \<leftarrow> get n ''invariant'' \<bind> of_string;
+    (inv, inv_vars) \<leftarrow>
+      compile_invariant' clocks vars inv |> err_msg (STR ''Failed to parse invariant!'');
+    assert (case inv_vars of bexp.true \<Rightarrow> True | _ \<Rightarrow> False)
+      (STR ''State invariants on nodes are not supported'');
+    Result ((name, ID), inv)
+  }"
+
+definition convert_edge where
+  "convert_edge clocks vars e \<equiv> do {
+    e \<leftarrow> of_object e;
+    source \<leftarrow> get e ''source'' \<bind> of_nat;
+    target \<leftarrow> get e ''target'' \<bind> of_nat;
+    guard  \<leftarrow> get e ''guard''  \<bind> of_string;
+    label  \<leftarrow> get e ''label''  \<bind> of_string;
+    update \<leftarrow> get e ''update'' \<bind> of_string;
+    label  \<leftarrow> if label = STR '''' then STR '''' |> Sil |> Result else
+      parse scan_action label |> err_msg (STR ''Failed to parse label in '' + label);
+    (g, check)  \<leftarrow> compile_invariant' clocks vars guard |> err_msg (STR ''Failed to parse guard!'');
+    upd \<leftarrow> if update = STR '''' then Result [] else
+      parse scan_updates update |> err_msg (STR ''Failed to parse update in '' + update);
+    let resets = filter (\<lambda>x. fst x \<in> set clocks) upd;
+    assert
+      (list_all (\<lambda>(_, d). case d of exp.const x \<Rightarrow> x = 0 | _ \<Rightarrow> undefined) resets)
+      (STR ''Clock resets to values different from zero are not supported'');
+    let resets = map fst resets;
+    let upds = filter (\<lambda>x. fst x \<notin> set clocks) upd;
+    assert
+      (list_all (\<lambda>(x, _). x \<in> set vars) upds)
+      (STR ''Unknown variable in update: '' + update);
+    Result (source, check, g, label, upds, resets, target)
+  }"
+
+definition convert_automaton where
+  "convert_automaton clocks vars a \<equiv> do {
+    nodes \<leftarrow> get a ''nodes'' \<bind> of_array;
+    edges \<leftarrow> get a ''edges'' \<bind> of_array;
+    nodes \<leftarrow> combine_map (convert_node clocks vars) nodes;
+    let invs = map (\<lambda> ((_, n), g). (n, g)) nodes;
+    let names_to_ids = map fst nodes;
+    assert (map fst names_to_ids |> filter (\<lambda>s. s \<noteq> STR '''') |> distinct)
+      (STR ''Node names are ambiguous'' + (show (map fst names_to_ids) |> String.implode));
+    assert (map snd names_to_ids |> distinct) (STR ''Duplicate node id'');
+    let ids_to_names = map_of (map prod.swap names_to_ids);
+    let names_to_ids = map_of names_to_ids;
+    let committed = default [] (get a ''committed'' \<bind> of_array);
+    committed \<leftarrow> combine_map of_nat committed;
+    let urgent = default [] (get a ''urgent'' \<bind> of_array);
+    urgent \<leftarrow> combine_map of_nat urgent;
+    edges \<leftarrow> combine_map (convert_edge clocks vars) edges;
+    Result (names_to_ids, ids_to_names, (committed, urgent, edges, invs))
+  }"
+
+fun rename_locs_sexp where
+  "rename_locs_sexp f (not a) =
+    do {a \<leftarrow> rename_locs_sexp f a; not a |> Result}" |
+  "rename_locs_sexp f (imply a b) =
+    do {a \<leftarrow> rename_locs_sexp f a; b \<leftarrow> rename_locs_sexp f b; imply a b |> Result}" |
+  "rename_locs_sexp f (sexp.or a b) =
+    do {a \<leftarrow> rename_locs_sexp f a; b \<leftarrow> rename_locs_sexp f b; sexp.or a b |> Result}" |
+  "rename_locs_sexp f (and a b) =
+    do {a \<leftarrow> rename_locs_sexp f a; b \<leftarrow> rename_locs_sexp f b; and a b |> Result}" |
+  "rename_locs_sexp f (loc n x) = do {x \<leftarrow> f n x; loc n x |> Result}" |
+  "rename_locs_sexp f (eq a b) = Result (eq a b)" |
+  "rename_locs_sexp f (lt a b) = Result (lt a b)" |
+  "rename_locs_sexp f (le a b) = Result (le a b)" |
+  "rename_locs_sexp f (ge a b) = Result (ge a b)" |
+  "rename_locs_sexp f (gt a b) = Result (gt a b)"
+
+fun rename_locs_formula where
+  "rename_locs_formula f (formula.EX \<phi>) = rename_locs_sexp f \<phi> \<bind> Result o formula.EX" |
+  "rename_locs_formula f (EG \<phi>) = rename_locs_sexp f \<phi> \<bind> Result o EG" |
+  "rename_locs_formula f (AX \<phi>) = rename_locs_sexp f \<phi> \<bind> Result o AX" |
+  "rename_locs_formula f (AG \<phi>) = rename_locs_sexp f \<phi> \<bind> Result o AG" |
+  "rename_locs_formula f (Leadsto \<phi> \<psi>) =
+    do {\<phi> \<leftarrow> rename_locs_sexp f \<phi>; \<psi> \<leftarrow> rename_locs_sexp f \<psi>; Leadsto \<phi> \<psi> |> Result}"
+
+definition convert :: "JSON \<Rightarrow>
+  ((nat \<Rightarrow> nat \<Rightarrow> String.literal) \<times> (String.literal \<Rightarrow> nat) \<times> String.literal list \<times>
+    (nat list \<times> nat list \<times>
+     (String.literal act, nat, String.literal, int, String.literal, int) transition list
+      \<times> (nat \<times> (String.literal, int) cconstraint) list) list \<times>
+   (String.literal \<times> int \<times> int) list \<times>
+   (nat, nat, String.literal, int) formula \<times> nat list \<times> (String.literal \<times> int) list
+  ) Error_List_Monad.result" where
+  "convert json \<equiv> do {
+    all \<leftarrow> of_object json;
+    automata \<leftarrow> get all ''automata'';
+    automata \<leftarrow> of_array automata;
+    let broadcast = default [] (do {x \<leftarrow> get all ''broadcast''; of_array x});
+    broadcast \<leftarrow> combine_map of_string broadcast;
+    let _ = trace_level 3
+      (\<lambda>_. return (STR ''Broadcast channels '' + String.implode (show broadcast)));
+    let bounds = default (STR '''') (do {
+      x \<leftarrow> get all ''vars''; of_string x}
+    );
+    bounds \<leftarrow> parse parse_bounds bounds |> err_msg (STR ''Failed to parse bounds'');
+    clocks \<leftarrow> get all ''clocks'';
+    clocks \<leftarrow> of_string clocks;
+    clocks \<leftarrow> parse (parse_list (lx_ws *-- ta_var_ident with String.implode)) clocks
+       |> err_msg (STR ''Failed to parse clocks'');
+    formula \<leftarrow> get all ''formula'';
+    formula \<leftarrow> of_string formula;
+    formula \<leftarrow> parse scan_formula formula |> err_msg (STR ''Failed to parse formula'');
+    automata \<leftarrow> combine_map of_object automata;
+    process_names \<leftarrow> combine_map (\<lambda>a. get a ''name'' \<bind> of_string) automata;
+    assert (distinct process_names) (STR ''Process names are ambiguous'');
+    assert (locs_of_formula formula \<subseteq> set process_names) (STR ''Unknown process name in formula'');
+    let process_names_to_index = List_Index.index process_names;
+    init_locs \<leftarrow> combine_map
+      (\<lambda>a. do {x \<leftarrow> get a ''initial''; x \<leftarrow> of_nat x; x |> Result})
+      automata;
+    let formula = formula.map_formula process_names_to_index id id id formula;
+    let vars = map fst bounds;
+    let init_vars = map (\<lambda>x. (x, 0::int)) vars;
+    names_automata \<leftarrow> combine_map (convert_automaton clocks vars) automata;
+    let automata = map (snd o snd) names_automata;
+    let names    = map fst names_automata;
+    let ids_to_names = map (fst o snd) names_automata;
+    let ids_to_names =
+      (\<lambda>p i. case (ids_to_names ! p) i of Some n \<Rightarrow> n | None \<Rightarrow> String.implode (show i));
+    formula \<leftarrow> rename_locs_formula (\<lambda>i. get (names ! i)) formula;
+    Result
+    (ids_to_names, process_names_to_index,
+     broadcast, automata, bounds, formula, init_locs, init_vars)
+}" for json
+
+
+
+paragraph \<open>Unsafe Glue Code for Printing\<close>
+
+code_printing
+  constant print \<rightharpoonup> (SML) "writeln _"
+       and        (OCaml) "print'_string _"
+code_printing
+  constant println \<rightharpoonup> (SML) "writeln _"
+       and          (OCaml) "print'_string _"
+
+definition parse_convert_run_print where
+  "parse_convert_run_print dc s \<equiv>
+   case parse json s \<bind> convert of
+     Error es \<Rightarrow> do {let _ = map println es; return ()}
+   | Result (ids_to_names, _, broadcast, automata, bounds, formula, L0, s0) \<Rightarrow> do {
+      r \<leftarrow> do_preproc_mc dc ids_to_names (broadcast, automata, bounds) L0 s0 formula;
+      case r of
+        Error es \<Rightarrow> do {let _ = map println es; return ()}
+      | Result s \<Rightarrow> do {let _ = println s; return ()}
+  }"
+
+definition parse_convert_run where
+  "parse_convert_run dc s \<equiv>
+   case
+      parse json s \<bind> (\<lambda>r.
+      let
+      s' = show r |> String.implode;
+      _  = trace_level 2 (\<lambda>_. return s')
+      in parse json s' \<bind> (\<lambda>r'.
+      assert (r = r') STR ''Parse-print-parse loop failed!'' \<bind> (\<lambda>_. convert r)))
+   of
+     Error es \<Rightarrow> return (Error es)
+   | Result (ids_to_names, _, broadcast, automata, bounds, formula, L0, s0) \<Rightarrow>
+      do_preproc_mc dc ids_to_names (broadcast, automata, bounds) L0 s0 formula
+"
+
+definition convert_run where
+  "convert_run dc json_data \<equiv>
+   case (
+      let
+      s' = show json_data |> String.implode;
+      _  = trace_level 2 (\<lambda>_. return s')
+      in parse json s' \<bind> (\<lambda>r'.
+      assert (json_data = r') STR ''Parse-print-parse loop failed!'' \<bind> (\<lambda>_. convert json_data)))
+   of
+     Error es \<Rightarrow> return (Error es)
+   | Result (ids_to_names, _, broadcast, automata, bounds, formula, L0, s0) \<Rightarrow>
+      do_preproc_mc dc ids_to_names (broadcast, automata, bounds) L0 s0 formula
+"
+
+text \<open>Eliminate Gabow statistics\<close>
+code_printing
+  code_module Gabow_Skeleton_Statistics \<rightharpoonup> (SML) \<open>\<close>
+code_printing
+  code_module AStatistics \<rightharpoonup> (SML) \<open>\<close>
+
+text \<open>Delete ``junk''\<close>
+code_printing code_module Bits_Integer \<rightharpoonup> (SML) \<open>\<close>
+
+code_printing
+  constant stat_newnode \<rightharpoonup> (SML) "(fn x => ()) _"
+| constant stat_start   \<rightharpoonup> (SML) "(fn x => ()) _"
+| constant stat_stop    \<rightharpoonup> (SML) "(fn x => ()) _"
+
+
+(* Fix to IArray theory *)
+code_printing
+  constant IArray.sub' \<rightharpoonup> (SML) "(Vector.sub o (fn (a, b) => (a, IntInf.toInt b)))"
+
+code_printing code_module "Logging" \<rightharpoonup> (SML)
+\<open>
+structure Logging : sig
+  val set_level : int -> unit
+  val trace : int -> (unit -> string) -> unit
+  val get_trace: unit -> (int * string) list
+end = struct
+  val level = Unsynchronized.ref 0;
+  val messages : (int * string) list ref = Unsynchronized.ref [];
+  fun set_level i = level := i;
+  fun get_trace () = !messages;
+  fun trace i f =
+    if i > !level
+    then ()
+    else
+      let
+        val s = f ();
+        val _ = messages := (i, s) :: !messages;
+      in () end;
+end
+\<close>
+and (Eval) \<open>\<close>
+code_reserved (Eval) Logging
+code_reserved (SML) Logging
+
+
+code_printing constant trace_level \<rightharpoonup> (SML)
+  "Logging.trace (IntInf.toInt (integer'_of'_int _)) (_ ())"
+and (Eval)
+  "(fn '_ => fn s => writeln (s ())) _ (_ ())"
+
+text \<open>To disable state tracing:\<close>
+cancel\<open>code_printing
+  constant "Show_State_Defs.tracei" \<rightharpoonup>
+      (SML)   "(fn n => fn show_state => fn show_clock => fn typ => fn x => ()) _ _ _"
+  and (OCaml) "(fun n show_state show_clock ty x -> -> ()) _ _ _"\<close>
+
+paragraph \<open>SML Code Export\<close>
+text_raw \<open>\label{export-munta}\<close>
+
+text \<open>Now we export the actual executable code for \munta. First for SML:\<close>
+
+export_code parse_convert_run Result Error
+  in SML module_name Model_Checker
+  file_prefix "Simple_Model_Checker"
+
+paragraph \<open>OCaml Code Export\<close>
+text_raw \<open>\label{export-ocaml}\<close>
+
+text \<open>This is how to do it for OCaml:\<close>
+export_code
+  convert_run Result Error String.explode int_of_integer nat_of_integer
+  JSON.Object JSON.Array JSON.String JSON.Int JSON.Nat JSON.Rat JSON.Boolean JSON.Null fract.Rat
+  in OCaml module_name Model_Checker
+  file_prefix "Simple_Model_Checker"
+
+paragraph \<open>Running Munta\<close>
+text_raw \<open>\label{run-munta}\<close>
+
+text \<open>
+Adding a bit of wrapper code, so that we can directly run \textsc{Munta} from within Isabelle:
+\<close>
+definition parse_convert_run_test where
+  "parse_convert_run_test dc s \<equiv> do {
+    x \<leftarrow> parse_convert_run dc s;
+    case x of
+      Error es \<Rightarrow> do {let _ = map println es; return (STR ''Fail'')}
+    | Result r \<Rightarrow> return r
+  }"
+
+ML \<open>
+  fun assert comp exp =
+    if comp = exp then () else error ("Assertion failed! expected: " ^ exp ^ " but got: " ^ comp)
+  fun test dc file =
+  let
+    val dir = Path.append @{master_dir} @{path "../benchmarks/"}
+    val file = Path.explode file |> Path.append dir
+    val s = File.read file
+  in
+    @{code parse_convert_run_test} dc s end
+\<close>
+
+text \<open>Now we can run \munta\ on a few examples:\<close>
+
+ML_val \<open>assert (test false "HDDI_02.muntax" ()) "Property is not satisfied!"\<close>
+ML_val \<open>assert (test true "HDDI_02.muntax" ()) "Model has no deadlock!"\<close>
+
+ML_val \<open>assert (test false "simple.muntax" ()) "Property is satisfied!"\<close>
+ML_val \<open>assert (test true "simple.muntax" ()) "Model has no deadlock!"\<close>
+
+ML_val \<open>assert (test false "light_switch.muntax" ()) "Property is satisfied!"\<close>
+ML_val \<open>assert (test true "light_switch.muntax" ()) "Model has no deadlock!"\<close>
+
+ML_val \<open>assert (test false "PM_test.muntax" ()) "Property is not satisfied!"\<close>
+ML_val \<open>assert (test true "PM_test.muntax" ()) "Model has a deadlock!"\<close>
+
+ML_val \<open>assert (test false "bridge.muntax" ()) "Property is satisfied!"\<close>
+ML_val \<open>assert (test true "bridge.muntax" ()) "Model has no deadlock!"\<close>
+
+ML_val \<open>assert (test false "fischer.muntax" ()) "Property is satisfied!"\<close>
+ML_val \<open>assert (test true "fischer.muntax" ()) "Model has no deadlock!"\<close>
+
+ML_val \<open>assert (test false "PM_all_4.muntax" ()) "Property is not satisfied!"\<close>
+ML_val \<open>assert (test true "PM_all_4.muntax" ()) "Model has no deadlock!"\<close>
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Munta_MLton/Munta_Compile_MLton.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Munta_MLton/Munta_Compile_MLton.thy
@@ -0,0 +1,115 @@
+section \<open>Build and Test Exported Program With MLton\<close>
+text_raw \<open>\label{build-mlton}\<close>
+
+theory Munta_Compile_MLton
+  imports Munta_Model_Checker.Simple_Network_Language_Export_Code
+begin
+
+\<comment> \<open>Produces a command for checking a single benchmark, e.g.:
+verbatim\<open>./check_benchmark.sh 568 "Property is not satisfied" ./munta -m benchmarks/PM_all_5.muntax\<close>
+\<close>
+ML \<open>
+fun mk_benchmark_check_gen munta mk_prop name num_states satisfied =
+  let
+    val prop = mk_prop satisfied
+    val benchmark = name ^ ".muntax"
+  in
+    \<comment> \<open>this line checks that number of states and property satisfaction are correct\<close>
+    "./check_benchmark.sh " ^ string_of_int num_states ^ " " ^ prop
+    \<comment> \<open>this line runs Munta on the actual benchmark\<close>
+    ^ " " ^ munta ^ " " ^ benchmark
+  end
+val mk_benchmark_check =
+  mk_benchmark_check_gen "./munta -m"
+    (fn satisfied =>
+      if satisfied then
+        verbatim\<open>"Property is satisfied!"\<close>
+      else
+        verbatim\<open>"Property is not satisfied!"\<close>)
+val mk_benchmark_check_dc =
+  mk_benchmark_check_gen "./munta -dc -m"
+    (fn satisfied =>
+      if satisfied then
+        verbatim\<open>"Model has a deadlock!"\<close>
+      else
+        verbatim\<open>"Model has no deadlock!"\<close>)
+val it  = mk_benchmark_check "PM_all_5" 568 false
+val it1 = mk_benchmark_check_dc "hddi_08" 466 false
+\<close>
+
+text \<open>Here is how to compile Munta with MLton and then run some benchmarks:\<close>
+
+\<comment> \<open>Commenting this out since the AFP submission machine sandboxing somehow breaks MLton.\<close>
+compile_generated_files "code/Simple_Model_Checker.ML" (in Simple_Network_Language_Export_Code)
+  external_files
+    \<open>Unsynchronized.sml\<close>
+    \<open>Writeln.sml\<close>
+    \<open>Util.sml\<close>
+    \<open>Muntax.sml\<close>
+    \<open>Mlton_Main.sml\<close>
+    \<open>library.ML\<close>
+    \<open>muntax.mlb\<close>
+    \<open>check_benchmark.sh\<close>
+    (in "../ML")
+  and
+    \<open>HDDI_02.muntax\<close>
+    \<open>HDDI_02_broadcast.muntax\<close>
+    \<open>HDDI_08_broadcast.muntax\<close>
+    \<open>PM_all_1.muntax\<close>
+    \<open>PM_all_2.muntax\<close>
+    \<open>PM_all_3.muntax\<close>
+    \<open>PM_all_4.muntax\<close>
+    \<open>PM_all_5.muntax\<close>
+    \<open>PM_all_6.muntax\<close>
+    \<open>PM_all_urgent.muntax\<close>
+    \<open>bridge.muntax\<close>
+    \<open>csma_05.muntax\<close>
+    \<open>csma_06.muntax\<close>
+    \<open>fischer.muntax\<close>
+    \<open>fischer_05.muntax\<close>
+    \<open>hddi_08.muntax\<close>
+    \<open>light_switch.muntax\<close> (in "../benchmarks")
+  export_files \<open>munta\<close> (exe)
+  where \<open>fn dir =>
+    let
+      val exec = Generated_Files.execute dir
+      val _ =
+        exec \<open>Preparation\<close>
+          "mv code/Simple_Model_Checker.ML Simple_Model_Checker.ML"
+      val _ =
+        exec \<open>Compilation\<close>
+          (cancel\<open>verbatim\<open>"$ISABELLE_MLTON" $ISABELLE_MLTON_OPTIONS\<close> ^\<close>
+           verbatim\<open>"$ISABELLE_MLTON" \<close> ^
+            \<comment> \<open>these additional settings have been copied from the AFP entry \<open>PAC_Checker\<close>\<close>
+            \<comment> \<open>they bring down \<open>benchmarks/PM_all_6.muntax\<close> from 1000 s to 180 s on an M1 Mac\<close>
+             "-const 'MLton.safe false' -verbose 1 -inline 700 -cc-opt -O3 " ^
+            \<comment> \<open>this one from \<open>PAC_Checker\<close> does not work on ARM64 though\<close>
+            cancel\<open>"-codegen native " ^\<close>
+            \<comment> \<open>these used to be the only setting for Munta\<close>
+            "-default-type int64 " ^
+            "-output munta " ^
+            "muntax.mlb")
+      val _ = exec \<open>Test HDDI_02\<close> (mk_benchmark_check "HDDI_02" 34 false)
+      val _ = exec \<open>Test HDDI_02_broadcast\<close> (mk_benchmark_check "HDDI_02_broadcast" 34 false)
+      val _ = exec \<open>Test HDDI_08_broadcast\<close> (mk_benchmark_check "HDDI_08_broadcast" 466 false)
+      val _ = exec \<open>Test PM_all_1\<close> (mk_benchmark_check "PM_all_4" 529 false)
+      val _ = exec \<open>Test PM_all_1\<close> (mk_benchmark_check "PM_all_1" 338 false)
+      val _ = exec \<open>Test PM_all_2\<close> (mk_benchmark_check "PM_all_2" 2828 false)
+      val _ = exec \<open>Test PM_all_3\<close> (mk_benchmark_check "PM_all_3" 3994 false)
+      val _ = exec \<open>Test PM_all_4\<close> (mk_benchmark_check "PM_all_4" 529 false)
+      val _ = exec \<open>Test PM_all_5\<close> (mk_benchmark_check "PM_all_5" 568 false)
+      \<comment> \<open>180 s on an M1 Mac\<close>
+      cancel\<open>val _ = exec \<open>Test PM_all_6\<close> (mk_benchmark_check "PM_all_6" 416245 false)\<close>
+      val _ = exec \<open>Test PM_all_urgent\<close> (mk_benchmark_check "PM_all_urgent" 8 true)
+      val _ = exec \<open>Test bridge\<close> (mk_benchmark_check "bridge" 73 true)
+      val _ = exec \<open>Test csma_05\<close> (mk_benchmark_check "csma_05" 8217 false)
+      val _ = exec \<open>Test csma_06\<close> (mk_benchmark_check "csma_06" 68417 false)
+      val _ = exec \<open>Test fischer\<close> (mk_benchmark_check "fischer" 4500 true)
+      val _ = exec \<open>Test fischer_05\<close> (mk_benchmark_check "fischer_05" 38579 false)
+      val _ = exec \<open>Test hddi_08\<close> (mk_benchmark_check "hddi_08" 466 false)
+      val _ = exec \<open>Test light_switch\<close> (mk_benchmark_check "light_switch" 2 true)
+      \<comment> \<open>a test for the deadlock checker\<close>
+      val _ = exec \<open>Test hddi_08\<close> (mk_benchmark_check_dc "hddi_08" 466 false)
+    in () end\<close>
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Networks/Networks.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Networks/Networks.thy
@@ -0,0 +1,699 @@
+theory Networks
+  imports Timed_Automata.Timed_Automata Munta_Base.Normalized_Zone_Semantics_Impl
+    Munta_Base.Reordering_Quantifiers Munta_Base.TA_Syntax_Bundles
+begin
+
+unbundle no_library_syntax
+
+section \<open>Networks of Timed Automata\<close>
+
+subsection \<open>Syntax and Operational Semantics\<close>
+
+text \<open>Input, output and internal transitions\<close>
+
+datatype 'a act = In 'a | Out 'a | Sil 'a
+
+datatype 'b label = Del | Act 'b | Syn 'b 'b
+
+type_synonym
+  ('a, 'b, 'c, 't, 's) nta = "('a act * 'b, 'c, 't, 's) ta list"
+
+inductive step_n ::
+  "('a, 'b, 'c, 't, 's) nta \<Rightarrow> 's list \<Rightarrow> ('c, ('t::time)) cval \<Rightarrow> 'b label
+  \<Rightarrow> 's list \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>N \<langle>_, _\<rangle> \<rightarrow>_ \<langle>_,_\<rangle>" [61,61,61,61,61,61] 61)
+where
+  step_n_t:
+    "\<lbrakk>\<forall> p \<in> {..<length N}. N!p \<turnstile> \<langle>L!p, u\<rangle> \<rightarrow>d \<langle>L!p, u \<oplus> d\<rangle>; d \<ge> 0\<rbrakk>
+    \<Longrightarrow> N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Del \<langle>L, u \<oplus> d\<rangle>" |
+  step_n_i:
+    "\<lbrakk>
+      N!p \<turnstile> l \<longrightarrow>g,(Sil a, b),r l'; u \<turnstile> g; \<forall> p \<in> {..<length N}. u' \<turnstile> inv_of (N!p) (L'!p);
+      L!p = l; p < length L; L' = L[p := l']; u' = [r\<rightarrow>0]u
+     \<rbrakk>
+    \<Longrightarrow> N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Act b \<langle>L', u'\<rangle>" |
+  step_n_s:
+    "\<lbrakk>N!p \<turnstile> l1 \<longrightarrow>g1,(In a, b1),r1 l1'; N!q \<turnstile> l2 \<longrightarrow>g2,(Out a, b2),r2 l2'; u \<turnstile> g1; u \<turnstile> g2;
+      \<forall> p \<in> {..<length N}. u' \<turnstile> inv_of (N!p) (L'!p);
+      L!p = l1; L!q = l2; p < length L; q < length L; p \<noteq> q;
+      L' = L[p := l1', q := l2']; u' = [(r1 @ r2)\<rightarrow>0]u
+     \<rbrakk> \<Longrightarrow> N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Syn b1 b2 \<langle>L', u'\<rangle>"
+
+inductive_cases[elim!]: "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>a \<langle>L', u'\<rangle>"
+
+inductive steps_n ::
+  "('a, 'b, 'c, 't, 's) nta \<Rightarrow> 's list \<Rightarrow> ('c, ('t::time)) cval
+  \<Rightarrow> 's list \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>N \<langle>_, _\<rangle> \<rightarrow>* \<langle>_,_\<rangle>" [61,61,61] 61)
+where
+  refl: "N \<turnstile>N \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l, Z\<rangle>" |
+  step: "N \<turnstile>N \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l', Z'\<rangle> \<Longrightarrow> N \<turnstile>N \<langle>l', Z'\<rangle> \<rightarrow>_ \<langle>l'', Z''\<rangle>
+        \<Longrightarrow> N \<turnstile>N \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l'', Z''\<rangle>"
+
+declare steps_n.intros[intro]
+
+lemma stepI2:
+  "N \<turnstile>N \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l'', Z''\<rangle>" if "N \<turnstile>N \<langle>l', Z'\<rangle> \<rightarrow>* \<langle>l'', Z''\<rangle>" "N \<turnstile>N \<langle>l, Z\<rangle> \<rightarrow>b \<langle>l', Z'\<rangle>"
+  using that
+  apply induction
+   apply (rule steps_n.step)
+    apply (rule refl)
+   apply assumption
+  apply simp
+  by (rule steps_n.step; assumption)
+
+lemma step_n_t_I:
+  "\<lbrakk>\<forall> p \<in> {..<length N}. u \<turnstile> inv_of (N!p) (L!p); \<forall> p \<in> {..<length N}. u \<oplus> d \<turnstile> inv_of (N!p) (L!p);
+    d \<ge> 0
+   \<rbrakk> \<Longrightarrow> N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Del \<langle>L, u \<oplus> d\<rangle>"
+by (auto intro: step_n_t)
+
+subsection \<open>Product Automaton\<close>
+
+abbreviation state_set :: "('a, 'c, 'time, 's) transition set \<Rightarrow> 's set" where
+  "state_set T \<equiv> fst ` T \<union> (snd o snd o snd o snd) ` T"
+
+lemma guard_concat:
+  assumes "\<forall> g \<in> set xs. u \<turnstile> g"
+  shows "u \<turnstile> concat xs"
+  using assms by (induction xs) (auto simp: clock_val_def)
+
+lemma guard_append:
+  assumes "u \<turnstile> g1" "u \<turnstile> g2"
+  shows "u \<turnstile> g1 @ g2"
+  using assms by (auto simp: clock_val_def)
+
+lemma concat_guard:
+  assumes "u \<turnstile> concat xs" "g \<in> set xs"
+  shows "u \<turnstile> g"
+using assms by (auto simp: list_all_iff clock_val_def)
+
+lemma guard_consE:
+  assumes "u \<turnstile> ac # cc"
+  obtains "u \<turnstile>a ac" "u \<turnstile> cc"
+  using assms unfolding clock_val_def by auto
+
+lemma guard_consI:
+  assumes "u \<turnstile>a ac" "u \<turnstile> cc"
+  shows "u \<turnstile> ac # cc"
+  using assms unfolding clock_val_def by auto
+
+lemma collect_clock_pairs_append_cases:
+  assumes "x \<in> collect_clock_pairs (g1 @ g2)"
+  shows "x \<in> collect_clock_pairs g1 \<or> x \<in> collect_clock_pairs g2"
+    using assms unfolding collect_clock_pairs_def by auto
+
+lemma list_update_nth_split:
+  assumes "j < length xs"
+  shows "P (xs[i := x] ! j) = ((i = j \<longrightarrow> P x) \<and> (i \<noteq> j \<longrightarrow> P (xs ! j)))"
+    using assms by (cases "i = j") auto
+
+locale Product_TA_Defs =
+  fixes N :: "('a, 'b, 'c, 't, 's) nta"
+begin
+
+  abbreviation "T \<equiv> map trans_of N"
+  abbreviation "I \<equiv> map inv_of N"
+
+  definition "states = {L. length L = length T \<and> (\<forall> p \<in> {..<length T}. L!p \<in> state_set (T!p))}"
+
+  definition
+    "product_trans_i =
+      {(L,g,(a,Act b),r,L[p:=l']) | L p g a b r l'.
+      L \<in> states \<and> (L!p, g, (Sil a, b), r, l') \<in> T!p \<and> p < length L}"
+
+  definition
+    "product_trans_s =
+      {(L,g1 @ g2,(a,Syn b1 b2),r1 @ r2,L[p:=l1',q:=l2']) | L p q g1 g2 a b1 b2 r1 r2 l1' l2'.
+        L \<in> states \<and> (L!p, g1, (In a, b1), r1, l1') \<in> T!p \<and> (L!q, g2, (Out a, b2), r2, l2') \<in> T!q
+        \<and> p < length L \<and> q < length L \<and> p \<noteq> q
+      }"
+
+  definition
+    "product_trans \<equiv> product_trans_i \<union> product_trans_s"
+
+lemma product_state_set_subs:
+  assumes "q < length N" "l \<in> state_set product_trans"
+  shows "l ! q \<in> state_set (T ! q)"
+  using assms
+  unfolding product_trans_def product_trans_i_def product_trans_s_def states_def
+  apply safe
+     apply (solves auto)
+    apply (solves auto)
+  apply simp
+   apply (subst list_update_nth_split; force)
+  apply simp
+  apply (subst list_update_nth_split)
+   apply simp
+  apply safe
+   apply (subst (asm) (2) list_update_nth_split; force)
+  apply (subst list_update_nth_split)
+   apply simp
+  by (subst (asm) (2) list_update_nth_split; force) \<comment> \<open>slow: 13s\<close>
+
+  definition
+    "product_invariant L \<equiv>
+    concat (map (\<lambda> p. if p < length I then (I!p) (L!p) else []) [0..<length L])"
+
+  definition product_ta :: "('a \<times> 'b label, 'c, 't, 's list) ta" where
+    "product_ta \<equiv> (product_trans, product_invariant)"
+
+  lemma collect_clki_product_invariant:
+    "Timed_Automata.collect_clki product_invariant = \<Union> (Timed_Automata.collect_clki ` set I)"
+    unfolding Timed_Automata.collect_clki_def product_invariant_def
+    apply (simp add: image_Union)
+    apply safe
+    subgoal
+      unfolding collect_clock_pairs_def by fastforce
+    apply clarsimp
+    subgoal premises prems for a b aa ba x
+    proof -
+      let ?L = "map (\<lambda> _. x) [0..<length N]"
+      let ?x = "collect_clock_pairs
+        (concat
+        (map (\<lambda>p. if p < length N then (I ! p) (?L ! p) else [])
+        [0..<length ?L]))"
+      show ?thesis
+        apply (intro exI[where x = ?x] conjI)
+         apply (rule exI; rule HOL.refl)
+        apply simp
+        using prems unfolding collect_clock_pairs_def by (fastforce dest: mem_nth)
+    qed
+    done
+
+  lemma states_length:
+    assumes "L \<in> states"
+    shows "length L = length N"
+    using assms unfolding states_def by auto
+
+  lemma collect_clkt_product_trans_subs:
+    "Timed_Automata.collect_clkt product_trans \<subseteq> \<Union> (Timed_Automata.collect_clkt ` set T)"
+    unfolding
+      Timed_Automata.collect_clkt_def product_trans_def product_trans_i_def product_trans_s_def
+    by (fastforce dest: collect_clock_pairs_append_cases states_length)
+
+  lemma collect_clkvt_product_trans_subs:
+    "collect_clkvt product_trans \<subseteq> \<Union> (collect_clkvt ` set T)"
+    unfolding collect_clkvt_def product_trans_def product_trans_i_def product_trans_s_def
+    by (fastforce dest: states_length)
+
+  lemma statesI_aux:
+    fixes L
+    assumes L: "L \<in> states"
+    assumes
+      "p < length T"
+      "(l, g, a, r, l') \<in> T ! p"
+    shows "(L[p := l]) \<in> states"
+    unfolding states_def apply clarsimp
+    using L apply (simp add: states_def)
+    apply safe
+    apply (subst list_update_nth_split)
+    using assms by (fastforce simp: states_def)+
+
+  lemma product_trans_s_alt_def:
+    "product_trans_s =
+      {(L, g, (a, Syn b1 b2), r, L') | L g a b1 b2 r L'. \<exists> p q g1 g2 r1 r2 l1' l2'.
+      L \<in> states \<and> p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+      g = g1 @ g2 \<and> r = r1 @ r2 \<and> L' = L[p:=l1', q:=l2']
+        \<and> (L!p, g1, (In a, b1), r1, l1') \<in> T!p \<and> (L!q, g2, (Out a, b2), r2, l2') \<in> T!q
+      }"
+    unfolding product_trans_s_def by (safe; metis)
+
+  context
+    assumes states_not_empty: "states \<noteq> {}"
+    assumes trans_complete:
+      "\<forall> p < length T. \<forall> t1 \<in> T ! p. case t1 of (l1, g1, (In a, b1), r1, l1')
+      \<Rightarrow> \<exists> q < length T. p \<noteq> q \<and> (\<exists> l2 g2 b2 r2 l2'. (l2, g2, (Out a, b2), r2, l2') \<in> T ! q) |
+      (l1, g1, (Out a, b1), r1, l1')
+      \<Rightarrow> \<exists> q < length T. p \<noteq> q \<and> (\<exists> l2 g2 b2 r2 l2'. (l2, g2, (In a, b2), r2, l2') \<in> T ! q) | _ \<Rightarrow> True"
+  begin
+
+  lemma collect_clkt_product_trans_sups:
+    "Timed_Automata.collect_clkt product_trans \<supseteq> \<Union> (Timed_Automata.collect_clkt ` set T)"
+  proof
+    fix x assume "x \<in> \<Union> (Timed_Automata.collect_clkt ` set T)"
+    then obtain trans l1 g1 a' b1 r1 l1' where x:
+      "trans \<in> set T" "(l1, g1, (a', b1), r1, l1') \<in> trans" "x \<in> collect_clock_pairs g1"
+      unfolding Timed_Automata.collect_clkt_def by force
+    then obtain p where p:
+        "p < length T" "T ! p = trans"
+      by (auto dest!: mem_nth)
+    from states_not_empty obtain L where L: "L \<in> states" by auto
+    have "length T = length L" by (auto simp: states_length[OF \<open>L \<in> states\<close>])
+    show "x \<in> Timed_Automata.collect_clkt product_trans"
+    proof (cases a')
+      case a': (In a)
+      with x p trans_complete obtain q l2 g2 b2 r2 l2' where trans2:
+        "q < length T" "(l2, g2, (Out a, b2), r2, l2') \<in> T ! q" "p \<noteq> q"
+        by atomize_elim fastforce
+      have "L[p := l1, q := l2] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) trans2 unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1" "?L ! q = l2"
+        using \<open>p \<noteq> q\<close> \<open>p < length T\<close> \<open>q < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      moreover note t = calculation
+      have "(?L, g1 @ g2, (a, Syn b1 b2), r1 @ r2, ?L[p := l1', q := l2']) \<in> product_trans_s"
+        unfolding product_trans_s_alt_def
+        apply clarsimp
+        using t p(1) x(1,2) trans2 unfolding p(2)[symmetric] a' \<open>length T = length L\<close>
+        by metis
+      moreover have "x \<in> collect_clock_pairs (g1 @ g2)"
+        using x(3) by (auto simp: collect_clock_pairs_def)
+      ultimately show ?thesis unfolding Timed_Automata.collect_clkt_def product_trans_def by force
+    next
+      case a': (Out a)
+      with x p trans_complete obtain q l2 g2 b2 r2 l2' where trans2:
+        "q < length T" "(l2, g2, (In a, b2), r2, l2') \<in> T ! q" "p \<noteq> q"
+        by atomize_elim fastforce
+      have "L[q := l2, p := l1] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) trans2 unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1" "?L ! q = l2"
+        using \<open>p \<noteq> q\<close> \<open>p < length T\<close> \<open>q < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      moreover note t = calculation
+      have "(?L, g2 @ g1, (a, Syn b2 b1), r2 @ r1, ?L[q := l2', p := l1']) \<in> product_trans_s"
+        unfolding product_trans_s_alt_def
+        apply clarsimp
+        using t p(1) x(1,2) trans2 unfolding p(2)[symmetric] a' \<open>length T = length L\<close>
+        by metis
+      moreover have "x \<in> collect_clock_pairs (g2 @ g1)"
+        using x(3) by (auto simp: collect_clock_pairs_def)
+      ultimately show ?thesis unfolding Timed_Automata.collect_clkt_def product_trans_def by force
+    next
+      case a': (Sil a)
+      have "L[p := l1] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1"
+        using \<open>p < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      ultimately have "(?L, g1, (a, Act b1), r1, ?L[p := l1']) \<in> product_trans_i"
+        using x p unfolding product_trans_i_def a' by (force simp: states_length[OF \<open>L \<in> states\<close>])
+      with x(3) show ?thesis unfolding Timed_Automata.collect_clkt_def product_trans_def by force
+    qed
+  qed
+
+  lemma collect_clkvt_product_trans_sups:
+    "collect_clkvt product_trans \<supseteq> \<Union> (collect_clkvt ` set T)"
+  proof
+    fix x assume "x \<in> \<Union> (collect_clkvt ` set T)"
+    then obtain trans l1 g1 a' b1 r1 l1' where x:
+      "trans \<in> set T" "(l1, g1, (a', b1), r1, l1') \<in> trans" "x \<in> set r1"
+      unfolding collect_clkvt_def by force
+    then obtain p where p:
+        "p < length T" "T ! p = trans"
+      by (auto dest!: mem_nth)
+    from states_not_empty obtain L where L: "L \<in> states" by auto
+    show "x \<in> collect_clkvt product_trans"
+    proof (cases a')
+      case a': (In a)
+      with x p trans_complete obtain q l2 g2 b2 r2 l2' where trans2:
+        "q < length T" "(l2, g2, (Out a, b2), r2, l2') \<in> T ! q" "p \<noteq> q"
+        by atomize_elim fastforce
+      moreover have "L[p := l1, q := l2] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) trans2 unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1" "?L ! q = l2"
+        using \<open>p \<noteq> q\<close> \<open>p < length T\<close> \<open>q < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      ultimately have
+        "(?L, g1 @ g2, (a, Syn b1 b2), r1 @ r2, ?L[p := l1', q := l2']) \<in> product_trans_s"
+        using p(1) x trans2 unfolding p(2)[symmetric] a' product_trans_s_def
+        by (fastforce simp: states_length[OF \<open>L \<in> states\<close>])
+      with x(3) show ?thesis unfolding collect_clkvt_def product_trans_def by force
+    next
+      case a': (Out a)
+      with x p trans_complete obtain q l2 g2 b2 r2 l2' where trans2:
+        "q < length T" "(l2, g2, (In a, b2), r2, l2') \<in> T ! q" "p \<noteq> q"
+        by atomize_elim fastforce
+      moreover have "L[q := l2, p := l1] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) trans2 unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1" "?L ! q = l2"
+        using \<open>p \<noteq> q\<close> \<open>p < length T\<close> \<open>q < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      ultimately have "(?L, g2 @ g1, (a, Syn b2 b1), r2 @ r1, ?L[q := l2', p := l1']) \<in> product_trans_s"
+        using p(1) x trans2 unfolding p(2)[symmetric] a' product_trans_s_def
+        apply (simp add: states_length[OF \<open>L \<in> states\<close>] del: ex_simps)
+        apply (tactic \<open>rearrange_ex_tac @{context} 1\<close>)
+        apply simp
+        apply (rule exI, rule conjI, assumption)
+        apply (simp only: ex_simps[symmetric])
+        apply (tactic \<open>rearrange_ex_tac @{context} 1\<close>)
+        apply simp
+        by (fastforce simp: states_length[OF \<open>L \<in> states\<close>])
+          (* fastforce can solve this on its own but not very quickly *)
+      with x(3) show ?thesis unfolding collect_clkvt_def product_trans_def by force
+    next
+      case a': (Sil a)
+      have "L[p := l1] \<in> states" (is "?L \<in> _")
+        using L p(1) x(1,2) unfolding p(2)[symmetric] by (auto intro!: statesI_aux)
+      moreover have "?L ! p = l1"
+        using \<open>p < length T\<close> \<open>L \<in> states\<close> by (auto dest: states_length)
+      ultimately have "(?L, g1, (a, Act b1), r1, ?L[p := l1']) \<in> product_trans_i"
+        using x p unfolding product_trans_i_def a' by (force simp: states_length[OF \<open>L \<in> states\<close>])
+      with x(3) show ?thesis unfolding collect_clkvt_def product_trans_def by force
+    qed
+  qed
+
+  lemma collect_clkt_product_trans:
+    "Timed_Automata.collect_clkt product_trans = \<Union> (Timed_Automata.collect_clkt ` set T)"
+    using collect_clkt_product_trans_sups collect_clkt_product_trans_subs by simp
+
+  lemma collect_clkvt_product_trans:
+    "collect_clkvt product_trans = \<Union> (collect_clkvt ` set T)"
+  using collect_clkvt_product_trans_sups collect_clkvt_product_trans_subs by simp
+
+  end (* End of context for syntactic 1-to-1 correspondence for transitions *)
+
+  context
+    assumes finite_trans: "\<forall> A \<in> set N. finite (trans_of A)"
+  begin
+
+  lemma finite_states:
+    "finite states"
+  proof -
+    have "states \<subseteq> {L. set L \<subseteq>
+          (\<Union> {fst ` trans_of (N ! p) | p. p < length N} \<union>
+            \<Union> {(snd \<circ> snd \<circ> snd \<circ> snd) ` trans_of (N ! p) | p. p < length N})
+          \<and> length L = length N}"
+      unfolding states_def
+      apply clarsimp
+      apply (drule mem_nth)
+      apply safe
+      subgoal for _ _ i
+        by (erule ballE[where x = i]) auto
+      done
+    moreover have "finite \<dots>" using finite_trans by - (rule finite_lists_length_eq; auto)
+    ultimately show ?thesis by (rule finite_subset)
+  qed
+
+  lemma finite_product_trans_i:
+    "finite product_trans_i"
+  proof -
+    let ?N = "\<Union> (trans_of ` set N)"
+    let ?S = "{(L, p, g, (a, Act b), r, l')|L p g a b r l'.
+      L \<in> states \<and> (L ! p, g, (Sil a, b), r, l') \<in> map trans_of N ! p \<and> p < length L}"
+    let ?R = "{(L, p, g, (a, Act b), r, l')|L p g a b r l'.
+      L \<in> states \<and> p < length L \<and> (g, (Sil a, b), r, l') \<in> snd ` ?N}"
+    have "?S \<subseteq> ?R" by (fastforce simp: states_length dest: nth_mem)
+    moreover have "finite ?R" using finite_states [[simproc add: finite_Collect]]
+      apply simp
+      apply (rule finite_imageI)
+      apply (rule finite_SigmaI)
+       apply (rule finite_subset[where B = "{(p, L). p < length L \<and> L \<in> states}"])
+        apply force
+      using states_length
+       apply -
+       apply (rule finite_subset[where B = "{(p, L). p < length N \<and> L \<in> states}"]; fastforce)
+      using finite_trans
+      apply (intro finite_vimageI)
+      unfolding inj_on_def by auto
+    ultimately have "finite ?S" by (rule finite_subset)
+    moreover have "product_trans_i = (\<lambda> (L, p, g, a, r, l'). (L, g, a, r, L[p := l'])) ` ?S"
+      unfolding product_trans_i_def by (auto 4 3 simp: image_Collect)
+    ultimately show ?thesis by auto
+  qed
+
+  lemma finite_Collect_bounded_ex_5' [simp]:
+  assumes "finite {(a,b,c,d,e) . P a b c d e}"
+  shows
+    "finite {x. \<exists>a b c d e. P a b c d e \<and> Q x a b c d e}
+    \<longleftrightarrow> (\<forall> a b c d e. P a b c d e \<longrightarrow> finite {x. Q x a b c d e})"
+  using assms finite_Collect_bounded_ex[OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e). Q x a b c d e"]
+  by clarsimp (* force, simp *)
+
+  lemma finite_product_trans_s:
+    "finite product_trans_s"
+  proof -
+    let ?N = "\<Union> (trans_of ` set N)"
+    have "(g1, (In a, b), r1, l1') \<in> snd ` ?N" if
+      "(L ! p, g1, (In a, b), r1, l1') \<in> map trans_of N ! p" "p < length L" "L \<in> states"
+      for L p g1 a b r1 l1'
+      using that by (auto simp: states_length dest: nth_mem)
+    let ?S = "{(L, p, q, g1, g2, (a, Syn b1 b2), r1, r2, l1', l2')|L p q g1 g2 a b1 b2 r1 r2 l1' l2'.
+      L \<in> states \<and> (L ! p, g1, (In a, b1), r1, l1') \<in> map trans_of N ! p \<and>
+        (L ! q, g2, (Out a, b2), r2, l2') \<in> map trans_of N ! q \<and> p < length L \<and> q < length L}"
+    define F1 where "F1 \<equiv> {(g1, a, b, r1, l1') | g1 a b r1 l1'. (g1, (In a, b), r1, l1') \<in> snd ` ?N}"
+    have fin1: "finite F1" unfolding F1_def using finite_trans
+      using [[simproc add: finite_Collect]] by (force simp: inj_on_def intro: finite_vimageI)
+    define F2 where "F2 \<equiv> {(g1, a, b, r1, l1') | g1 a b r1 l1'. (g1, (Out a, b), r1, l1') \<in> snd ` ?N}"
+    have fin2: "finite F2" unfolding F2_def using finite_trans
+      using [[simproc add: finite_Collect]] by (force simp: inj_on_def intro: finite_vimageI)
+    define R where "R \<equiv> {(L, p, q, g1, g2, (a, Syn b1 b2), r1, r2, l1', l2')|L p q g1 g2 a b1 b2 r1 r2 l1' l2'.
+      L \<in> states \<and> p < length L \<and> q < length L \<and> (g1, a, b1, r1, l1') \<in> F1 \<and> (g2, a, b2, r2, l2') \<in> F2}"
+
+    have R_alt_def: "R = {t. \<exists> L p q g1 a b1 r1 l1' g2 b2 r2 l2'.
+      L \<in> states \<and> p < length L \<and> q < length L
+      \<and> (g1, a, b1, r1, l1') \<in> F1 \<and> (g2, a, b2, r2, l2') \<in> F2
+      \<and> t = (L, p, q, g1, g2, (a, Syn b1 b2), r1, r2, l1', l2')
+      }"
+      unfolding R_def by auto
+    have "?S \<subseteq> R" by (fastforce simp: R_alt_def F1_def F2_def states_length dest: nth_mem)
+    moreover have "finite R"
+      unfolding R_alt_def
+      using fin1 fin2 finite_states
+      using [[simproc add: finite_Collect]]
+      by (auto simp: inj_on_def intro: finite_vimageI intro!: finite_imageI)
+    ultimately have "finite ?S" by (rule finite_subset)
+    moreover have
+      "product_trans_s
+      \<subseteq> (\<lambda> (L, p, q, g1, g2, a, r1, r2, l1', l2').
+          (L, g1 @ g2, a, r1 @ r2, L[p := l1', q := l2'])) ` ?S"
+      unfolding product_trans_s_def by (simp add: image_Collect) blast
+    ultimately show ?thesis by (auto intro: finite_subset)
+  qed
+
+  lemma finite_trans_of_product:
+    shows "finite (trans_of product_ta)"
+    using finite_product_trans_s finite_product_trans_i
+    unfolding product_ta_def trans_of_def product_trans_def
+    by auto
+
+  end (* End of context for finiteness of trans *)
+
+  lemma inv_of_product[simp]:
+      "inv_of product_ta = product_invariant"
+    unfolding product_ta_def inv_of_def by simp
+
+  lemma concat_map_if_aux:
+    assumes "(m :: nat) \<ge> n"
+    shows "concat (map (\<lambda> i. if i < n then f i else []) [0..<m])
+         = concat (map (\<lambda> i. if i < n then f i else []) [0..<n])"
+    using assms by (induction rule: dec_induct) auto
+
+  lemma finite_invariant_of_product[folded inv_of_product]:
+    assumes "\<forall> A \<in> set N. finite (range (inv_of A))"
+    shows "finite (range product_invariant)"
+  proof -
+    let ?N = "\<Union> (range ` inv_of ` set N)"
+    let ?X = "{I. set I \<subseteq> ?N \<and> length I \<le> length N}"
+    have "[] \<in> ?X" by auto
+    have "finite ?N"
+      using assms
+      by auto
+    then have "finite ?X" by (rule finite_lists_length_le)
+    moreover have "range product_invariant \<subseteq> concat ` ?X"
+    proof
+      fix x assume "x \<in> range product_invariant"
+      then obtain L where L:
+            "x = concat
+                       (map (\<lambda>p. if p < length (map inv_of N)
+                                  then (map inv_of N ! p) (L ! p) else [])
+                         [0..<length L])"
+        unfolding product_invariant_def by auto
+      show "x \<in> concat ` ?X"
+      proof (cases "length L \<le> length N")
+        case True
+        then show ?thesis using L by (fastforce dest: nth_mem)
+      next
+        case False
+        then have "x = concat
+                       (map (\<lambda>p. if p < length (map inv_of N)
+                                  then (map inv_of N ! p) (L ! p) else [])
+                         [0..<length N])"
+          using L by (auto intro: concat_map_if_aux)
+        then show ?thesis by (fastforce dest: nth_mem)
+      qed
+    qed
+    ultimately show ?thesis by - (drule finite_subset; auto)
+  qed
+
+end (* End locale for product TA definition *)
+
+locale Product_TA_Defs' =
+  Product_TA_Defs N for N :: "('a, 'b, 'c, 't :: time, 's) nta"
+begin
+
+  lemma product_invariantD:
+    assumes "u \<turnstile> product_invariant L" "p < length N" "length L = length N"
+    shows "u \<turnstile> inv_of (N ! p) (L!p)"
+  using assms unfolding product_invariant_def by (force intro: concat_guard)
+
+  lemma states_step:
+    "L' \<in> states" if "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>a \<langle>L', u'\<rangle>" "L \<in> states"
+    using that unfolding states_def apply safe
+        apply simp
+       apply simp
+      apply force
+     apply (subst list_update_nth_split)
+      apply simp
+     apply (rule conjI)
+      apply clarsimp
+    subgoal premises prems for p g a r l'
+      using prems(3-6) by force
+     apply clarsimp
+     apply (solves auto)
+    apply (subst list_update_nth_split)
+     apply (simp; fail)
+    apply safe
+     apply simp
+    subgoal premises prems
+      using prems(3-7) by force
+    apply (subst list_update_nth_split)
+     apply simp
+    apply safe
+     apply simp
+    subgoal premises prems
+      using prems(3-6) by force
+    by auto
+
+  lemma states_steps:
+    "L' \<in> states" if "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>" "L \<in> states"
+    using that apply (induction N \<equiv> N _ _ _ _ rule: steps_n.induct)
+     apply assumption
+    apply simp
+    by (rule states_step; assumption)
+
+end
+
+lemma steps_altI:
+  "A \<turnstile> \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l'', Z''\<rangle>" if
+  "A \<turnstile> \<langle>l, Z\<rangle> \<rightarrow>* \<langle>l', Z'\<rangle>" "A \<turnstile> \<langle>l', Z'\<rangle> \<rightarrow> \<langle>l'', Z''\<rangle>"
+  using that by (induction; blast)
+
+(* Network + valid start state *)
+locale Product_TA =
+  Product_TA_Defs' N for N :: "('a, 'b, 'c, 't :: time, 's) nta" +
+  fixes L :: "'s list"
+  assumes states[intro]: "L \<in> states"
+begin
+
+  lemma
+    len[simp]: "length L = length N"
+  using states unfolding states_def by auto
+
+  lemma product_delay_complete:
+    assumes step: "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Del \<langle>L', u'\<rangle>"
+    obtains d where "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>d \<langle>L', u'\<rangle>"
+  using step proof cases
+    case prems: (step_n_t d)
+    from prems have *:
+      "\<forall>p\<in>{..<length N}. u \<oplus> d \<turnstile> inv_of (N ! p) (L ! p)"
+    by (auto elim!: step_t.cases)
+    from prems * show ?thesis
+    by (fastforce simp add: product_ta_def inv_of_def product_invariant_def[abs_def]
+                  intro!: guard_concat that
+       )
+  qed
+
+  lemma product_int_complete:
+    assumes step: "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Act b \<langle>L', u'\<rangle>"
+    obtains a where "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, Act b) \<langle>L', u'\<rangle>"
+    using step proof cases
+    case prems: (step_n_i p l g a r l')
+    from prems show ?thesis
+      apply -
+      apply (rule that)
+      apply (rule step_a.intros[where g = g and r = r])
+      by (fastforce
+          simp: product_trans_def product_trans_i_def product_invariant_def product_ta_def
+          trans_of_def inv_of_def
+          intro: guard_concat
+          )+
+  qed
+
+  lemma product_sync_complete:
+    assumes step: "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Syn b1 b2 \<langle>L', u'\<rangle>"
+    obtains a where "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, Syn b1 b2) \<langle>L', u'\<rangle>"
+    using step proof cases
+    case prems: (step_n_s p l1 g1 a r1 l1' q l2 g2 r2 l2')
+    from prems show ?thesis
+      apply -
+      apply (rule that)
+      apply (rule step_a.intros[where g = "g1 @ g2" and a = "(a, Syn b1 b2)" and r = "r1 @ r2"])
+      subgoal premises prems
+        apply
+          (clarsimp simp add:
+            product_trans_def product_trans_s_def product_invariant_def product_ta_def
+            trans_of_def inv_of_def
+            )
+        apply (erule allE[where x = p])
+        using \<open>p < _\<close> apply simp
+        apply (erule allE[where x = q])
+        using prems by (fastforce simp: trans_of_def)
+      by (fastforce
+            simp: product_invariant_def product_ta_def inv_of_def
+            intro: guard_concat guard_append
+         )+
+  qed
+
+  lemma product_complete:
+    assumes step: "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>b \<langle>L', u'\<rangle>"
+    shows "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>"
+    using step
+    by (cases b) (rule product_delay_complete product_int_complete product_sync_complete, simp, blast)+
+
+  lemma product_ta_cases:
+    assumes "product_ta \<turnstile> L \<longrightarrow>g,a,r L'"
+    shows "(L, g, a, r, L') \<in> product_trans_i \<or> (L, g, a, r, L') \<in> product_trans_s"
+  using assms unfolding product_ta_def trans_of_def product_trans_def by auto
+
+  lemma product_delay_sound:
+    assumes step: "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>d \<langle>L', u'\<rangle>"
+    shows "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Del \<langle>L', u'\<rangle>"
+  using assms by cases (force intro!: step_n_t product_invariantD len)
+
+  lemma product_action_sound:
+    assumes step: "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, b) \<langle>L', u'\<rangle>"
+    shows "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>b \<langle>L', u'\<rangle>"
+    using assms
+    apply cases
+    apply simp
+    apply (drule product_ta_cases)
+    apply (erule disjE)
+    unfolding product_trans_i_def product_trans_s_def
+     apply (solves \<open>auto intro!: product_invariantD step_n_i\<close>)
+    by (auto intro!: product_invariantD step_n_s simp: clock_val_def)
+
+  lemma product_sound:
+    assumes step: "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>"
+    obtains a where "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>a \<langle>L', u'\<rangle>"
+    using step by cases (force dest!: product_action_sound product_delay_sound)+
+
+  lemma states_product_step[intro]:
+    "L' \<in> states" if "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>"
+    by (auto intro: product_sound that elim!: states_step)
+
+  lemma product_steps_sound:
+    "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>" if "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>"
+    using that states proof (induction A \<equiv> product_ta _ _ _ _ rule: steps.induct)
+    case (refl l u)
+    then show ?case by blast
+  next
+    case prems: (step l u l' u' l'' u'')
+    interpret interp: Product_TA N l by standard (rule prems)
+    from prems have *: "N \<turnstile>N \<langle>l', u'\<rangle> \<rightarrow>* \<langle>l'',u''\<rangle>" by auto
+    from prems show ?case by - (erule interp.product_sound, rule stepI2, rule *)
+  qed
+
+  lemma product_steps_complete:
+    "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>" if "N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>"
+    using that states proof (induction N \<equiv> N _ _ _ _ rule: steps_n.induct)
+    case (refl l Z)
+    then show ?case by blast
+  next
+    case prems: (step l Z l' Z' _ l'' Z'')
+    interpret interp: Product_TA N l' by standard (blast intro: prems states_steps)
+    from prems show ?case by - (assumption | rule steps_altI interp.product_complete)+
+  qed
+
+  lemma product_correct:
+    "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle> \<longleftrightarrow> N \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>* \<langle>L', u'\<rangle>"
+    by (metis product_steps_complete product_steps_sound)
+
+  end (* End context: network + valid start state *)
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Networks/State_Networks.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Networks/State_Networks.thy
@@ -0,0 +1,782 @@
+theory State_Networks
+  imports Networks Munta_Base.Normalized_Zone_Semantics_Impl
+    Munta_Base.More_Methods
+begin
+
+unbundle no_library_syntax
+
+section \<open>Networks of Timed Automata With Discrete State\<close>
+
+subsection \<open>Syntax and Operational Semantics\<close>
+
+text \<open>
+  We extend Networks of Timed Automata with arbitrary shared (global) state.
+  Syntactically, this extension is very simple.
+  We can just use the free action label slot to annotate edges with a guard
+  and an update function on discrete states.
+  The slightly more clumsy part is adding invariants for discrete states
+  by directly specifying an invariant annotating function.
+\<close>
+
+type_synonym
+  ('a, 'c, 'time, 's, 'st) transition =
+  "'s * ('st \<Rightarrow> ('c, 'time) cconstraint) * 'a * ('st \<Rightarrow> 'c list) * 's"
+
+type_synonym
+  ('a, 'c, 'time, 's, 'st) sta = "('a, 'c, 'time, 's, 'st) transition set * ('c, 'time, 's) invassn"
+
+type_synonym
+  ('a, 'c, 't, 's, 'st) snta =
+  "('a act \<times> ('st \<Rightarrow> bool) \<times> ('st \<Rightarrow> 'st option), 'c, 't, 's, 'st) sta list \<times> ('s \<Rightarrow> 'st \<Rightarrow> bool) list"
+
+(*
+type_synonym
+  ('a, 'c, 'time, 's) unta = "programc \<times> ('a act, 'c, 'time, 's) uta list"
+
+type_synonym
+  ('a, 'c, 't, 's, 'st) snta =
+  "('a, ('st \<Rightarrow> bool) \<times> ('st \<Rightarrow> 'st), 'c, 't, 's) nta \<times> ('s \<Rightarrow> 'st \<Rightarrow> bool) list"
+*)
+
+text \<open>
+  Semantic states now consist of three things:
+  a list of process locations, the shared state, and a clock valuation.
+  The semantic extension then is also obvious: we can take the same transitions
+  as in the network without shared state, however we have to add state updates
+  and checks for guards on the shared state.
+  The updates on discrete state for synchronizing transitions are in the same order as in UPPAAL
+  (output before input).
+\<close>
+
+datatype 'b label = Del | Act 'b | Syn 'b
+
+inductive step_sn ::
+  "('a, 'c, 't, 's, 'st) snta \<Rightarrow> 's list \<Rightarrow> 'st \<Rightarrow> ('c, ('t::time)) cval \<Rightarrow> 'a label
+  \<Rightarrow> 's list \<Rightarrow> 'st \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+  ("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>_ \<langle>_, _, _\<rangle>" [61,61,61,61,61] 61)
+where
+  step_sn_t:
+    "(N, I) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L, s, u \<oplus> d\<rangle>"
+    if "\<forall> p \<in> {..<length N}. u \<oplus> d \<turnstile> snd (N ! p) (L ! p)"
+       "d \<ge> 0" "length N = length I" |
+  step_sn_i:
+    "(N, I) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Act a \<langle>L', s', u'\<rangle>"
+    if "(l, g, (Sil a, c, m), f, l') \<in> fst (N!p)"
+       "u \<turnstile> g s" "\<forall> p \<in> {..<length N}. u' \<turnstile> snd (N!p) (L'!p)"
+       "r = f s"
+       "L!p = l" "p < length L" "L' = L[p := l']" "u' = [r\<rightarrow>0]u"
+       "c s" "\<forall> p < length I. (I ! p) (L' ! p) s'" "Some s' = m s"
+       "length N = length I" |
+  step_sn_s:
+    "(N, I) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>"
+    if "(l1, g1, (In a, ci, mi), f1, l1') \<in> fst (N!p)"
+       "(l2, g2, (Out a, co, mo), f2, l2') \<in> fst (N!q)" "u \<turnstile> g1 s" "u \<turnstile> g2 s"
+       "\<forall> p \<in> {..<length N}. u' \<turnstile> snd (N!p) (L'!p)"
+       "r1 = f1 s" "r2 = f2 s"
+       "L!p = l1" "L!q = l2" "p < length L" "q < length L" "p \<noteq> q"
+       "L' = L[p := l1', q := l2']" "u' = [(r1 @ r2)\<rightarrow>0]u"
+       "ci s" "co s" "\<forall> p < length I. (I ! p) (L' ! p) s'"
+       "Some so = mo s" "Some s' = mi so" "length N = length I"
+
+inductive_cases[elim!]: "N \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>"
+
+inductive steps_sn ::
+  "('a, 'c, 't, 's, 'st) snta \<Rightarrow> 's list \<Rightarrow> 'st \<Rightarrow> ('c, ('t::time)) cval
+  \<Rightarrow> 's list \<Rightarrow> 'st \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>* \<langle>_, _, _\<rangle>" [61, 61, 61,61,61] 61)
+where
+  refl: "N \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L, s, u\<rangle>" |
+  step: "N \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<Longrightarrow> N \<turnstile> \<langle>L', s', u'\<rangle> \<rightarrow>l \<langle>L'', s'', u''\<rangle>
+        \<Longrightarrow> N \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+
+declare steps_sn.intros[intro]
+
+lemma stepI2:
+  "N \<turnstile> \<langle>l, s, u\<rangle> \<rightarrow>* \<langle>l'', s'', u''\<rangle>" if
+  "N \<turnstile> \<langle>l', s', u'\<rangle> \<rightarrow>* \<langle>l'', s'', u''\<rangle>" "N \<turnstile> \<langle>l, s, u\<rangle> \<rightarrow>a \<langle>l', s', u'\<rangle>"
+  using that
+  apply induction
+   apply rule
+    apply (rule refl)
+   apply assumption
+  apply simp
+  by (rule; assumption)
+
+abbreviation state_set :: "('a, 'c, 't, 's, 'st) transition set \<Rightarrow> 's set" where
+  "state_set T \<equiv> fst ` T \<union> (snd o snd o snd o snd) ` T"
+
+subsection \<open>Product Automaton\<close>
+
+locale Prod_TA_Defs =
+  fixes A :: "('a, 'c, 't, 's, 'st) snta"
+begin
+
+definition
+  "T_s p s = {(l, g s, a, f s, l') | l g a f l'. (l, g, a, f, l') \<in> fst (fst A ! p)}"
+
+definition
+  "N_s s = map (\<lambda> p. (T_s p s, snd (fst A ! p))) [0..<length (fst A)]"
+
+abbreviation "P \<equiv> snd A"
+
+definition "p \<equiv> length (fst A)"
+
+abbreviation "product s \<equiv> Product_TA_Defs.product_ta (N_s s)"
+
+abbreviation "T' s \<equiv> trans_of (product s)"
+abbreviation "I' s \<equiv> inv_of (product s)"
+
+definition
+  "prod_trans_i =
+    {((L, s), g, a, r, (L', s')) | L s g c a r m L' s'.
+     (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+     \<and> (L, g, (a, Networks.label.Act (c, m)), r, L') \<in> T' s \<and> c s \<and> Some s' = m s}"
+
+definition
+  "prod_trans_s =
+    {((L, s), g, a, r, (L', s')) | L s g ci co a r mi mo L' s' so.
+      ci s \<and> co s
+      \<and> (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+      \<and> (L, g, (a, Networks.label.Syn (ci, mi) (co, mo)), r, L') \<in> T' s
+      \<and> Some so = mo s
+      \<and> Some s' = mi so
+    }"
+
+  definition
+    "prod_trans \<equiv> prod_trans_i \<union> prod_trans_s"
+
+  definition
+    "prod_invariant \<equiv> \<lambda> (L, s). I' s L"
+
+  definition prod_ta :: "('a, 'c, 't, 's list \<times> 'st) ta" where
+    "prod_ta \<equiv> (prod_trans, prod_invariant)"
+
+  lemma prod_ta_cases:
+    assumes "prod_ta \<turnstile> L \<longrightarrow>g,a,r L'"
+    shows "(L, g, a, r, L') \<in> prod_trans_i \<or> (L, g, a, r, L') \<in> prod_trans_s"
+    using assms unfolding prod_ta_def trans_of_def prod_trans_def by auto
+
+  lemma inv_of_simp:
+    "inv_of prod_ta (L, s) = I' s L"
+    unfolding inv_of_def prod_ta_def prod_invariant_def by simp
+
+  lemma I'_simp:
+    "I' s L = I' s' L"
+    unfolding Product_TA_Defs.product_ta_def inv_of_def Product_TA_Defs.product_invariant_def N_s_def
+    apply simp
+    apply (rule arg_cong[where f = concat])
+    by simp
+
+  lemma collect_clki_prod_invariant:
+    "Timed_Automata.collect_clki prod_invariant = Timed_Automata.collect_clki (I' s)"
+    unfolding prod_invariant_def Timed_Automata.collect_clki_def
+    apply (simp split: prod.split)
+    apply safe
+     apply (subst (asm) I'_simp[where s' = s])
+    by auto
+
+  lemma collect_clki_prod_invariant':
+    "Timed_Automata.collect_clki prod_invariant
+    \<subseteq> \<Union> {Timed_Automata.collect_clki (snd (fst A ! p)) | p. p < length (fst A)}"
+    unfolding collect_clki_prod_invariant[of s]
+    unfolding inv_of_def Product_TA_Defs.product_ta_def
+    unfolding Product_TA_Defs.product_invariant_def
+    unfolding inv_of_def N_s_def
+    unfolding Timed_Automata.collect_clki_def
+    unfolding collect_clock_pairs_def
+    by auto
+
+  lemma collect_clkt_prod_trans_subs:
+    "Timed_Automata.collect_clkt prod_trans \<subseteq> Timed_Automata.collect_clkt (\<Union> (T' ` UNIV))"
+    unfolding Timed_Automata.collect_clkt_def prod_trans_def prod_trans_i_def prod_trans_s_def
+    by fastforce
+
+  lemma collect_clkvt_prod_trans_subs:
+    "collect_clkvt prod_trans \<subseteq> collect_clkvt (\<Union> (T' ` UNIV))"
+    unfolding collect_clkvt_def prod_trans_def prod_trans_i_def prod_trans_s_def by fastforce
+
+lemma T_simp:
+  "T ! q = trans_of (N ! q)" if "q < length N"
+  using that oops
+
+  (*
+lemma prod_state_set_subs:
+  assumes "l \<in> state_set T'" "q < p"
+  shows "l ! q \<in> state_set (trans_of (N ! q))"
+  using assms
+  apply (simp only: T_simp[symmetric] p_def)
+  by (rule product_state_set_subs; simp add: product_ta_def trans_of_def)
+*)
+
+abbreviation "N \<equiv> fst A"
+
+context
+  fixes Q
+  assumes finite_state:
+    "\<forall> l. \<forall> q < p. (P ! q) l s \<longrightarrow> Q s"
+    "finite {s. Q s}"
+      and finite_trans: "\<forall> A \<in> set N. finite (fst A)"
+      and p_gt_0: "p > 0"
+begin
+
+  lemma finite_state':
+    "finite {s. \<forall>q<p. (P ! q) (L ! q) s}" (is "finite ?S")
+  proof -
+    from p_gt_0 obtain q where "q < p" by blast
+    then have "?S \<subseteq> {s. Q s}" using finite_state(1) by auto
+    moreover have "finite \<dots>" by (rule finite_state(2))
+    ultimately show ?thesis by (rule finite_subset)
+  qed
+
+  lemma finite_trans':
+    "\<forall>A\<in>set (N_s s). finite (trans_of A)"
+  unfolding N_s_def apply clarsimp
+    unfolding trans_of_def T_s_def
+    apply simp
+    apply (drule nth_mem)
+    using finite_trans
+    using [[simproc add: finite_Collect]]
+    apply simp
+    apply (rule finite_imageI)
+    apply (rule finite_vimageI)
+     apply simp
+    unfolding inj_on_def by auto
+
+  lemma finite_states:
+    "finite (Product_TA_Defs.states (N_s s))"
+    using finite_trans' by (rule Product_TA_Defs.finite_states)
+
+  lemma
+    "finite (T' s)"
+    using finite_trans' by (rule Product_TA_Defs.finite_trans_of_product)
+
+  (* Duplicated proof, what is the better way? *)
+  lemma finite_product_1:
+    "finite (T' s)"
+    unfolding product_def
+    unfolding trans_of_def Product_TA_Defs.product_ta_def
+    apply simp
+    unfolding Product_TA_Defs.product_trans_def
+  proof safe
+    have "Product_TA_Defs.product_trans_i (N_s s)
+        \<subseteq> {(L, g, (a, Networks.label.Act (aa, b)), r, L[p := l']) |L p g a aa b r l'.
+            L \<in> Product_TA_Defs.states (N_s s) \<and> p < length (N_s s) \<and>
+            (L ! p, g, (Sil a, aa, b), r, l') \<in> \<Union> (trans_of ` set (N_s s))}"
+      unfolding Product_TA_Defs.product_trans_i_def
+      by (fastforce simp: Product_TA_Defs.states_length)
+    moreover have "finite \<dots>"
+      apply defer_ex
+      using finite_states[of s] apply clarsimp
+      apply (subst finite_Collect_bounded_ex_6)
+      subgoal premises prems for y y'
+      proof -
+        have "
+              {(a, b, c, d, e, f). \<exists>x\<in>set (N_s s). x \<turnstile> y ! y' \<longrightarrow>a,(Sil b, c, d),e f}
+            = {xx. \<exists> x a b c d e f. x\<in>set (N_s s) \<and> x \<turnstile> y ! y' \<longrightarrow>a,(Sil b, c, d),e f
+                   \<and> xx = (a, b, c, d, e, f)}"
+          by force
+        moreover have "finite \<dots>" (* finite_Collect_bounded_ex is not crucial here *)
+          using finite_trans'[of s]
+          using [[simproc add: finite_Collect]]
+          by (auto simp: inj_on_def intro: finite_vimageI simp del: finite_Collect_bounded_ex)
+        ultimately show ?thesis by simp
+      qed
+      by auto
+    ultimately show "finite (Product_TA_Defs.product_trans_i (N_s s))" by (rule finite_subset)
+  next
+    have "Product_TA_Defs.product_trans_s (N_s s)
+        \<subseteq> {(L, g1 @ g2, (a, Networks.label.Syn b1 b2), r1 @ r2, L[p := l1', q := l2']) |
+              L p q g1 g2 a b1 b2 r1 r2 l1' l2'.
+              L \<in> Product_TA_Defs.states (N_s s) \<and>
+              p < length (N_s s) \<and> q < length (N_s s) \<and>
+              (L ! p, g1, (In a, b1), r1, l1') \<in> map trans_of (N_s s) ! p \<and>
+              (L ! q, g2, (Out a, b2), r2, l2') \<in> map trans_of (N_s s) ! q \<and> p \<noteq> q}"
+      unfolding Product_TA_Defs.product_trans_s_def
+      by (fastforce simp: Product_TA_Defs.states_length)
+    moreover have "finite \<dots>"
+      apply defer_ex
+      using finite_states[of s]
+      apply clarsimp
+      subgoal
+        apply (mini_existential, simp)
+        apply (mini_existential, simp)
+        apply (mini_existential, simp)
+        apply (mini_existential, simp)
+        apply (subst finite_Collect_bounded_ex_6)
+        subgoal
+          using [[simproc add: finite_Collect]] finite_trans'[of s]
+          by (auto simp: inj_on_def intro: finite_vimageI)
+        apply safe
+        apply (subst finite_Collect_bounded_ex_5)
+        subgoal
+          using [[simproc add: finite_Collect]] finite_trans'[of s]
+          by (auto 4 3 simp: simp: inj_on_def intro: finite_vimageI)
+        by auto
+      done
+    ultimately show "finite (Product_TA_Defs.product_trans_s (N_s s))" by (rule finite_subset)
+  qed
+
+  lemma prod_trans_i_alt_def:
+    "prod_trans_i =
+      {((L, s), g, a, r, (L', s')) | L s g c a r m L' s'.
+       (L, g, (a, Networks.label.Act (c, m)), r, L') \<in> T' s \<and>
+       (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+       \<and> c s \<and> Some s' = m s}"
+    unfolding prod_trans_i_def by (safe; metis)
+
+  lemma Some_finite:
+    "finite {x. Some x = y}"
+    using not_finite_existsD by fastforce
+
+  lemma finite_prod_trans:
+    "finite prod_trans" if "p > 0"
+    unfolding prod_trans_def
+  proof safe
+    have "prod_trans_i \<subseteq>
+        {((L, s), g, a, r, (L', s')) | L s g c a r m L' s'.
+         Q s \<and>
+         (L, g, (a, Networks.label.Act (c, m)), r, L') \<in> T' s \<and>
+         (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+         \<and> c s \<and> Some s' = m s}
+      "
+      unfolding prod_trans_i_alt_def
+      using finite_state(1) p_gt_0 by force
+    moreover have
+      "finite \<dots>"
+      apply defer_ex
+      apply (mini_existential, simp only: ex_simps)
+      using finite_state(2) apply clarsimp
+      apply (subst finite_Collect_bounded_ex_7)
+      using [[simproc add: finite_Collect]] finite_state' finite_product_1
+      by (auto 4 3 simp: inj_on_def intro: finite_vimageI)
+    ultimately show "finite prod_trans_i" by (rule finite_subset)
+  next
+    have "prod_trans_s \<subseteq>
+        {((L, s), g, a, r, (L', s')) | L s g ci co a r mi mo L' s' so.
+          Q s \<and>
+          product s \<turnstile> L \<longrightarrow>g,(a, Networks.label.Syn (ci, mi) (co, mo)),r L' \<and>
+          (\<forall>q<p. (P ! q) (L ! q) s) \<and> (\<forall>q<p. (P ! q) (L' ! q) s') \<and>
+          ci s \<and> co s \<and> Some so = mo s \<and> Some s' = mi so}
+      "
+      unfolding prod_trans_s_def
+      using finite_state(1) p_gt_0 by fastforce
+    moreover have
+      "finite \<dots>"
+      apply defer_ex
+      apply (mini_existential, simp only: ex_simps)
+      using finite_state(2) apply clarsimp
+      apply (subst finite_Collect_bounded_ex_9)
+
+      subgoal
+        using [[simproc add: finite_Collect]] finite_state' finite_product_1
+        by (auto 4 3 simp: inj_on_def intro: finite_vimageI)[]
+
+      apply safe
+      subgoal for s a b c d e f g h i
+        apply (rule finite_subset[where B =
+              "(\<lambda> s'. ((a, s), b, e, f, i, s')) ` { s'. \<exists> so. Some so = h s \<and> Some s' = g so}"
+              ])
+         apply force
+        apply (rule finite_imageI)
+        apply (subst finite_Collect_bounded_ex)
+        by (force intro: Some_finite)+
+      done
+    ultimately show "finite prod_trans_s" by (rule finite_subset)
+  qed
+
+end (* End of context for finiteness of automaton *)
+
+  abbreviation "states' s \<equiv> Product_TA_Defs.states (N_s s)"
+
+  lemma N_s_length:
+    "length (N_s s) = p"
+    unfolding N_s_def p_def by simp
+
+end (* End locale for product TA definition *)
+
+locale Prod_TA_Defs' =
+  Prod_TA_Defs A for A :: "('a, 'c, 't :: time, 's, 'st) snta"
+begin
+
+lemma A_unfold:
+  "A \<equiv> (N, P)"
+  by auto
+
+lemma network_step_delay:
+  assumes step: "(N, P) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>" and len: "length L = p"
+  shows "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.label.Del \<langle>L', u'\<rangle>"
+  subgoal
+    using step
+    apply cases
+    subgoal
+      apply simp
+      apply (rule step_n_t)
+      subgoal
+        unfolding N_s_def by (auto simp: inv_of_def)
+      apply assumption
+      done
+    done
+  done
+
+lemma network_step_silent:
+  assumes step: "(N, P) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Act a \<langle>L', s', u'\<rangle>" and len: "length L = p"
+  obtains a where "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.label.Act a \<langle>L', u'\<rangle>"
+  subgoal premises prems
+    using step
+    apply cases
+    subgoal
+      apply (rule prems)
+      apply (rule step_n_i)
+      unfolding N_s_def T_s_def by (auto 4 0 simp: trans_of_def inv_of_def len p_def)
+    done
+  done
+
+lemma network_step_sync:
+  assumes step: "(N, P) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>" and len: "length L = p"
+  obtains a b where "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.label.Syn a b \<langle>L', u'\<rangle>"
+  subgoal premises prems
+    using step
+    apply cases
+    subgoal
+      subgoal premises A
+        apply (rule prems)
+        apply (rule step_n_s)
+                   defer
+                   defer
+                   apply (rule A; fail)
+                  apply (rule A(4); fail)
+        subgoal
+          using A unfolding N_s_def by (auto simp: inv_of_def len)
+                defer
+                defer
+                apply (rule A; fail)
+               apply (rule A(11); fail)
+        using A unfolding N_s_def T_s_def by (auto 4 0 simp: trans_of_def len p_def)
+      done
+    done
+  done
+
+lemma network_step:
+  assumes step: "(N, P) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" and len: "length L = p"
+  obtains a where "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>a \<langle>L', u'\<rangle>"
+  subgoal
+    using step
+    apply (cases a; simp)
+      apply (rule that, erule network_step_delay[OF _ len, simplified])
+     apply (erule network_step_silent[OF _ len, simplified], erule that)
+    apply (erule network_step_sync[OF _ len, simplified], erule that)
+    done
+  done
+
+lemma trans_of_N_s_1:
+  "(fst ` trans_of (N_s s ! q)) = fst ` fst (N ! q)" if "q < p"
+  using that unfolding trans_of_def N_s_def p_def T_s_def by (auto 0 7 simp: image_iff)
+
+lemma trans_of_N_s_2:
+  "((snd o snd o snd o snd) ` trans_of (N_s s ! q)) = (snd o snd o snd o snd) ` fst (N ! q)" if "q < p"
+  using that unfolding trans_of_def N_s_def p_def T_s_def by force
+
+lemma
+  "fst ` trans_of (N_s s ! q) = fst ` trans_of (N_s s' ! q)" if "q < p"
+  using that by (simp add: trans_of_N_s_1)
+
+lemma states'_simp:
+  "states' s = states' s'"
+  unfolding Product_TA_Defs.states_def using trans_of_N_s_1 trans_of_N_s_2 by (simp add: N_s_length)
+
+  lemma states_step:
+    "L' \<in> states' s" if "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" "L \<in> states' s"
+  proof -
+    interpret Product_TA_Defs' "N_s s" .
+    from \<open>L \<in> _\<close> have "L \<in> states" .
+    from \<open>L \<in> _\<close> have "length L = p" by (simp add: N_s_length states_length)
+    with network_step[folded A_unfold, OF that(1)] obtain a where
+      "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>a \<langle>L',u'\<rangle>"
+      by auto
+    then show ?thesis using that(2) by (rule states_step)
+  qed
+
+  lemma states_steps:
+    "L' \<in> states' s'" if "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" "L \<in> states' s"
+    using that proof (induction A \<equiv> A _ _ _ _ _ _ rule: steps_sn.induct)
+    case (refl L s u)
+    then show ?case by assumption
+  next
+    case (step L s u L' s' u' L'' s'' u'')
+    with states_step[of L' s' u' L'' s'' u''] states'_simp show ?case by blast
+  qed
+
+  lemma inv_step:
+    "\<forall>p<length P. (P ! p) (L' ! p) s'" if
+    "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" "\<forall>p<length P. (P ! p) (L ! p) s"
+    using that by (cases) auto
+
+  lemma inv_steps:
+    "\<forall>p<length P. (P ! p) (L' ! p) s'" if
+    "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" "\<forall>p<length P. (P ! p) (L ! p) s"
+    using that by (induction A \<equiv> A _ _ _ _ _ _ rule: steps_sn.induct) (auto dest: inv_step)
+
+end
+
+(* Network + valid start state *)
+locale Prod_TA =
+  Prod_TA_Defs' A for A :: "('a, 'c, 't :: time, 's, 'st) snta" +
+  fixes L :: "'s list" and s :: 'st
+  assumes states[intro]: "L \<in> states' s"
+  assumes Len: "length N = length P"
+      and inv: "\<forall>p<length P. (P ! p) (L ! p) s"
+begin
+
+  sublocale Product_TA "N_s s" L by standard rule
+
+  lemma inv_prod_simp:
+    "inv_of prod_ta (l, s') = Product_TA_Defs.product_invariant (N_s s') l" if "length l = p"
+    unfolding prod_ta_def prod_invariant_def Product_TA_Defs.product_ta_def N_s_def inv_of_def
+    using that by (simp add: p_def)
+
+  lemma inv_of_N_simp:
+    "map inv_of (N_s s') ! q = I ! q" if "q < p"
+    using that unfolding inv_of_def N_s_def p_def by simp
+
+  lemma product_inv_prod_simp:
+    "inv_of prod_ta (l, s') = I' s l" if "length l = p"
+    using that
+    apply (simp add: inv_prod_simp)
+    apply (simp add: N_s_length inv_of_def Product_TA_Defs.product_invariant_def)
+    apply (rule arg_cong[where f = concat])
+    apply (clarsimp cong: map_cong)
+    by (subst inv_of_N_simp; simp)
+
+  lemma product_inv_prod[intro]:
+    "u \<turnstile> inv_of prod_ta (l, s')" if "u \<turnstile> inv_of product_ta l" "length l = p"
+    using that by (simp add: product_inv_prod_simp)
+
+  lemma A_simp[simp]:
+    "N' = N" "P' = P" if "A = (N', P')"
+    using that by auto
+
+  lemma length_L[intro]:
+    "length L = p"
+    by (simp add: N_s_length)
+
+  lemma prod_complete_delay:
+    assumes step: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+    obtains d where "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>"
+  using step proof cases
+    case prems: (step_sn_t N d P)
+    note [simp] = A_simp[OF prems(1)]
+    from prems have "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.Del \<langle>L', u'\<rangle>"
+      unfolding N_s_def by (auto 4 3 simp: inv_of_def intro: step_n_t)
+    with prems show ?thesis
+      by (auto 4 4
+          intro: that
+          simp: product_inv_prod_simp[OF length_L]
+          elim!: product_delay_complete step_t.cases)
+  qed
+
+  lemma prod_complete_silent:
+    assumes step: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Act a \<langle>L', s', u'\<rangle>"
+    obtains a where "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+  using step proof cases
+    case prems: (step_sn_i l g c m f l' N q r I)
+    note [simp] = A_simp[OF prems(1)]
+    from prems(13) have [simp]: "length P = p" by (simp add: p_def)
+    have "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.label.Act (c, m) \<langle>L', u'\<rangle>"
+      apply (rule step_n_i)
+      using prems unfolding N_s_def T_s_def by (auto 3 0 simp: trans_of_def inv_of_def N_s_length)
+    with \<open>length P = p\<close> obtain b where
+      "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(b, Networks.label.Act (c, m)) \<langle>L', u'\<rangle>"
+      by (clarsimp elim!: product_int_complete)
+    with prems inv obtain g r where step:
+      "((L, s), g, b, r, (L', s')) \<in> prod_trans_i"
+      "u \<turnstile> g" "[r\<rightarrow>0]u = u'" "u' \<turnstile> inv_of product_ta L'"
+        apply atomize_elim
+      unfolding prod_trans_i_def by - (erule step_a.cases; auto)
+    then have "((L, s), g, b, r, (L', s')) \<in> trans_of prod_ta"
+      by (simp add: prod_trans_def trans_of_def prod_ta_def)
+    moreover have "length L' = p"
+      using length_L prems(8) by auto
+    ultimately show ?thesis
+      apply -
+      apply (rule that)
+      apply rule
+      using step(2-) by force+
+  qed
+
+  lemma prod_complete_sync:
+    assumes step: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>"
+    obtains a where "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+  using step proof cases
+    case prems: (step_sn_s l1 g1 ci mi f1 l1' N q1 l2 g2 co mo f2 l2' q2 r1 r2 I so)
+    note [simp] = A_simp[OF prems(1)]
+    from prems(21) have [simp]: "length P = p" by (simp add: p_def)
+    (* Clean *)
+    have "N_s s \<turnstile>N \<langle>L, u\<rangle> \<rightarrow>Networks.label.Syn (ci, mi) (co, mo) \<langle>L', u'\<rangle>"
+        apply (rule step_n_s)
+                   defer
+                   defer
+                   apply (rule prems; fail)
+                  apply (rule prems(5); fail)
+        subgoal
+          using prems unfolding N_s_def by (auto simp: inv_of_def)
+                defer
+                defer
+                apply (rule prems; fail)
+               apply (rule prems(12); fail)
+        using prems unfolding N_s_def T_s_def by (auto 3 0 simp: trans_of_def p_def N_s_length)
+    with \<open>length P = p\<close> obtain a where
+      "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, Networks.label.Syn (ci, mi) (co, mo)) \<langle>L', u'\<rangle>"
+      by (auto elim!: product_sync_complete)
+    with prems inv obtain g r where step:
+        "((L, s), g, a, r, (L', s')) \<in> prod_trans_s"
+        "u \<turnstile> g" "[r\<rightarrow>0]u = u'" "u' \<turnstile> inv_of product_ta L'"
+        apply atomize_elim
+      unfolding prod_trans_s_def by - (erule step_a.cases; auto; blast) (* Slow *)
+    then have "((L, s), g, a, r, (L', s')) \<in> trans_of prod_ta"
+      by (simp add: prod_trans_def trans_of_def prod_ta_def)
+    moreover have "length L' = p"
+      using length_L \<open>L' = _\<close> by auto
+    ultimately show ?thesis
+      apply -
+      apply (rule that)
+      apply rule
+      using step(2-) by force+
+  qed
+    
+  lemma prod_complete:
+    assumes step: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+    shows "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+    using step
+    by (cases a; simp; blast elim!: prod_complete_delay prod_complete_silent prod_complete_sync)
+
+  lemma A_unfold:
+    "A = (N, P)"
+    by simp
+
+  lemma prod_sound'_delay:
+    assumes step: "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>"
+    shows "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle> \<and> product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>
+           \<and> (\<forall>p<length P. (P ! p) (L' ! p) s')"
+    using assms proof cases
+    case prems: 1
+    then have "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>d \<langle>L', u'\<rangle>" unfolding inv_of_simp by fast
+    moreover from product_delay_sound[OF this] prems(1-3) have "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+      apply simp
+      apply (subst A_unfold)
+      apply (rule step_sn_t)
+      by (auto simp: N_s_def inv_of_def step_t.simps N_s_length p_def Len intro: \<open>0 \<le> d\<close>)
+    ultimately show ?thesis using prems inv by fast
+  qed
+
+  lemma prod_sound'_action:
+    assumes step: "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+    obtains a where "(A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> a \<noteq> Del) \<and> product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>
+             \<and> (\<forall>p<length P. (P ! p) (L' ! p) s')"
+    using assms
+  proof cases
+    case prems: (1 g r)
+    from Len have [simp]: "length P = p" by (simp add: p_def)
+    from prems(1)[THEN prod_ta_cases] show ?thesis
+    proof (rule disjE, goal_cases)
+      case 1
+      then obtain c m where *:
+        "Some s' = m s" "\<forall>q<p. (P ! q) (L ! q) s" "\<forall>q<p. (P ! q) (L' ! q) s'"
+        "product_ta \<turnstile> L \<longrightarrow>g,(a, Networks.label.Act (c, m)),r L'" "c s"
+        unfolding prod_trans_i_def by auto
+      with prems have "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, Networks.label.Act (c, m)) \<langle>L', u'\<rangle>"
+        unfolding inv_of_simp by (metis I'_simp step_a.intros)
+      moreover from product_action_sound[OF this] prems(3-4) obtain a where
+        "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Act a \<langle>L', s', u'\<rangle>"
+        apply safe
+        apply (simp add: N_s_def trans_of_def N_s_length T_s_def)
+        apply (simp only: ex_simps[symmetric])
+        apply (erule exE, erule exE)
+        apply (rule that)
+        apply (subst A_unfold)
+        apply (rule step_sn_i)
+                   apply fast
+        using *(3) by (auto simp: N_s_def inv_of_def p_def \<open>Some s' = m s\<close> intro: \<open>c s\<close>)
+      ultimately show ?thesis using * by (auto intro: that)
+    next
+      case 2
+      then obtain ci co mi mo si where *:
+        "Some s' = mi si" "Some si = mo s" "\<forall>q<p. (P ! q) (L ! q) s" "\<forall>q<p. (P ! q) (L' ! q) s'"
+        "product_ta \<turnstile> L \<longrightarrow>g,(a, Networks.label.Syn (ci, mi) (co, mo)),r L'"
+        "ci s" "co s"
+        unfolding prod_trans_s_def by auto
+      with prems have "product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow>(a, Networks.label.Syn (ci, mi) (co, mo)) \<langle>L', u'\<rangle>"
+        unfolding inv_of_simp by (metis I'_simp step_a.intros)
+      moreover from product_action_sound[OF this] prems(3-4) obtain a where
+        "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>"
+        apply safe
+        apply (simp add: N_s_def trans_of_def N_s_length T_s_def)
+        apply (simp only: ex_simps[symmetric])
+        apply (erule exE, erule exE, erule exE, erule exE)
+        apply (rule that)
+        apply (subst A_unfold)
+        apply (rule step_sn_s)
+                           apply fast
+                          apply blast
+        using *(4) by (auto simp: N_s_def inv_of_def p_def \<open>Some s' = _\<close> \<open>Some si = _\<close> intro: *(6-)) (* Slow *)
+      ultimately show ?thesis using * by (intro that conjI) auto
+    qed
+  qed
+
+  lemma prod_sound':
+    assumes step: "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+    obtains a where "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> product_ta \<turnstile> \<langle>L, u\<rangle> \<rightarrow> \<langle>L', u'\<rangle>
+             \<and> (\<forall>p<length P. (P ! p) (L' ! p) s')"
+    using assms apply cases
+     apply (erule prod_sound'_action; blast intro: that)
+    apply (rule that, erule prod_sound'_delay)
+    done
+
+  lemmas prod_sound = prod_sound'[THEN conjunct1]
+  lemmas prod_inv_1 = prod_sound'[THEN conjunct2, THEN conjunct1]
+  lemmas prod_inv_2 = prod_sound'[THEN conjunct2, THEN conjunct2]
+
+  lemma states_prod_step[intro]:
+    "L' \<in> states" if "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+    by (blast intro: prod_inv_1[OF that])
+
+  lemma inv_prod_step[intro]:
+    "\<forall>p<length P. (P ! p) (L' ! p) s'" if "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+    using that by (blast intro: prod_inv_2)
+
+  lemma prod_steps_sound:
+    assumes step: "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>"
+    shows "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>"
+    using step states inv
+  proof (induction A \<equiv> prod_ta l \<equiv> "(L, s)" _ l' \<equiv> "(L', s')" _ arbitrary: L s rule: steps.induct)
+    case (refl u)
+    then show ?case by blast
+  next
+    case prems: (step u l' u' u'' L s)
+    obtain L'' s'' where "l' = (L'', s'')" by force
+    interpret interp: Prod_TA A L s by (standard; rule prems Len)
+    from prems(3)[OF \<open>l' = _\<close>] prems(1,2,4-) have *: "A \<turnstile> \<langle>L'', s'', u'\<rangle> \<rightarrow>* \<langle>L', s', u''\<rangle>"
+      unfolding \<open>l' = _\<close>
+      by (metis Prod_TA_Defs'.states'_simp interp.states_prod_step interp.inv_prod_step)
+    show ?case
+      using * prems by (auto simp: \<open>l' = _\<close> intro: interp.prod_sound stepI2)
+  qed
+
+  lemma prod_steps_complete:
+    "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>" if "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>"
+    using that states inv proof (induction A \<equiv> A L _ _ _ _ _ rule: steps_sn.induct)
+    case (refl L s u)
+    then show ?case by blast
+  next
+    case prems: (step L s u L' s' u' L'' s'' u'')
+    interpret interp: Prod_TA A L' s' apply standard
+      using prems by - (assumption | rule Prod_TA_Defs'.states_steps Len Prod_TA_Defs'.inv_steps)+
+    from prems show ?case by - (rule steps_altI, auto intro!: interp.prod_complete)
+  qed
+
+  lemma prod_correct:
+    "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<longleftrightarrow> A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>"
+    by (metis prod_steps_complete prod_steps_sound)
+
+  end (* End context: network + valid start state *)
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Parsing/JSON_Parsing.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Parsing/JSON_Parsing.thy
@@ -0,0 +1,161 @@
+section \<open>JSON Parsing\<close>
+theory JSON_Parsing
+  imports Lexer
+begin
+
+definition [consuming]:
+  "brace_open \<equiv> lx_ws *-- exactly ''{''"
+
+definition [consuming]:
+ "brace_close \<equiv> lx_ws *-- exactly ''}''"
+
+definition [consuming]:
+  "bracket_open \<equiv> lx_ws *-- exactly ''[''"
+
+definition [consuming]:
+  "bracket_close \<equiv> lx_ws *-- exactly '']''"
+
+definition [consuming]:
+  "colon \<equiv> lx_ws *-- exactly '':''"
+
+definition [consuming]:
+  "comma \<equiv> lx_ws *-- exactly '',''"
+
+definition json_character :: "(char,char) parser" where [consuming]:
+  \<comment>\<open>This is both, too permissive and too strict compared to JSON\<close>
+  "json_character \<equiv> do {
+    x\<leftarrow>get;
+    if x \<notin> set [CHR ''\"'', CHR 0x92, CHR ''\<newline>'']
+    then return x
+    else err_expecting (shows_string ''JSON string character'') }"
+
+definition [consuming]:
+  "json_string = exactly ''\"'' *-- repeat json_character --* exactly ''\"''"
+  \<comment>\<open>Less permissive than JSON strings\<close>
+
+definition [consuming]:
+  "identifier = exactly ''\"'' *-- repeat1 json_character --* exactly ''\"''"
+
+fun nats_to_nat :: "nat \<Rightarrow> nat list \<Rightarrow> nat" where
+  "nats_to_nat x [] = x" |
+  "nats_to_nat x (n # ns) = nats_to_nat (10 * x + n) ns"
+
+(* definition [consuming]:
+  "lx_rat_fract \<equiv> repeat1 lx_digit'
+    with (\<lambda>ns. Fract (int (nats_to_nat 0 ns)) (int (10 ^ length ns)))"
+
+definition lx_rat[consuming]:
+  "lx_rat \<equiv> lx_int -- exactly ''.'' *-- lx_rat_fract
+    with (\<lambda> (a, b). if a \<ge> 0 then of_int a + b else of_int a - b)" *)
+
+datatype fract = Rat bool int int
+
+definition lx_rat[consuming]:
+  "lx_rat \<equiv> lx_int -- exactly ''.'' *-- (repeat1 lx_digit' with int o nats_to_nat 0)
+    with (\<lambda> (a, b). if a \<ge> 0 then Rat True a b else Rat False a b)"
+
+
+definition
+  "parse_list a \<equiv> chainL1 (a with (\<lambda>a. [a])) (do {comma; return (@)})"
+
+abbreviation
+  "parse_list' a \<equiv> parse_list a \<parallel> lx_ws with (\<lambda>_. [])"
+
+lemma parse_list_cong[fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> A l2 = A' l2"
+  assumes "l=l'"
+  shows "parse_list A l = parse_list A' l'"
+  unfolding parse_list_def chainL1_def
+  apply (intro bind_cong repeat_cong assms order_refl)
+  by auto
+
+datatype JSON =
+  is_obj: Object "(string \<times> JSON) list"
+| is_array: Array "JSON list"
+| is_string: String string \<comment>\<open>An Isabelle string rather than a JSON string\<close>
+| is_int: Int int \<comment>\<open>The number type is split into natural Isabelle/HOL types\<close>
+| is_nat: Nat nat
+| is_rat: Rat fract
+| is_boolean: Boolean bool \<comment>\<open>True and False are contracted to one constructor\<close>
+| is_null: Null
+
+definition [consuming]:
+  "atom \<equiv> lx_ws *-- (
+     json_string       with String
+   \<parallel> lx_rat            with JSON.Rat
+   \<parallel> lx_nat            with Nat
+   \<parallel> lx_int            with JSON.Int
+   \<parallel> exactly ''true''  with (\<lambda>_. Boolean True)
+   \<parallel> exactly ''false'' with (\<lambda>_. Boolean False)
+   \<parallel> exactly ''null''  with (\<lambda>_. Null)
+)"
+
+fun json and dict and seq where
+  "json ::= atom \<parallel> seq \<parallel> dict" |
+  "dict ::=
+    brace_open *--
+      parse_list' (lx_ws *-- identifier -- colon *-- json)
+    --* brace_close with Object" |
+  "seq  ::=
+    bracket_open *--
+      parse_list' json
+    --* bracket_close with Array"
+
+
+paragraph \<open>Test Cases\<close>
+
+definition test where "test \<equiv> STR
+''
+{
+ \"a\": [\"b\", {\"a\":\"a\",\"b\":[\"b\" ,\"d\"]}] ,
+
+ \"c\" : \"c\",
+
+ \"ab\" : 1,
+ \"abcd1\" : -11,
+
+ \"abc\" : 11.1,
+\"_abc-\" : true,
+\"ab_c\" : false,
+\"a-bc\":null
+}''
+"
+
+value [code]
+  "parse_all lx_ws json test"
+
+definition test2 where "test2 \<equiv>
+ STR ''{
+    \"info\": \"Derived from the Uppaal benchmarks found at https://www.it.uu.se/research/group/darts/uppaal/benchmarks/\",
+    \"automata\": [
+        {
+            \"name\": \"ST2\",
+            \"initial\": 9,
+            \"nodes\": [
+                {
+                    \"id\": 14,
+                    \"name\": \"y_idle\",
+                    \"x\": 458.2894592285156,
+                    \"y\": 442.4790954589844,
+                    \"invariant\": \"\"
+                }
+            ],
+            \"edges\": [
+                {
+                    \"source\": 9,
+                    \"target\": 12,
+                    \"guard\": \"\",
+                    \"label\": \"tt1?\",
+                    \"update\": \"y2 := 0, x2 := 0\"
+                }
+            ]
+        }
+    ],
+    \"clocks\": \"x0, x1, y1, z1, x2, y2, z2\",
+    \"vars\": \"id[-1:2]\",
+    \"formula\": \"E<> ((ST1.z_sync || ST1.z_async)  && (ST2.z_sync || ST2.z_async)  )\"
+}''"
+
+value [code] "parse_all lx_ws json test2"
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Parsing/Lexer.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Parsing/Lexer.thy
@@ -0,0 +1,42 @@
+(* Author: Peter Lammich *)
+section \<open>Basic Lexing\<close>
+
+theory Lexer
+  imports Parser_Combinator
+begin
+
+\<comment> \<open>We start by defining a lexer\<close>
+definition "lx_ws \<equiv> repeat (any char_wspace)"
+abbreviation "token \<equiv> gen_token lx_ws"
+
+definition [consuming]: "tk_plus = token (exactly ''+'')"
+definition [consuming]: "tk_times \<equiv> token (exactly ''*'')"
+definition [consuming]: "tk_minus \<equiv> token (exactly ''-'')"
+definition [consuming]: "tk_div \<equiv> token (exactly ''/'')"
+definition [consuming]: "tk_power \<equiv> token (exactly ''^'')"
+definition [consuming]: "tk_lparen \<equiv> token (exactly ''('')"
+definition [consuming]: "tk_rparen \<equiv> token (exactly '')'')"
+
+abbreviation "additive_op \<equiv> 
+  tk_plus \<then> return (+)
+\<parallel> tk_minus \<then> return (-)"
+abbreviation "multiplicative_op \<equiv>
+  tk_times \<then> return (*)
+\<parallel> tk_div \<then> return (div)"
+abbreviation "power_op \<equiv>
+  tk_power \<then> return (\<lambda>a b. a^nat b)"
+
+abbreviation "lx_digit' \<equiv> lx_digit with (\<lambda>d. nat_of_char d - nat_of_char CHR ''0'')"
+
+\<comment> \<open>We convert the string to a number while parsing, using a parameterized parser.
+   NB: Allows leading zeroes\<close>
+fun lx_nat_aux :: "nat \<Rightarrow> (char,nat) parser" where
+  " lx_nat_aux acc ::= do { d \<leftarrow> lx_digit'; lx_nat_aux (10*acc + d) }
+  \<parallel> return acc"
+
+definition [consuming]: "lx_nat \<equiv> lx_digit' \<bind> lx_nat_aux"
+definition [consuming]: "lx_int \<equiv> exactly ''-'' *-- lx_nat with uminus o int \<parallel> lx_nat with int"
+
+definition [consuming]: "tk_int \<equiv> token lx_int"
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Parsing/Munta_Error_Monad_Add.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Parsing/Munta_Error_Monad_Add.thy
@@ -0,0 +1,56 @@
+(* Author: Peter Lammich *)
+section \<open>More Material on the Error Monad\<close>
+
+theory Munta_Error_Monad_Add
+imports
+  Certification_Monads.Check_Monad
+  Show.Show_Instances
+begin
+(* TODO: Move *)
+abbreviation "assert_opt \<Phi> \<equiv> if \<Phi> then Some () else None"
+
+definition "lift_opt m e \<equiv> case m of Some x \<Rightarrow> Error_Monad.return x | None \<Rightarrow> Error_Monad.error e"
+
+lemma lift_opt_simps[simp]:
+  "lift_opt None e = error e"
+  "lift_opt (Some v) e = return v"
+  by (auto simp: lift_opt_def)
+
+(* TODO: Move *)
+lemma reflcl_image_iff[simp]: "R=``S = S\<union>R``S" by blast
+
+named_theorems return_iff
+
+lemma bind_return_iff[return_iff]: "Error_Monad.bind m f = Inr y \<longleftrightarrow> (\<exists>x. m = Inr x \<and> f x = Inr y)"
+  by auto
+
+lemma lift_opt_return_iff[return_iff]: "lift_opt m e = Inr x \<longleftrightarrow> m=Some x"
+  unfolding lift_opt_def by (auto split: option.split)
+
+lemma check_return_iff[return_iff]: "check \<Phi> e = Inr uu \<longleftrightarrow> \<Phi>"
+  by (auto simp: check_def)
+
+
+lemma check_simps[simp]:
+  "check True e = succeed"
+  "check False e = error e"
+  by (auto simp: check_def)
+
+lemma Let_return_iff[return_iff]: "(let x=v in f x) = Inr w \<longleftrightarrow> f v = Inr w" by simp
+
+
+abbreviation ERR :: "shows \<Rightarrow> (unit \<Rightarrow> shows)" where "ERR s \<equiv> \<lambda>_. s"
+abbreviation ERRS :: "string \<Rightarrow> (unit \<Rightarrow> shows)" where "ERRS s \<equiv> ERR (shows s)"
+
+
+lemma error_monad_bind_split:
+"P (bind m f) \<longleftrightarrow> (\<forall>v. m = Inl v \<longrightarrow> P (Inl v)) \<and> (\<forall>v. m = Inr v \<longrightarrow> P (f v))"
+by (cases m) auto
+
+lemma error_monad_bind_split_asm:
+"P (bind m f) \<longleftrightarrow> \<not> (\<exists>x. m = Inl x \<and> \<not> P (Inl x) \<or> (\<exists>x. m = Inr x \<and> \<not> P (f x)))"
+  by (cases m) auto
+
+lemmas error_monad_bind_splits =error_monad_bind_split error_monad_bind_split_asm
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Parsing/Parser_Combinator.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Parsing/Parser_Combinator.thy
@@ -0,0 +1,546 @@
+(* Author: Peter Lammich *)
+section \<open>Parser Combinators\<close>
+
+theory Parser_Combinator
+imports
+  "HOL-Library.Monad_Syntax"
+  "HOL-Library.Char_ord"
+  "HOL-Library.Code_Target_Nat"
+  "Certification_Monads.Error_Monad"
+  Munta_Error_Monad_Add 
+  "Show.Show"
+  "HOL-Library.Rewrite"
+begin
+
+  (**
+    Parser Combinators, based on Sternagel et Thiemann's Parser_Monad, with the following additional features/differences:
+
+    * Setup for the function package, to handle recursion
+      parser uses fuel to ensure termination.
+    * Uses (unit \<Rightarrow> shows) for error messages instead of String (lazy computation of messages, more comfortable due to shows)
+    * Everything defined over generic token type
+    * Some fancy combinators
+        a \<parallel> b    choice, type a = type b
+        --[f]    sequential composition, combine results with f
+        --       = --[Pair]
+        --*      seq, ignore right result
+        *--      seq, ignore left result
+
+    TODO/FIXME
+
+    * Currently, the bind and repeat operation dynamically check whether input is consumed and then fail.
+      At least for bind (no input is generated), we could try to encode this information into the parser's type.
+      However, interplay with function package is not clear at this point :(
+      Possible solution: Fixed-point based recursion combinator and partial_function. We could then do totality proof afterwards.
+
+
+  *)
+
+
+subsection \<open>Type Definitions\<close>
+
+datatype 'a len_list = LL (ll_fuel: nat) (ll_list: "'a list")
+definition "ll_from_list l \<equiv> LL (length l) l"
+
+lemma [measure_function]: "is_measure ll_fuel" ..
+
+text \<open>
+  A parser takes a list of tokes and returns either an error message or
+  a result together with the remaining tokens.
+\<close>
+
+type_synonym
+  ('t, 'a) parser = "'t len_list \<Rightarrow> (unit \<Rightarrow> shows) + ('a \<times> 't len_list)"
+
+text \<open>A \emph{consuming parser} (cparser for short) consumes at least one token of input.\<close>
+definition is_cparser :: "('t,'a) parser \<Rightarrow> bool"
+where
+  "is_cparser p \<longleftrightarrow> (\<forall>l l' x. p l = Inr (x,l') \<longrightarrow> ll_fuel l' < ll_fuel l)"
+
+lemma is_cparserI:
+  assumes "\<And>l l' x. p l = Inr (x, l') \<Longrightarrow> ll_fuel l' < ll_fuel l"
+  shows "is_cparser p"
+  using assms unfolding is_cparser_def by blast
+
+lemma is_cparserE:
+  assumes "is_cparser p"
+    and "(\<And>l l' x. p l = Inr (x, l') \<Longrightarrow> ll_fuel l' < ll_fuel l) \<Longrightarrow> P"
+  shows "P"
+  using assms by (auto simp: is_cparser_def)
+
+lemma is_cparser_length[simp, dest]:
+  assumes "p l = Inr (x, l')" and "is_cparser p"
+  shows "ll_fuel l' < ll_fuel l"
+  using assms by (blast elim: is_cparserE)
+
+text \<open>Used by fundef congruence rules\<close>
+definition "PCONG_INR x ts' \<equiv> Inr (x,ts')"
+
+lemma PCONG_EQ_D[dest!]:
+  assumes "p l = PCONG_INR x l'"
+  assumes "is_cparser p"
+  shows "ll_fuel l' < ll_fuel l"
+  using assms unfolding PCONG_INR_def by auto
+
+named_theorems parser_rules
+
+lemmas parser0_rules = disjI1 disjI2 allI impI conjI
+
+ML \<open>
+  structure Parser_Combinator = struct
+
+    val cfg_simproc = Attrib.setup_config_bool @{binding parser_simproc} (K true)
+
+    val cfg_debug = Attrib.setup_config_bool @{binding parser_debug} (K false)
+
+    fun trace_tac ctxt msg = if Config.get ctxt cfg_debug then print_tac ctxt msg else all_tac
+    fun trace_tac' ctxt msg i =
+      if Config.get ctxt cfg_debug then
+        print_tac ctxt (msg ^ " on subgoal " ^ Int.toString i)
+      else all_tac
+
+    fun prove_cparser_step_tac ctxt =
+      let
+        val p_rls = Named_Theorems.get ctxt @{named_theorems parser_rules}
+      in
+        trace_tac' ctxt "prove_cparser_step" THEN'
+        (
+          resolve_tac ctxt (@{thms parser0_rules} @ p_rls)
+          ORELSE' SOLVED' (asm_simp_tac ctxt)
+        )
+      end
+
+
+    fun prove_cparser_tac ctxt =
+      trace_tac ctxt "prove_cparser" THEN
+      DEPTH_SOLVE (FIRSTGOAL (prove_cparser_step_tac ctxt))
+
+    fun add_cparser_def def_thm context = let
+      val ctxt = Context.proof_of context
+      val orig_ctxt = ctxt
+
+      val ctxt = Config.put cfg_simproc false ctxt
+
+      val (def_thm, ctxt) = yield_singleton (apfst snd oo Variable.import true) def_thm ctxt
+
+      val lhs = def_thm
+        |> Local_Defs.meta_rewrite_rule ctxt
+        |> (fst o Logic.dest_equals o Thm.prop_of)
+
+      val T = fastype_of lhs
+      val cp_stmt =
+           Const (@{const_name is_cparser}, T --> HOLogic.boolT)$lhs
+        |> HOLogic.mk_Trueprop
+
+      fun is_cparser_free (@{const Trueprop} $ (Const (@{const_name is_cparser},_) $ Free _)) = true
+        | is_cparser_free _ = false
+
+      fun is_goal_ok st =
+        Thm.prop_of st |> Logic.strip_imp_prems
+        |> map (Logic.strip_assums_concl)
+        |> forall is_cparser_free
+
+      val cp_thm =
+        cp_stmt |> Thm.cterm_of ctxt
+        |> Goal.init
+        |> SINGLE (
+            unfold_tac ctxt [def_thm] THEN
+            trace_tac ctxt "cparser def proof" THEN
+            DEPTH_FIRST is_goal_ok (
+              FIRSTGOAL (prove_cparser_step_tac ctxt)
+            )
+          )
+
+      val cp_thm = case cp_thm of
+        NONE => error "Could not prove any is_cparser theorem: Empty result sequence"
+      | SOME thm =>
+          Goal.conclude thm
+
+
+      val cp_thm =
+           singleton (Variable.export ctxt orig_ctxt) cp_thm
+        |> Drule.zero_var_indexes
+
+      (*
+      val cp_thm =
+        Goal.prove ctxt [] [] cp_stmt (fn {context, ...} => tac context)
+      |> singleton (Variable.export ctxt orig_ctxt)
+      *)
+
+      val context = Named_Theorems.add_thm @{named_theorems parser_rules} cp_thm context
+    in
+      context
+    end
+  end
+\<close>
+
+
+attribute_setup consuming = \<open>
+  Scan.succeed (Thm.declaration_attribute (Parser_Combinator.add_cparser_def))
+\<close>
+
+simproc_setup is_cparser_prover ("is_cparser p") = \<open>fn _ => fn ctxt => fn ct =>
+  if Config.get ctxt Parser_Combinator.cfg_simproc then
+    let
+      open Parser_Combinator
+      val t = Thm.term_of ct
+      val stmt = Logic.mk_equals (t,@{term True})
+
+      val _ = if Config.get ctxt cfg_debug then
+          (Pretty.block [Pretty.str "is_cparser simproc invoked on: ", Syntax.pretty_term ctxt t, Pretty.fbrk, Syntax.pretty_term ctxt stmt]) |> Pretty.string_of |> tracing
+        else ()
+
+      val ctxt = Config.put Parser_Combinator.cfg_simproc false ctxt
+
+      val othm = try (Goal.prove ctxt [] [] stmt) (fn {context=ctxt, ...} =>
+        FIRSTGOAL (resolve_tac ctxt @{thms HOL.Eq_TrueI})
+        THEN TRY (Parser_Combinator.prove_cparser_tac ctxt)
+      )
+
+      val _ =
+        if Config.get ctxt cfg_debug andalso is_none othm then
+          (Pretty.block [Pretty.str "is_cparser simproc failed on: ", Syntax.pretty_term ctxt t, Pretty.fbrk, Syntax.pretty_term ctxt stmt]) |> Pretty.string_of |> tracing
+        else ()
+
+      (*
+      val _ = case othm of
+        NONE => (Pretty.block [Pretty.str "is_cparser simproc failed on: ", Syntax.pretty_term ctxt t, Pretty.fbrk, Syntax.pretty_term ctxt stmt]) |> Pretty.string_of |> warning
+     | SOME _ => ();
+      *)
+    in
+      othm
+    end
+  else
+    NONE
+\<close>
+
+text \<open>Wrapping a parser to dynamically assert that it consumes tokens.\<close>
+definition ensure_cparser :: "('t,'a) parser \<Rightarrow> ('t,'a) parser" where
+  "ensure_cparser p \<equiv> \<lambda>ts. do {
+    (x, ts') \<leftarrow> p ts;
+    if (ll_fuel ts' < ll_fuel ts) then Error_Monad.return (x,ts')
+    else Error_Monad.error (\<lambda>_. shows ''Dynamic parser check failed'')
+  }"
+
+lemma ensure_cparser_cparser[parser_rules]: "is_cparser (ensure_cparser p)"
+  apply (rule is_cparserI)
+  unfolding ensure_cparser_def
+  by (auto simp: Error_Monad.bind_def split: sum.splits if_splits)
+
+lemma ensure_cparser_cong[fundef_cong]:
+  assumes "l=l'"
+  assumes "p l = p' l'"
+  shows "ensure_cparser p l = ensure_cparser p' l'"
+  using assms by (auto simp: ensure_cparser_def)
+
+
+abbreviation bnf_eq :: "('t,'a) parser \<Rightarrow> ('t,'a) parser \<Rightarrow> prop" (infix "::=" 2) where
+  "bnf_eq p1 p2 \<equiv> (\<And>l. p1 l = p2 l)"
+
+
+
+subsection \<open>Monad-Setup for Parsers\<close>
+
+definition return :: "'a \<Rightarrow> ('t, 'a) parser"
+where
+  "return x = (\<lambda>ts. Error_Monad.return (x, ts))"
+
+definition error_aux :: "(unit \<Rightarrow> shows) \<Rightarrow> ('t, 'a) parser"
+where
+  "error_aux e = (\<lambda>_. Error_Monad.error e)"
+
+abbreviation error where "error s \<equiv> error_aux (ERR s)"
+abbreviation "error_str s \<equiv> error_aux (ERRS s)"
+
+definition update_error :: "('t,'a) parser \<Rightarrow> (shows \<Rightarrow> shows) \<Rightarrow> ('t,'a) parser"
+  where "update_error m f l \<equiv> m l <+? (\<lambda>e _. f (e ()))"
+
+definition ensure_parser :: "('t,'a) parser \<Rightarrow> ('t,'a) parser" where
+  "ensure_parser p \<equiv> \<lambda>ts. do {
+    (x, ts') \<leftarrow> p ts;
+    if (ll_fuel ts' \<le> ll_fuel ts) then Error_Monad.return (x,ts')
+    else Error_Monad.error (ERRS ''Dynamic parser check failed'')
+  }"
+
+definition bind :: "('t, 'a) parser \<Rightarrow> ('a \<Rightarrow> ('t, 'b) parser) \<Rightarrow> ('t, 'b) parser"
+where
+  "bind m f \<equiv> \<lambda>ts. do {
+    (x, ts) \<leftarrow> ensure_parser m ts;
+    ensure_parser (f x) ts
+  }"
+
+definition get :: "('t,'t) parser" where
+  "get ll \<equiv> case ll of LL (Suc n) (x#xs) \<Rightarrow> Error_Monad.return (x,LL n xs) | _ \<Rightarrow> (Error_Monad.error (\<lambda>_. shows_string ''Expecting more input''))"
+
+definition get_tokens :: "('t,'t list) parser" where
+  "get_tokens \<equiv> \<lambda>ll. Error_Monad.return (ll_list ll,ll)"
+
+adhoc_overloading
+      Monad_Syntax.bind \<rightleftharpoons> bind
+  and Error_Syntax.update_error \<rightleftharpoons> update_error
+
+(* TODO: Specialize to parser type? *)
+lemma let_cong' [fundef_cong]:
+  "M = N \<Longrightarrow> l=l' \<Longrightarrow> (\<And>x. x = N \<Longrightarrow> f x l' = g x l') \<Longrightarrow> Let M f l = Let N g l'"
+  unfolding Let_def by blast
+
+lemma if_cong' [fundef_cong]:
+  assumes "b = c"
+    and "l=l'"
+    and "c \<Longrightarrow> x l' = u l'"
+    and "\<not> c \<Longrightarrow> y l' = v l'"
+  shows "(if b then x else y) l = (if c then u else v) l'"
+  using assms by simp
+
+lemma split_cong' [fundef_cong]:
+  "l=l' \<Longrightarrow> (\<And>x y. (x, y) = q \<Longrightarrow> f x y l' = g x y l' ) \<Longrightarrow> p = q \<Longrightarrow> case_prod f p l = case_prod g q l'"
+  by (auto simp: split_def)
+
+lemma bind_cong [fundef_cong]:
+  fixes m1 :: "('t, 'a) parser"
+  assumes "m1 ts2 = m2 ts2"
+    and "\<And> y ts. \<lbrakk> m2 ts2 = PCONG_INR y ts; ll_fuel ts \<le> ll_fuel ts2\<rbrakk> \<Longrightarrow> f1 y ts = f2 y ts"
+    and "ts1 = ts2"
+  shows "((m1 \<bind> f1) ts1) = ((m2 \<bind> f2) ts2)"
+  using assms
+  unfolding bind_def PCONG_INR_def
+  by (auto simp: Error_Monad.bind_def ensure_parser_def split: sum.split prod.split)
+
+lemma is_cparser_bind[parser_rules]:
+  assumes "is_cparser p \<or> (\<forall>x. is_cparser (q x))"
+  shows "is_cparser (p \<bind> q)"
+  apply (rule is_cparserI)
+  using assms unfolding is_cparser_def bind_def Error_Monad.bind_def ensure_parser_def
+  by (fastforce split: sum.splits if_splits)
+
+lemma return_eq[simp]: "return x l = Inr y \<longleftrightarrow> y=(x,l)"
+  unfolding return_def by auto
+
+
+lemma is_cparser_error[parser_rules]: "is_cparser (error_aux e)"
+  by (auto simp: error_aux_def intro: is_cparserI)
+
+lemma is_cparser_get[parser_rules]:
+  "is_cparser get"
+  apply (rule is_cparserI)
+  apply (auto simp: get_def split: len_list.splits nat.splits list.splits)
+  done
+
+lemma monad_laws[simp]:
+  "bind m return = ensure_parser m"
+  "bind (return x) f = ensure_parser (f x)"
+  "bind (bind m f) g = bind m (\<lambda>x. bind (f x) g)"
+  "ensure_parser (ensure_parser m) = ensure_parser m"
+  "bind (ensure_parser m) f = bind m f"
+  "bind m (\<lambda>x. ensure_parser (f x)) = bind m f"
+  unfolding bind_def return_def ensure_parser_def Error_Monad.bind_def
+  by (auto split: if_splits sum.splits prod.splits intro!: ext)
+
+subsection \<open>More Combinators\<close>
+
+definition err_expecting_aux :: "(unit \<Rightarrow> shows) \<Rightarrow> ('t::show, 'a) parser"
+where
+  "err_expecting_aux msg = do { ts \<leftarrow> get_tokens; error
+    (shows_string ''expecting '' o msg () o shows_string '', but found: '' o shows_quote (shows (take 100 ts)))}"
+
+abbreviation err_expecting :: "shows \<Rightarrow> ('t::show, 'a) parser"
+where
+  "err_expecting msg \<equiv> err_expecting_aux (\<lambda>_. msg)"
+
+abbreviation "err_expecting_str msg \<equiv> err_expecting (shows_string msg)"
+
+definition "eoi \<equiv> do {
+  tks \<leftarrow> get_tokens; if tks=[] then return () else err_expecting_str ''end of input'' }"
+
+
+definition alt :: "(_,'a) parser \<Rightarrow> (_,'b) parser \<Rightarrow> (_,'a+'b) parser" where
+  "alt p1 p2 l \<equiv> try   do { (r,l) \<leftarrow> p1 l; Error_Monad.return (Inl r, l) }
+                 catch (\<lambda>e1. (try do { (r,l) \<leftarrow> p2 l; Error_Monad.return (Inr r, l) }
+                 catch (\<lambda>e2. Error_Monad.error (\<lambda>_. e1 () o shows ''\<newline>  | '' o e2 ()))))"
+
+fun sum_join where
+  "sum_join (Inl x) = x" | "sum_join (Inr x) = x"
+
+abbreviation alt' :: "(_,'a) parser \<Rightarrow> (_,'a) parser \<Rightarrow> (_,'a) parser" (infixr "\<parallel>" 53)
+  where "alt' p q \<equiv> alt p q \<bind> return o sum_join"
+
+abbreviation gseq :: "('t,'a) parser \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('t,'b) parser \<Rightarrow> ('t,'c) parser" ("_--[_]_" [61,0,60] 60)
+  where "gseq p f q \<equiv> p \<bind> (\<lambda>a. q \<bind> (\<lambda>b. return (f a b)))" (* TODO/FIXME: Do-notation and abbreviation generate additional type vars here *)
+
+abbreviation seq :: "('t,'a) parser \<Rightarrow> ('t,'b) parser \<Rightarrow> ('t,'a\<times>'b) parser" (infixr "--" 60)
+  where "seq p q \<equiv> p --[Pair] q"
+
+abbreviation seq_ignore_left :: "('t,'a) parser \<Rightarrow> ('t,'b) parser \<Rightarrow> ('t,'b) parser" (infixr "*--" 60)
+  where "p *-- q \<equiv> p --[\<lambda>_ x. x] q"
+
+abbreviation seq_ignore_right :: "('t,'a) parser \<Rightarrow> ('t,'b) parser \<Rightarrow> ('t,'a) parser" (infixr "--*" 60)
+  where "p --* q \<equiv> p --[\<lambda>x _. x] q"
+
+abbreviation map_result :: "('t,'a) parser \<Rightarrow> ('a\<Rightarrow>'b) \<Rightarrow> ('t,'b) parser" (infixr "with" 54)
+  where "p with f \<equiv> p \<bind> return o f"
+
+definition "exactly ts \<equiv> foldr (\<lambda>t p. do { x\<leftarrow>get; if x=t then p \<bind> return o (#) x else error id}) ts (return [])
+    \<parallel> err_expecting (shows_string ''Exactly '' o shows ts)"
+
+declare err_expecting_aux_def[consuming]
+
+lemma alt_is_cparser[parser_rules]:
+  "is_cparser p \<Longrightarrow> is_cparser q \<Longrightarrow> is_cparser (alt p q)"
+  apply (rule is_cparserI)
+  unfolding alt_def
+  by (auto simp: Error_Monad.bind_def split: sum.splits)
+
+lemma alt_cong[fundef_cong]:
+  "\<lbrakk> l=l'; p1 l = p1' l'; \<And>e. p1' l' = Inl e \<Longrightarrow> p2 l = p2' l' \<rbrakk> \<Longrightarrow> alt p1 p2 l = alt p1' p2' l'"
+  unfolding alt_def by (auto split: sum.splits simp: Error_Monad.bind_def)
+
+lemma [parser_rules]: "ts\<noteq>[] \<Longrightarrow> is_cparser (exactly ts)"
+  by (cases ts) (auto simp: exactly_def intro: parser_rules)
+
+
+abbreviation optional :: "'a \<Rightarrow> ('t,'a) parser \<Rightarrow> ('t,'a) parser" where
+  "optional dflt p \<equiv> p \<parallel> return dflt"
+
+abbreviation maybe :: "('t,'a) parser \<Rightarrow> ('t,'a option) parser" ("(_?)" [1000] 999)
+  where "p? \<equiv> p with Some \<parallel> return None"
+
+
+subsubsection \<open>Repeat\<close>
+
+fun repeat :: "('t,'a) parser \<Rightarrow> ('t,'a list) parser" where
+  "repeat p ::= optional [] (ensure_cparser p --[(#)] repeat p)"
+
+abbreviation "repeat1 p \<equiv> p --[(#)] repeat p"
+
+
+declare repeat.simps[simp del]
+
+lemma repeat_cong[fundef_cong]:
+  assumes "\<And>nts. \<lbrakk> ll_fuel nts \<le> ll_fuel l' \<rbrakk> \<Longrightarrow> p (nts) = p' (nts)"
+  assumes "l=l'"
+  shows "repeat p l = repeat p' l'"
+  using assms(1)
+  unfolding \<open>l=l'\<close>
+  apply (induction p l' rule: repeat.induct)
+  apply (rewrite in "_=\<hole>" repeat.simps)
+  apply (rewrite in "\<hole>=_" repeat.simps)
+  apply (intro alt_cong bind_cong)
+  apply (auto simp: ensure_cparser_def PCONG_INR_def)
+  done
+
+subsubsection \<open>Left and Right Associative Chaining\<close>
+text \<open>Parse a sequence of \<open>A\<close> separated by \<open>F\<close>,
+  and then fold the sequence with the results of \<open>F\<close>,
+  either left or right associative.
+
+  Example: Assume we have the input \<open>x1 o1 \<dots> on-1 xn\<close>,
+    and the result of parsing the \<open>x\<close> with \<open>A\<close> and the \<open>o\<close> with \<open>F\<close>
+    are \<open>ai\<close> and \<open>+i\<close>.
+    Then, \<open>chainL1\<close> returns \<open>(\<dots>((a1 +1 a2) +2 a3) +3 \<dots>) \<close>
+    and \<open>chainR1\<close> returns \<open>a1 +1 (a2 +2 (a3 +3 \<dots>)\<dots>) \<close>
+\<close>
+context
+  fixes A :: "('t,'a) parser"
+  fixes F :: "('t,'a\<Rightarrow>'a\<Rightarrow>'a) parser"
+begin
+  definition chainL1 :: "('t,'a) parser" where
+    "chainL1 \<equiv> do {
+      x \<leftarrow> A;
+      xs \<leftarrow> repeat (F --[\<lambda>f b a. f a b] A);
+      return (foldl (\<lambda>a f. f a) x xs)
+    }"
+
+  qualified fun fold_shiftr :: "('a \<Rightarrow> (('a\<Rightarrow>'a\<Rightarrow>'a)\<times>'a) list \<Rightarrow> 'a)" where
+    "fold_shiftr a [] = a"
+  | "fold_shiftr a ((f,b)#xs) = f a (fold_shiftr b xs)"
+
+  definition chainR1 :: "('t,'a) parser" where
+    "chainR1 \<equiv> do {
+      x \<leftarrow> A;
+      xs \<leftarrow> repeat (F -- A);
+      return (fold_shiftr x xs)
+    } "
+
+end
+
+
+lemma chainL1_cong[fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> A l2 = A' l2"
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> F l2 = F' l2"
+  assumes "l=l'"
+  shows "chainL1 A F l = chainL1 A' F' l'"
+  unfolding chainL1_def
+  apply (intro bind_cong repeat_cong assms order_refl)
+  by auto
+
+lemma chainR1_cong[fundef_cong]:
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> A l2 = A' l2"
+  assumes "\<And>l2. ll_fuel l2 \<le> ll_fuel l' \<Longrightarrow> F l2 = F' l2"
+  assumes "l=l'"
+  shows "chainR1 A F l = chainR1 A' F' l'"
+  unfolding chainR1_def
+  apply (intro bind_cong repeat_cong assms order_refl)
+  by auto
+
+
+
+
+subsection \<open>Lexing Utilities\<close>
+
+definition tk_with_prop' :: "shows \<Rightarrow> ('t::show \<Rightarrow> bool) \<Rightarrow> ('t,'t) parser" where
+  [consuming]: "tk_with_prop' errmsg \<Phi> \<equiv> do {
+    x\<leftarrow>get;
+    if \<Phi> x then return x
+    else err_expecting errmsg
+  }"
+
+abbreviation tk_with_prop :: "('t::show \<Rightarrow> bool) \<Rightarrow> ('t,'t) parser" where
+  "tk_with_prop \<equiv> tk_with_prop' id"
+
+definition range :: "'t::{linorder,show} \<Rightarrow> 't \<Rightarrow> ('t,'t) parser" where
+  [consuming]: "range a b \<equiv> do {
+    x\<leftarrow>get;
+    if a\<le>x \<and> x\<le>b then return x
+    else err_expecting (shows_string ''Token in range '' o shows a o shows_string '' - '' o shows b) }"
+
+definition any :: "'t::show list \<Rightarrow> ('t,'t) parser" where
+  [consuming]: "any ts \<equiv> do { t\<leftarrow>get; if t\<in>set ts then return t else err_expecting (shows_string ''One of '' o shows ts) }"
+
+definition "gen_token ws p \<equiv> ws *-- p"
+
+lemma [parser_rules]: "is_cparser p \<Longrightarrow> is_cparser (gen_token ws p)"
+  unfolding gen_token_def by simp
+
+subsubsection \<open>Characters\<close>
+abbreviation (input) "char_tab \<equiv> CHR 0x09"
+abbreviation (input) "char_carriage_return \<equiv> CHR 0x0D"
+abbreviation (input) "char_wspace \<equiv> [CHR '' '', CHR ''\<newline>'', char_tab, char_carriage_return]"
+
+text \<open>Some standard idioms\<close>
+definition [consuming]: "lx_lowercase \<equiv> (range CHR ''a'' CHR ''z'' )"
+definition [consuming]: "lx_uppercase \<equiv> (range CHR ''A'' CHR ''Z'' )"
+definition [consuming]: "lx_alpha \<equiv> (lx_lowercase \<parallel> lx_uppercase)"
+definition [consuming]: "lx_digit \<equiv> (range CHR ''0'' CHR ''9'' )"
+abbreviation "lx_alphanum \<equiv> lx_alpha \<parallel> lx_digit"
+
+subsection \<open>Code Generator Setup\<close>
+declare monad_laws[code_unfold]
+lemma bind_return_o_unfold[code_unfold]: "(m \<bind> return o f) = do { x\<leftarrow>m; return (f x)}" by (auto simp: o_def)
+declare split[code_unfold] (* TODO: Should this be code_unfold by default? *)
+
+
+subsection \<open>Utilities for Parsing\<close>
+
+text \<open>Project out remainder token sequence\<close>
+fun show_pres where
+  "show_pres (Inr (ll,_)) = Inr ll"
+| "show_pres (Inl e) = Inl e"
+
+text \<open>Parse complete input, parameterized by parser for trailing whitespace\<close>
+definition "parse_all ws p \<equiv> show_pres o (p --* ws --* eoi) o ll_from_list o String.explode"
+
+definition "parse_all_implode ws p s \<equiv> parse_all ws p s <+? (\<lambda>msg. String.implode (msg () ''''))"
+
+definition "parse_all_implode_nows p s \<equiv> parse_all_implode (return ()) p s"
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/ROOT
--- /dev/null
+++ thys/Munta_Model_Checker/ROOT
@@ -0,0 +1,71 @@
+chapter AFP
+
+session Munta_Base in TA_Impl = Timed_Automata +
+  options [timeout = 600]
+  sessions
+    Show
+    Worklist_Algorithms
+    Transition_Systems_and_Automata
+    LTL_Master_Theorem
+    Certification_Monads
+    Gabow_SCC
+    AutoCorres2
+    FinFun
+  directories
+    "../library"
+  theories [document = false]
+    Refine_Imperative_HOL.IICF
+    "HOL-Library.IArray"
+  theories
+    "../library/TA_Syntax_Bundles"
+    "../library/ML_Util"
+    "../library/More_Methods"
+    "../library/Abstract_Term"
+    "../library/Reordering_Quantifiers"
+    "../library/Bijective_Embedding"
+    "../library/Tracing"
+    "../library/Printing"
+    "../library/Trace_Timing"
+    "../library/Error_List_Monad"
+    "../library/Temporal_Logics"
+    "../library/Subsumption_Graphs"
+
+    Normalized_Zone_Semantics_Impl_Refine
+    Normalized_Zone_Semantics_Impl_Extra
+
+    Certification_Monads.Error_Syntax
+    Gabow_SCC.Gabow_SCC_Code
+    AutoCorres2.Subgoals
+    FinFun.FinFun
+
+session Munta_Model_Checker = Munta_Base +
+  options [timeout = 1800]
+  sessions
+    "HOL-Library"
+    Show
+  directories
+    Uppaal_Networks
+    Simple_Networks
+    Deadlock
+    Model_Checking
+    Parsing
+    Networks
+    Munta_MLton
+  theories [document=false]
+    "Simple_Networks/Munta_Tagging"
+    "Parsing/JSON_Parsing"
+  theories
+    "Networks/Networks"
+    "Networks/State_Networks"
+    "Uppaal_Networks/UPPAAL_Model_Checking"
+    "Simple_Networks/Simple_Network_Language_Model_Checking"
+    "Deadlock/Deadlock_Impl"
+    "Deadlock/Deadlock_Checking"
+    "Model_Checking/Simple_Network_Language_Export_Code"
+  theories [condition = ISABELLE_MLTON]
+    "Munta_MLton/Munta_Compile_MLton"
+  document_files
+    "root.tex"
+    "root.bib"
+  export_files (in ".") [1]
+    "Munta_Model_Checker.Munta_Compile_MLton:*"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Munta_Tagging.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Munta_Tagging.thy
@@ -0,0 +1,590 @@
+section \<open>Tagged facts\<close>
+theory Munta_Tagging
+  imports Main "HOL-Eisbach.Eisbach"
+  keywords "usingT"  :: prf_decl % "proof"
+       and "preferT" :: prf_script % "proof"
+       and "print_tags" :: diag
+begin
+
+subsection \<open>Tags\<close>
+
+definition TAG ("_ \<bar> _" [8, 8] 9) where
+  \<open>TAG t x = x\<close>
+
+definition
+  "TAG' t x = TAG t x"
+
+lemmas TAG = TAG'_def[symmetric]
+
+lemmas add_tag = TAG_def[symmetric]
+
+named_theorems tagged
+
+
+subsection \<open>Prototyping Our Methods With Eisbach\<close>
+
+lemma rm_TAG:
+  assumes "TAG t x" "TAG t x = TAG' t x"
+  shows True
+  by auto
+
+method tags0 uses del keep =
+  (unfold keep)?,
+  ((drule rm_TAG, rule del)+)?,
+  (unfold keep[symmetric])?
+
+lemma
+  assumes "TAG 1 False" "TAG (Some STR ''remove'') (x > 0)"
+  shows False
+  using assms
+  apply -
+  apply (tags0 del: TAG[of "Some t" for t] TAG[of 1] keep: TAG[of 1])
+  unfolding TAG_def .
+
+ML \<open>
+fun unfold_tac' ctxt rewrs =
+  let
+    val rewrs = map (Local_Defs.abs_def_rule ctxt) rewrs
+  in
+    rewrite_goal_tac ctxt rewrs
+  end
+
+fun insert'_meth ths ctxt = SIMPLE_METHOD' (Method.insert_tac ctxt ths)
+
+fun unfold'_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (unfold_tac' ctxt ths 1))
+\<close>
+
+method_setup insert' = \<open>Attrib.thms >> insert'_meth\<close> "insert theorems into first subgoal"
+
+method_setup unfold' = \<open>Attrib.thms >> unfold'_meth\<close> "unfold definitions in first subgoal"
+
+method tags' uses keep declares tagged =
+  use nothing in \<open>insert' method_facts tagged\<close>,
+  use nothing in \<open>tags0 del: TAG keep: keep, unfold' TAG_def, ((thin_tac True)+)?\<close>
+
+lemma
+  assumes "TAG ''tagged'' (x > 0)"
+  shows "x > 0"
+  using assms
+  apply (tags' keep: TAG[of "''tagged''"] tagged: assms)
+  apply assumption
+  done
+
+method tags uses keep declares tagged =
+   match conclusion in "tag \<bar> P" for tag P \<Rightarrow>
+     \<open>tags' keep: keep TAG[of tag] tagged: tagged\<close>
+ | tags' keep: keep tagged: tagged
+
+lemma
+  assumes "''tagged'' \<bar> (x > 0)" and [tagged]: "''test'' \<bar> P" and "''test'' \<bar> Q" 
+  shows "''tagged'' \<bar> x > 0"
+  using assms(1)
+  apply (tags keep: TAG[of "''test''"] tagged: assms(3))
+  apply assumption
+  done
+
+method etag for tag uses keep declares tagged =
+  tags keep: keep TAG[of tag] tagged: tagged
+
+lemma
+  assumes "''tagged'' \<bar> (x > 0)" and [tagged]: "''test'' \<bar> P" and "''test'' \<bar> Q" 
+  shows "''tagged'' \<bar> x > 0"
+  using assms(1)
+  apply (etag "''test''" tagged: assms(3))
+  apply assumption
+  done
+
+
+subsection \<open>Recalling Facts by Tag\<close>
+
+lemmas add_TAG = TAG_def[THEN eq_reflection, symmetric]
+
+lemma re_TAG:
+  "(t \<bar> P) \<equiv> (t_new \<bar> P)"
+  unfolding TAG_def .
+
+ML \<open>
+fun mk_add_tag_thm t =
+  instantiate\<open>'t = \<open>Thm.ctyp_of_cterm t\<close> and t = t in
+    lemma (schematic) \<open>y \<equiv> t \<bar> y\<close> for t :: 't by (rule add_TAG)\<close>
+fun mk_re_tag_thm t =
+  instantiate\<open>'t = \<open>Thm.ctyp_of_cterm t\<close> and t_new = t in
+    lemma (schematic) \<open>(t_old \<bar> P) \<equiv> (t_new \<bar> P)\<close> for t_new :: 't by (rule re_TAG)\<close>
+fun lift_conv conv =
+  conv |> Conv.rewr_conv |> Conv.arg_conv |> Conv.concl_conv 100 |> K |> Conv.params_conv 100
+fun add_tag_conv ct = mk_add_tag_thm ct |> lift_conv
+fun re_tag_conv ct  = mk_re_tag_thm ct |> lift_conv
+fun tag_thm ct ctxt = Conv.fconv_rule (add_tag_conv ct ctxt)
+fun re_tag_thm ct ctxt = Conv.fconv_rule (Conv.try_conv (re_tag_conv ct ctxt))
+
+fun extract_tag_trm trm =
+  case trm of
+    Const\<open>Trueprop for Const\<open>TAG _ _ for tag _\<close>\<close> => SOME tag
+  | _ => NONE
+
+val empty = (Vartab.empty, Vartab.empty)
+
+fun matches thy pat trm = (Pattern.match thy (pat, trm) empty; true)
+  handle Pattern.MATCH => false
+
+fun matches_tag0 thy tag =
+  matches thy (instantiate\<open>tag in term (schematic) \<open>Trueprop (tag \<bar> P)\<close>\<close>)
+
+fun matches_tag thy tag goal =
+  case extract_tag_trm goal of
+      NONE => false
+    | SOME t => matches thy tag t
+
+fun filter_thms thy pat = List.filter (fn thm => matches thy pat (Thm.prop_of thm))
+
+fun filter_tagged thy tag =
+  List.filter (fn thm => matches_tag thy tag (Thm.prop_of thm))
+
+val tagged = Named_Theorems.check @{context} ("tagged", Position.none)
+
+fun get_tagged ctxt tag =
+  filter_tagged (Proof_Context.theory_of ctxt) tag (Named_Theorems.get ctxt tagged)
+
+fun err ctxt trm pos msg =
+  error (msg ^ Position.here pos ^ ":\n" ^ Syntax.string_of_term ctxt trm);
+
+fun check_single ctxt tag pos thms = case thms of
+  [x] => [x]
+| [] => err ctxt tag pos "No fact with tag"
+| _  => err ctxt tag pos "Ambiguous tag"
+
+fun get_tagged_single ctxt pos tag = get_tagged ctxt tag |> check_single ctxt tag pos
+
+fun check_at_least_one ctxt tag pos thms = case thms of
+  [] => err ctxt tag pos "No fact with tag"
+| xs => xs
+
+fun get_tagged_at_least_one ctxt pos tag = get_tagged ctxt tag |> check_at_least_one ctxt tag pos
+
+fun unfold_tags ctxt = Local_Defs.unfold ctxt @{thms TAG_def}
+
+fun filter_thms_attr ((opt_num, keep_tags), s) = Thm.rule_attribute [] (fn context => fn _ =>
+  let
+    val pos = Syntax.read_input_pos s
+    val ctxt = Context.proof_of context
+    val pat = Proof_Context.read_term_pattern ctxt s
+    val facts = (
+      if opt_num = NONE then
+        get_tagged_single ctxt pos pat
+      else
+        get_tagged_at_least_one ctxt pos pat)
+    val post_fact = if keep_tags then I else unfold_tags ctxt
+  in
+    case opt_num of
+      NONE => List.hd facts |> post_fact
+    | SOME i =>
+      if length facts >= i andalso i > 0 then
+        Library.nth facts (i - 1)
+      else err ctxt pat pos ("Index " ^ string_of_int i ^ " out of range")
+  end)
+
+fun tag_thms_attr ((keep_tags,retag), s) = Thm.mixed_attribute (fn (context, thm) =>
+  let
+    val ctxt = Context.proof_of context
+    val ct = Syntax.read_term ctxt s |> Thm.cterm_of ctxt
+    val thm = if retag then re_tag_thm ct ctxt thm else tag_thm ct ctxt thm
+    val context = Named_Theorems.add_thm tagged thm context
+    val thm = if keep_tags then thm else unfold_tags ctxt thm
+  in
+    (context, thm)
+  end)
+
+val untagged_attr = Thm.rule_attribute [] (fn context => unfold_tags (Context.proof_of context))
+
+fun get_tagged_state strict keep_tags s: Proof.state -> Proof.state = fn st =>
+  let
+    val pos = Syntax.read_input_pos s
+    val ctxt = Proof.context_of st
+    val tag = Proof_Context.read_term_pattern ctxt s
+    val thms = (if strict then get_tagged_single else get_tagged_at_least_one) ctxt pos tag
+    val thms = if keep_tags then thms else map (unfold_tags ctxt) thms
+  in Proof.using [[(thms, [])]] st end
+
+fun get_tagged_trans ((non_strict, keep_tags), xs) =
+  fold (get_tagged_state (not non_strict) keep_tags) xs
+
+val auto_insert_tac = Subgoal.FOCUS (fn {context = ctxt, concl, ...} =>
+  case Thm.term_of concl of
+    Const\<open>Trueprop for Const\<open>TAG _ _ for tag _\<close>\<close> =>
+      Method.insert_tac ctxt (get_tagged ctxt tag) 1
+  | _ => no_tac
+)
+
+fun TRY' tac = tac ORELSE' K all_tac;
+
+fun insert_tagged_meth xs ctxt =
+  let
+    val facts = flat (map (fn (pos, tag) => get_tagged_at_least_one ctxt pos tag) xs)
+    val tac = Method.insert_tac ctxt facts THEN' TRY' (auto_insert_tac ctxt)
+  in
+    SIMPLE_METHOD' tac
+  end
+
+fun untag_tac ctxt = unfold_tac' ctxt @{thms TAG_def}
+\<close>
+
+ML \<open>
+val position_parser: Position.T parser = Scan.ahead Parse.not_eof >> Token.pos_of
+val position_ctxt_parser: Position.T context_parser = Scan.lift position_parser
+val parse_non_strict: bool parser = Scan.optional (Args.$$$ "-" >> K true) false
+\<close>
+
+ML \<open>
+Outer_Syntax.command command_keyword\<open>usingT\<close> "use tagged facts"
+  (parse_non_strict -- Args.mode "keep" -- Scan.repeat1 Parse.embedded_inner_syntax
+    >> (Toplevel.proof o get_tagged_trans))
+\<close>
+
+method_setup insertT =
+  \<open>Scan.repeat (position_ctxt_parser -- Args.term_pattern) >> insert_tagged_meth\<close>
+  "insert tagged facts"
+
+method_setup untag =
+  \<open>Args.context >> (fn _ => fn ctxt => SIMPLE_METHOD' (untag_tac ctxt))\<close> "remove tags"
+
+attribute_setup get_tagged =
+  \<open>Scan.lift (Scan.option (Args.parens Parse.nat) -- Args.mode "keep" -- Parse.embedded_inner_syntax)
+    >> filter_thms_attr\<close>
+
+attribute_setup tag =
+  \<open>Scan.lift (Args.mode "keep" -- Args.mode "re" -- Parse.embedded_inner_syntax) >> tag_thms_attr\<close>
+
+attribute_setup untagged =
+  \<open>Scan.succeed () >> (fn _ => untagged_attr)\<close>
+
+notepad begin
+  fix t1 t2 :: 'a and P Q :: bool
+  assume F: False
+
+  have [tagged]: "t1 \<bar> P" "Some t2 \<bar> Q"
+    using F by auto
+
+  have False
+    using [[get_tagged \<open>t1\<close>]] cancel\<open>[[get_tagged \<open>t2\<close>]]\<close> [[get_tagged \<open>Some _\<close>]]
+    using F by simp
+
+  have False thm tagged
+    usingT cancel\<open>\<open>''hi''\<close>\<close> \<open>Some _\<close> \<open>t1\<close>
+    using F by simp
+
+  have False
+    apply (insertT cancel\<open>\<open>''hi''\<close>\<close> \<open>Some _\<close> \<open>t1\<close>)
+    using F by simp
+
+  have [tagged]: "t1 \<bar> True" unfolding TAG_def ..
+
+  thm [[get_tagged (1) \<open>t1\<close>]]
+
+  \<comment> \<open>Note that these next three are not actually retrievable later.\<close>
+  have test[tag (re) \<open>''tag''\<close>]: "True" "x = x" "tt \<bar> y = y" if "x > 0" for x y tt :: nat
+    unfolding TAG_def by auto
+  thm tagged thm test
+
+  have test[tag \<open>''tag''\<close>]: "True" "x = x" "tt \<bar> y = y" if "x > 0" for x y tt :: nat
+    unfolding TAG_def by auto
+  thm tagged thm test
+
+  have test[tag (keep) \<open>''tag''\<close>]: "True" "x = x" "tt \<bar> y = y" if "x > 0" for x y tt :: nat
+    unfolding TAG_def by auto thm test
+
+  thm tagged[untagged] thm tagged
+
+  have "False \<Longrightarrow> t1 \<bar> True"
+    apply (tactic \<open>auto_insert_tac @{context} 1\<close>)
+    by untag
+
+  have "t1 \<bar> True"
+    apply (insertT \<open>Some _\<close> cancel\<open>\<open>''hi''\<close>\<close> cancel\<open>\<open>t1\<close>\<close>)
+    by untag
+
+  have "t1 \<bar> True" "t2 \<bar> True"
+    apply (untag, rule HOL.TrueI)+
+    usingT - cancel\<open>\<open>''hi''\<close>\<close> \<open>Some _\<close> \<open>t1\<close>
+    done
+
+end
+
+
+
+subsection \<open>Deleting Assumptions\<close>
+
+\<comment> \<open>How @{method thin_tac} does it.\<close>
+lemma
+  "x > 0 \<Longrightarrow> y > 0 \<Longrightarrow> z > 0"
+  apply (tactic "eresolve_tac @{context} [Drule.thin_rl] 1")
+  oops
+
+lemma
+  "x > 0 \<Longrightarrow> y > 0 \<Longrightarrow> z > 0"
+  apply (erule Pure.thin_rl[where V = "y > 0"])
+  oops
+
+
+lemma thin_tag_rl:
+  "TAG t V \<Longrightarrow> PROP W \<Longrightarrow> PROP W"
+  by (rule thin_rl)
+
+ML \<open>
+fun thin_tag_tac ctxt s fixes =
+  Rule_Insts.eres_inst_tac ctxt [((("t", 0), Position.none), s)] fixes @{thm thin_tag_rl};
+\<close>
+
+method_setup thin_tag =
+  \<open>(Args.goal_spec -- Scan.lift (Parse.embedded_inner_syntax -- Parse.for_fixes) >>
+      (fn (quant, (prop, fixes)) => fn ctxt =>
+        SIMPLE_METHOD'' quant (thin_tag_tac ctxt prop fixes)))\<close>
+  "filter me"
+
+lemma
+  "TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG r U \<Longrightarrow> P"
+  apply (thin_tag \<open>Some _\<close>)
+  oops
+
+ML \<open>
+fun del_tags_tac ctxt = REPEAT1 (eresolve_tac ctxt @{thms thin_tag_rl} 1)
+\<close>
+
+method_setup del_tags =
+  \<open>Args.context >> (fn _ => fn ctxt => SIMPLE_METHOD (del_tags_tac ctxt))\<close> "test"
+
+lemma
+  "TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG r U \<Longrightarrow> P"
+  apply del_tags
+  oops
+
+
+subsection \<open>Main Tactics\<close>
+
+ML \<open>
+fun instantiate_TAG ctxt (t: term) =
+let
+  val ct = Thm.cterm_of ctxt t
+in (* XXX How inefficient is this? *)
+  instantiate\<open>'t = \<open>Thm.ctyp_of_cterm ct\<close> and t = ct in
+    lemma (schematic) \<open>(t \<bar> P) = TAG' t P\<close> for t :: 't by (rule TAG)\<close>
+end
+
+fun protect_tag_tac ctxt s fixes i =
+  let
+    val thm = Rule_Insts.where_rule ctxt [((("t", 0), Position.none), s)] fixes @{thm TAG};
+  in REPEAT1 (CHANGED_PROP (unfold_tac' ctxt [thm] i)) end
+
+fun protect_tags_tac ctxt props fixes i =
+  let
+    fun mk_thm s =
+      Rule_Insts.where_rule ctxt [((("t", 0), Position.none), s)] fixes @{thm TAG}
+    val thms = map mk_thm props
+  in REPEAT1 (CHANGED_PROP (unfold_tac' ctxt thms i)) end
+
+fun filter_tags_tac ctxt props fixes =
+  TRY (protect_tags_tac ctxt props fixes 1)
+  THEN TRY (del_tags_tac ctxt)
+  THEN TRY (unfold_tac' ctxt @{thms TAG'_def} 1)
+
+fun mk_auto_insert_tac tac = SUBGOAL (fn (concl, i) =>
+  case concl |> Logic.strip_assums_concl |> extract_tag_trm of
+    SOME tag => tac [tag] i
+  | _ => tac [] i
+)
+
+fun auto_filter_tags_and_insert_tac0 strict ctxt props fixes tags0 =
+  let
+    val pos_tags = map
+      (fn s => (Syntax.read_input_pos s, Proof_Context.read_term_pattern ctxt s))
+      props
+    val get_fact =
+      if strict then
+        fn (pos, t) => get_tagged_at_least_one ctxt pos t
+      else
+        fn (_, t) => get_tagged ctxt t
+    val facts = flat (map get_fact pos_tags) @ flat (map (get_tagged ctxt) tags0)
+    val protect_eqs = map (instantiate_TAG ctxt) tags0
+  in
+    REPEAT (CHANGED_PROP (unfold_tac' ctxt protect_eqs 1))
+    THEN filter_tags_tac ctxt props fixes
+    THEN TRY (Method.insert_tac ctxt facts 1)
+  end
+
+fun auto_filter_tags_and_insert_tac strict ctxt props fixes =
+  mk_auto_insert_tac
+    (fn tags => fn _ => auto_filter_tags_and_insert_tac0 strict ctxt props fixes tags) 1
+
+val props_fixes_parser: (string list * (binding * string option * mixfix) list) context_parser =
+  Scan.lift (Scan.repeat Parse.embedded_inner_syntax -- Parse.for_fixes)
+\<close>
+
+method_setup protect_tag =
+  \<open>(Scan.lift (Parse.embedded_inner_syntax -- Parse.for_fixes) >>
+      (fn (prop, fixes) => fn ctxt => SIMPLE_METHOD' (protect_tag_tac ctxt prop fixes)))\<close>
+  "TAG \<rightarrow> TAG' (single)"
+
+method_setup protect_tags =
+  \<open>props_fixes_parser >>
+     (fn (props, fixes) => fn ctxt => SIMPLE_METHOD' (protect_tags_tac ctxt props fixes))\<close>
+  "TAG \<rightarrow> TAG' (multiple)"
+
+method_setup filter_tags =
+  \<open>props_fixes_parser >>
+     (fn (props, fixes) => fn ctxt => SIMPLE_METHOD (filter_tags_tac ctxt props fixes))\<close>
+  "Filter tagged facts"
+
+method_setup tag0 =
+  \<open>Scan.lift parse_non_strict -- props_fixes_parser >>
+     (fn (non_strict, (props, fixes)) => fn ctxt =>
+        SIMPLE_METHOD (auto_filter_tags_and_insert_tac (not non_strict) ctxt props fixes))\<close>
+  "Filter tagged facts and insert matching tagged facts, taking goal tag into account"
+
+method_setup tag =
+  \<open>Scan.lift parse_non_strict -- props_fixes_parser >>
+     (fn (non_strict, (props, fixes)) => fn ctxt =>
+        SIMPLE_METHOD (DETERM (
+               auto_filter_tags_and_insert_tac (not non_strict) ctxt props fixes
+          THEN unfold_tac' ctxt @{thms TAG_def} 1)))\<close>
+  "Like tags0 but removing tags in the end"
+
+lemma
+  "TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG r U \<Longrightarrow> P"
+  apply (protect_tags \<open>Some x\<close> t for x, del_tags, unfold' TAG'_def)
+  oops
+
+lemma
+  "TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG r U \<Longrightarrow> P"
+  apply (filter_tags \<open>Some x\<close> t for x)
+  oops
+
+lemma [tagged]: "Some ''s'' \<bar> True" unfolding TAG_def ..
+lemma [tagged]: "Some u \<bar> True" unfolding TAG_def ..
+
+lemma
+  "TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG r U \<Longrightarrow> P"
+  \<comment> \<open>XXX: The first three fail to recall facts from \<open>tagged\<close>.
+      We would need to read the variables (and generalize them again) like in \<open>rule_insts.ML\<close>.\<close>
+cancel\<open>  apply (tag0 \<open>Some x\<close> t for x)
+  apply (tag \<open>Some u\<close> t for u)
+  apply (tag \<open>Some s\<close> t for s :: string)\<close>
+  apply (insertT \<open>Some _\<close>)
+  apply (tag \<open>Some _\<close>)
+  oops
+
+lemma [tagged]: "''hi'' \<bar> True" unfolding TAG_def ..
+
+lemma
+  "\<And>tt. TAG t V \<Longrightarrow> TAG (Some p) P \<Longrightarrow> TAG ''hi'' U \<Longrightarrow> TAG ''hi'' P"
+  apply (tag- t)
+  oops
+
+
+subsection \<open>Reordering Subgoals\<close>
+
+ML \<open>
+(* Find index of first subgoal in thm whose conclusion matches pred *)
+fun find_subgoal pred thm =
+  let
+    val subgoals = Thm.prems_of thm
+  in
+    Library.find_index (fn goal => pred (Logic.strip_assums_concl goal)) subgoals
+  end
+
+fun prefer_by_pred_tac pred: tactic = fn thm =>
+  let
+    val i = find_subgoal pred thm + 1
+  in
+    if i > 0 then Tactic.prefer_tac i thm else no_tac thm
+  end
+
+fun prefer_by_tag_tac0 thy tag: tactic =
+  prefer_by_pred_tac (matches_tag0 thy tag)
+
+fun prefer_by_tag_tac thy tag: tactic =
+  prefer_by_pred_tac (matches_tag thy tag)
+\<close>
+
+lemma
+  "TAG p P" "TAG q Q" "TAG r R" "TAG p P2"
+  apply -
+     apply (tactic \<open>prefer_by_tag_tac0 @{theory} term\<open>r\<close>\<close>)
+     apply (tactic \<open>prefer_by_tag_tac0 @{theory} term\<open>p\<close>\<close>)
+  apply (tactic \<open>prefer_by_tag_tac @{theory} term\<open>q\<close>\<close>)
+  oops
+
+ML \<open>
+fun prefer_tag pat st =
+  let
+    val st = Proof.assert_no_chain st
+    val thy = Proof.theory_of st
+  in
+    Proof.refine_singleton (Method.Basic (fn _ => METHOD (fn _ => prefer_by_tag_tac thy pat))) st
+  end
+
+val prefer_tag_trans: string -> Proof.state -> Proof.state = fn s => fn st =>
+  let
+    val ctxt = Proof.context_of st
+    val tag = Proof_Context.read_term_pattern ctxt s
+  in prefer_tag tag st end
+\<close>
+
+ML \<open>Outer_Syntax.command command_keyword\<open>preferT\<close> "select subgoal by tag"
+  (Parse.embedded_inner_syntax >> (Toplevel.proof o prefer_tag_trans))
+\<close>
+
+lemma
+  "TAG p P" "TAG q Q" "TAG r R" "TAG p P2"
+     preferT \<open>r\<close>
+  oops
+
+lemma
+  "\<And>t. TAG p P \<and> TAG q Q"
+  apply (rule conjI)
+   preferT \<open>q\<close>
+  oops
+
+
+subsection \<open>Diagnostic Commands\<close>
+
+ML \<open>
+val extract_tag_thm = extract_tag_trm o Thm.prop_of
+
+fun tags_of_tagged ctxt =
+  let
+    val thms = Named_Theorems.get ctxt tagged
+  in
+    Library.map_filter extract_tag_thm thms
+  end
+
+fun pretty_term_item ctxt trm = Pretty.item [Syntax.pretty_term ctxt trm]
+
+fun string_of_tags ctxt () =
+  let
+    val tags = tags_of_tagged ctxt
+    val pretty_tags = map (pretty_term_item ctxt) tags
+    val pretty = Pretty.blk (0, Pretty.fbreaks pretty_tags)
+  in
+    Pretty.string_of pretty
+  end
+
+\<comment> \<open>From \<open>Pure.thy\<close>\<close>
+val opt_modes =
+  Scan.optional (keyword\<open>(\<close> |-- Parse.!!! (Scan.repeat1 Parse.name --| keyword\<open>)\<close>)) [];
+
+\<comment> \<open>From \<open>isar_cmd.ml\<close>\<close>
+fun print_item string_of (modes, arg) = Toplevel.keep (fn state =>
+  Print_Mode.with_modes modes (fn () => writeln (string_of state arg)) ());
+
+val print_tags = print_item (string_of_tags o Toplevel.context_of);
+\<close>
+
+ML \<open>
+val _ =
+  Outer_Syntax.command command_keyword\<open>print_tags\<close>
+    "print tags of theorems in collection 'tagged'"
+    (opt_modes -- Scan.succeed () >> print_tags);
+\<close>
+
+print_tags
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Expressions.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Expressions.thy
@@ -0,0 +1,120 @@
+theory Simple_Expressions
+  imports Main
+begin
+
+datatype ('a, 'b) bexp =
+  true |
+  not "('a, 'b) bexp" |
+  "and" "('a, 'b) bexp" "('a, 'b) bexp" |
+  or "('a, 'b) bexp" "('a, 'b) bexp" |
+  imply "('a, 'b) bexp" "('a, 'b) bexp" | \<comment> \<open>Boolean connectives\<close>
+  eq "('a, 'b) exp" "('a, 'b) exp" |
+  le "('a, 'b) exp" "('a, 'b) exp" |
+  lt "('a, 'b) exp" "('a, 'b) exp" |
+  ge "('a, 'b) exp" "('a, 'b) exp" |
+  gt "('a, 'b) exp" "('a, 'b) exp"
+and ('a, 'b) exp =
+  const 'b | var 'a | if_then_else "('a, 'b) bexp" "('a, 'b) exp" "('a, 'b) exp" |
+  binop "'b \<Rightarrow> 'b \<Rightarrow> 'b" "('a, 'b) exp" "('a, 'b) exp" | unop "'b \<Rightarrow> 'b" "('a, 'b) exp"
+
+inductive check_bexp :: "('a \<rightharpoonup> 'b) \<Rightarrow> ('a, 'b :: linorder) bexp \<Rightarrow> bool \<Rightarrow> bool" and is_val
+  where
+  "check_bexp s true True" |
+  "check_bexp s (not e) (\<not> b)" if "check_bexp s e b" |
+  "check_bexp s (and e1 e2) (a \<and> b)" if "check_bexp s e1 a" "check_bexp s e2 b" |
+  "check_bexp s (or e1 e2) (a \<or> b)" if "check_bexp s e1 a" "check_bexp s e2 b" |
+  "check_bexp s (imply e1 e2) (a \<longrightarrow> b)" if "check_bexp s e1 a" "check_bexp s e2 b" |
+  "check_bexp s (eq a b) (x = v)" if "is_val s a v" "is_val s b x" |
+  "check_bexp s (le a b) (v \<le> x)" if "is_val s a v" "is_val s b x" |
+  "check_bexp s (lt a b) (v < x)" if "is_val s a v" "is_val s b x" |
+  "check_bexp s (ge a b) (v \<ge> x)" if "is_val s a v" "is_val s b x" |
+  "check_bexp s (gt a b) (v > x)" if "is_val s a v" "is_val s b x" |
+  "is_val s (const c) c" |
+  "is_val s (var x)   v" if "s x = Some v" |
+  "is_val s (if_then_else b e1 e2) (if bv then v1 else v2)"
+  if "is_val s e1 v1" "is_val s e2 v2" "check_bexp s b bv" |
+  "is_val s (binop f e1 e2) (f v1 v2)" if "is_val s e1 v1" "is_val s e2 v2" |
+  "is_val s (unop f e) (f v)" if "is_val s e v"
+
+inductive_simps check_bexp_simps:
+  "check_bexp s bexp.true bv"
+  "check_bexp s (bexp.not b) bv"
+  "check_bexp s (bexp.and b1 b2) bv"
+  "check_bexp s (bexp.or b1 b2) bv"
+  "check_bexp s (bexp.imply b1 b2) bv"
+  "check_bexp s (le i x) bv"
+  "check_bexp s (lt i x) bv"
+  "check_bexp s (ge i x) bv"
+  "check_bexp s (eq i x) bv"
+  "check_bexp s (gt i x) bv"
+
+inductive_simps is_val_simps:
+  "is_val s (const c) d"
+  "is_val s (var x) v"
+  "is_val s (if_then_else b e1 e2) v"
+  "is_val s (binop f e1 e2) v"
+  "is_val s (unop f e) v"
+
+inductive_cases check_bexp_elims:
+  "check_bexp s bexp.true bv"
+  "check_bexp s (bexp.not b) bv"
+  "check_bexp s (bexp.and b1 b2) bv"
+  "check_bexp s (bexp.or b1 b2) bv"
+  "check_bexp s (bexp.imply b1 b2) bv"
+  "check_bexp s (le i x) bv"
+  "check_bexp s (lt i x) bv"
+  "check_bexp s (ge i x) bv"
+  "check_bexp s (eq i x) bv"
+  "check_bexp s (gt i x) bv"
+
+inductive_cases is_val_elims:
+  "is_val s (const c) d"
+  "is_val s (var x) v"
+  "is_val s (if_then_else b e1 e2) v"
+  "is_val s (binop f e1 e2) v"
+  "is_val s (unop f e) v"
+
+
+
+type_synonym
+  ('a, 'b) upd = "('a * ('a, 'b) exp) list"
+
+definition is_upd where
+  "is_upd s upd s' \<equiv> \<exists>x e v. upd = (x, e) \<and> is_val s e v \<and> s' = s(x := Some v)" for upd
+
+inductive is_upds where
+  "is_upds s [] s" |
+  "is_upds s (x # xs) s''" if "is_upd s x s'" "is_upds s' xs s''"
+
+inductive_cases is_upds_NilE:
+  "is_upds s [] s'"
+and is_upds_ConsE:
+  "is_upds s (e # es) s'"
+
+inductive_simps is_upds_Nil_iff: "is_upds s [] s'"
+  and is_upds_Cons_iff: "is_upds s (f # upds) s'"
+
+lemma is_upds_appendI:
+  assumes "is_upds s upds1 s'" "is_upds s' upds2 s''"
+  shows "is_upds s (upds1 @ upds2) s''"
+  using assms by induction (auto intro: is_upds.intros)
+
+lemma is_upds_appendE:
+  assumes "is_upds s (upds1 @ upds2) s''"
+  obtains s' where "is_upds s upds1 s'" "is_upds s' upds2 s''"
+  using assms by (induction upds1 arbitrary: s) (auto intro: is_upds.intros elim!: is_upds_ConsE)
+
+lemma is_upds_NilD:
+  "s' = s" if "is_upds s [] s'"
+  using that by (rule is_upds_NilE)
+
+lemma is_upds_all_NilD:
+  assumes "is_upds s (concat upds) s'" "(\<forall>xs\<in>set upds. xs = [])"
+  shows "s' = s"
+  using assms by (simp del: Nil_eq_concat_conv flip: concat_eq_Nil_conv) (rule is_upds_NilD)
+
+lemma is_upd_const_simp:
+  "is_upd s (x, const c) s' \<longleftrightarrow> s' = s(x := Some c)"
+  unfolding is_upd_def by (simp add: is_val_simps)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language.thy
@@ -0,0 +1,547 @@
+theory Simple_Network_Language
+  imports Main
+    Munta_Model_Checker.State_Networks
+    Munta_Model_Checker.UPPAAL_State_Networks_Impl
+    FinFun.FinFun
+    Simple_Expressions
+    Munta_Tagging
+begin
+
+section \<open>Simple Networks of Automata With Broadcast Channels and Committed Locations\<close>
+
+no_notation top_assn ("true")
+
+abbreviation concat_map where
+  "concat_map f xs \<equiv> concat (map f xs)"
+
+
+type_synonym
+  ('c, 't, 's) invassn = "'s \<Rightarrow> ('c, 't) cconstraint"
+
+type_synonym
+  ('a, 's, 'c, 't, 'x, 'v) transition =
+  "'s \<times> ('x, 'v) bexp \<times> ('c, 't) cconstraint \<times> 'a \<times> ('x, 'v) upd \<times> 'c list \<times> 's"
+
+type_synonym
+  ('a, 's, 'c, 't, 'x, 'v) sta =
+  "'s set \<times> 's set \<times> ('a, 's, 'c, 't, 'x, 'v) transition set \<times> ('c, 't, 's) invassn"
+
+type_synonym
+  ('a, 's, 'c, 't, 'x, 'v) nta = "'a set \<times> ('a act, 's, 'c, 't, 'x, 'v) sta list \<times> ('x \<rightharpoonup> 'v * 'v)"
+
+context begin
+
+qualified definition conv_t where "conv_t \<equiv> \<lambda> (l,b,g,a,f,r,l'). (l,b,conv_cc g,a,f,r,l')"
+
+qualified definition conv_A where "conv_A \<equiv> \<lambda> (C, U, T, I). (C, U, conv_t ` T, conv_cc o I)"
+
+definition conv where
+  "conv \<equiv> \<lambda>(broadcast, automata, bounds). (broadcast, map conv_A automata, bounds)"
+
+end
+
+datatype 'b label = Del | Internal 'b | Bin 'b | Broad 'b
+
+definition bounded where
+  "bounded bounds s \<equiv> dom s = dom bounds \<and>
+    (\<forall>x \<in> dom s. fst (the (bounds x)) \<le> the (s x) \<and> the (s x) \<le> snd (the (bounds x)))"
+
+definition committed :: "('a, 's, 'c, 't, 'x, 'v) sta \<Rightarrow> 's set" where
+  "committed A \<equiv> fst A"
+
+definition urgent :: "('a, 's, 'c, 't, 'x, 'v) sta \<Rightarrow> 's set" where
+  "urgent A \<equiv> fst (snd A)"
+
+definition trans :: "('a, 's, 'c, 't, 'x, 'v) sta \<Rightarrow> ('a, 's, 'c, 't, 'x, 'v) transition set"
+  where
+  "trans A \<equiv> fst (snd (snd A))"
+
+definition inv :: "('a, 's, 'c, 't, 'x, 'v) sta \<Rightarrow> ('c, 't, 's) invassn" where
+  "inv A \<equiv> snd (snd (snd A))"
+
+no_notation step_sn ("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>_ \<langle>_, _, _\<rangle>" [61,61,61,61,61] 61)
+no_notation steps_sn ("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>* \<langle>_, _, _\<rangle>" [61, 61, 61,61,61] 61)
+
+datatype 'a tag = ANY 'a | TRANS 'a | SEL 'a | SEND 'a | RECV 'a
+
+inductive step_u ::
+  "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval
+  \<Rightarrow> 'a label \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>_ \<langle>_, _, _\<rangle>" [61,61,61,61,61] 61)
+where
+  step_t:
+    "\<lbrakk>
+      ''target invariant'' \<bar> \<forall>p < length N. u \<oplus> d \<turnstile> inv (N ! p) (L ! p);
+      ''nonnegative''      \<bar> d \<ge> 0;
+      ''urgent''           \<bar> (\<exists>p < length N. L ! p \<in> urgent (N ! p)) \<longrightarrow> d = 0;
+      ''bounded''          \<bar> bounded B s
+     \<rbrakk>
+    \<Longrightarrow> (broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L, s, u \<oplus> d\<rangle>" |
+  step_int:
+    "\<lbrakk>
+      TRANS ''silent'' \<bar> (l, b, g, Sil a, f, r, l') \<in> trans (N ! p);
+      ''committed'' \<bar> l \<in> committed (N ! p) \<or> (\<forall>p < length N. L ! p \<notin> committed (N ! p));
+      ''bexp''      \<bar> check_bexp s b True;
+      ''guard''     \<bar> u \<turnstile> g;
+      ''target invariant'' \<bar> \<forall>p < length N. u' \<turnstile> inv (N ! p) (L' ! p);
+      ''loc''           \<bar> L!p = l;
+      ''range''         \<bar> p < length L;
+      ''new loc''       \<bar> L' = L[p := l'];
+      ''new valuation'' \<bar> u' = [r\<rightarrow>0]u;
+      ''is_upd''        \<bar> is_upds s f s';
+      ''bounded''       \<bar> bounded B s'
+    \<rbrakk>
+    \<Longrightarrow> (broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Internal a \<langle>L', s', u'\<rangle>" |
+  step_bin:
+    "\<lbrakk>
+      ''not broadcast'' \<bar> (a \<notin> broadcast);
+      TRANS ''in''  \<bar> (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N ! p);
+      TRANS ''out'' \<bar> (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N ! q);
+      ''committed'' \<bar>
+        l1 \<in> committed (N ! p) \<or> l2 \<in> committed (N ! q) \<or> (\<forall>p < length N. L ! p \<notin> committed (N ! p));
+      ''bexp'' \<bar> check_bexp s b1 True; ''bexp'' \<bar> check_bexp s b2 True;
+      ''guard'' \<bar> u \<turnstile> g1; ''guard'' \<bar> u \<turnstile> g2;
+      ''target invariant'' \<bar> \<forall>p < length N. u' \<turnstile> inv (N ! p) (L' ! p);
+      RECV ''loc'' \<bar> L!p = l1; SEND ''loc'' \<bar> L!q = l2;
+      RECV ''range'' \<bar> p < length L; SEND ''range'' \<bar> q < length L;
+      ''different'' \<bar> p \<noteq> q;
+      ''new loc''       \<bar> L' = L[p := l1', q := l2'];
+      ''new valuation'' \<bar> u' = [r1@r2\<rightarrow>0]u;
+      ''upd'' \<bar> is_upds s f1 s'; ''upd'' \<bar> is_upds s' f2 s'';
+      ''bounded'' \<bar> bounded B s''
+    \<rbrakk>
+    \<Longrightarrow> (broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Bin a \<langle>L', s'', u'\<rangle>" |
+  step_broad:
+    "\<lbrakk>
+      ''broadcast'' \<bar> a \<in> broadcast;
+      TRANS ''out'' \<bar> (l, b, g, Out a, f, r, l') \<in> trans (N ! p);
+      TRANS ''in''  \<bar> (\<forall>p \<in> set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N ! p));
+      ''committed'' \<bar> (l \<in> committed (N ! p) \<or> (\<exists>p \<in> set ps. L ! p \<in> committed (N ! p))
+      \<or> (\<forall>p < length N. L ! p \<notin> committed (N ! p)));
+      ''bexp''    \<bar> check_bexp s b True;
+      ''bexp''    \<bar> \<forall>p \<in> set ps. check_bexp s (bs p) True;
+      ''guard''   \<bar> u \<turnstile> g;
+      ''guard''   \<bar> \<forall>p \<in> set ps. u \<turnstile> gs p;
+      ''maximal'' \<bar> \<forall>q < length N. q \<notin> set ps \<and> p \<noteq> q
+        \<longrightarrow> (\<forall>b g f r l'. (L!q, b, g, In a, f, r, l') \<in> trans (N ! q)
+        \<longrightarrow> \<not> check_bexp s b True \<or> \<not> u \<turnstile> g);
+      ''target invariant'' \<bar> \<forall>p < length N. u' \<turnstile> inv (N ! p) (L' ! p);
+      SEND ''loc''       \<bar> L!p = l;
+      SEND ''range''     \<bar> p < length L;
+      SEL ''range''      \<bar> set ps \<subseteq> {0..<length N};
+      SEL ''not sender'' \<bar> p \<notin> set ps;
+      SEL ''distinct''   \<bar> distinct ps;
+      SEL ''sorted''     \<bar> sorted ps;
+      ''new loc'' \<bar> L' = (fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'];
+      ''new valuation'' \<bar> u' = [r@concat (map rs ps)\<rightarrow>0]u;
+      ''upd''     \<bar> is_upds s f s';
+      ''upds''    \<bar> is_upds s' (concat_map fs ps) s'';
+      ''bounded'' \<bar> bounded B s''
+    \<rbrakk>
+    \<Longrightarrow> (broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Broad a \<langle>L', s'', u'\<rangle>"
+lemmas [intro?] = step_u.intros[unfolded TAG_def]
+
+text \<open>Comments:
+item Should there be an error transition + state if states run of bounds or updates are undefined?
+Then one could run a reachability check for the error state.
+item Should the state be total?
+item Note that intermediate states are allowed to run out of bounds
+\<close>
+
+definition steps_u ::
+  "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval
+  \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>* \<langle>_, _, _\<rangle>" [61,61,61,61,61] 61)
+where
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<equiv>
+    (\<lambda> (L, s, u) (L', s', u'). \<exists>a. A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>)** (L, s, u) (L', s', u')"
+
+paragraph \<open>Misc\<close>
+
+lemma clock_val_concat_iff:
+  \<open>u \<turnstile> concat gs \<longleftrightarrow> (\<forall>g \<in> set gs. u \<turnstile> g)\<close>
+  by (auto intro: guard_concat concat_guard)
+
+lemma clock_val_append_iff:
+  \<open>u \<turnstile> g1 @ g2 \<longleftrightarrow> u \<turnstile> g1 \<and> u \<turnstile> g2\<close>
+proof -
+  have "u \<turnstile> g1 @ g2 \<longleftrightarrow> u \<turnstile> concat [g1, g2]"
+    by simp
+  also have "... \<longleftrightarrow> u \<turnstile> g1 \<and> u \<turnstile> g2"
+    unfolding clock_val_concat_iff by simp
+  finally show ?thesis .
+qed
+
+
+subsection \<open>Product Construction\<close>
+
+locale Prod_TA_Defs =
+  fixes A :: "('a, 's, 'c, 't, 'x, 'v :: linorder) nta"
+begin
+
+definition
+  "broadcast = fst A"
+
+definition
+  "N i = fst (snd A) ! i"
+
+definition
+  "bounds = snd (snd A)"
+
+definition \<comment>\<open>Number of processes\<close>
+  "n_ps = length (fst (snd A))"
+
+definition states  :: "'s list set" where
+  "states \<equiv> {L. length L = n_ps \<and>
+    (\<forall> i. i < n_ps \<longrightarrow> L ! i \<in> (\<Union> (l, e, g, a, r, u, l') \<in> (trans (N i)). {l, l'}))}"
+
+definition
+  "prod_inv \<equiv> \<lambda>(L, s). if L \<in> states then concat (map (\<lambda>i. inv (N i) (L ! i)) [0..<n_ps]) else []"
+
+definition
+  "trans_int =
+    {((L, s), g, Internal a, r, (L', s')) | L s l b g f p a r l' L' s'.
+      (l, b, g, Sil a, f, r, l') \<in> trans (N p) \<and>
+      (l \<in> committed (N p) \<or> (\<forall>p < n_ps. L ! p \<notin> committed (N p))) \<and>
+      L!p = l \<and> p < length L \<and> L' = L[p := l'] \<and> is_upds s f s' \<and> check_bexp s b True \<and>
+      L \<in> states \<and> bounded bounds s \<and> bounded bounds s'
+    }"
+
+definition
+  "trans_bin =
+    {((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+      L s L' s' s'' a p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'.
+      a \<notin> broadcast \<and>
+      (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+      (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+      (l1 \<in> committed (N p) \<or> l2 \<in> committed (N q) \<or> (\<forall>p < n_ps. L ! p \<notin> committed (N p))) \<and>
+      L!p = l1 \<and> L!q = l2 \<and> p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+      check_bexp s b1 True \<and> check_bexp s b2 True \<and>
+      L' = L[p := l1', q := l2'] \<and> is_upds s f1 s' \<and> is_upds s' f2 s'' \<and>
+      L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+    }"
+
+definition
+  "trans_broad =
+    {((L, s), g @ concat (map gs ps), Broad a, r @ concat (map rs ps), (L', s'')) |
+    L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps.
+      a \<in> broadcast  \<and>
+      (l, b, g, Out a, f, r, l') \<in> trans (N p) \<and>
+      (\<forall>p \<in> set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+      (l \<in> committed (N p) \<or> (\<exists>p \<in> set ps. L ! p \<in> committed (N p))
+      \<or> (\<forall>p < n_ps. L ! p \<notin> committed (N p))) \<and>
+      (\<forall>q < n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+        \<not> (\<exists>b g f r l'. (L!q, b, g, In a, f, r, l') \<in> trans (N q) \<and> check_bexp s b True)) \<and>
+      L!p = l \<and>
+      p < length L \<and> set ps \<subseteq> {0..<n_ps} \<and> p \<notin> set ps \<and> distinct ps \<and> sorted ps \<and>
+      check_bexp s b True \<and> (\<forall>p \<in> set ps. check_bexp s (bs p) True) \<and>
+      L' = (fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'] \<and>
+      is_upds s f s' \<and> is_upds s' (concat_map fs ps) s'' \<and>
+      L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+    }"
+
+definition
+  "trans_prod = trans_int \<union> trans_bin \<union> trans_broad"
+
+definition
+  "prod_ta = (trans_prod, prod_inv :: ('s list \<times> ('x \<Rightarrow> 'v option) \<Rightarrow> _))"
+
+lemma inv_of_prod[simp]:
+  "inv_of prod_ta = prod_inv"
+  unfolding prod_ta_def inv_of_def by simp
+
+lemma trans_of_prod[simp]:
+  "trans_of prod_ta = trans_prod"
+  unfolding prod_ta_def trans_of_def by simp
+
+lemma A_split:
+  \<open>A = (broadcast, map N [0..<n_ps], bounds)\<close>
+  unfolding broadcast_def N_def bounds_def n_ps_def by (cases A) (simp add: List.map_nth)
+
+lemma map_N_n_ps_simp:
+  "map N [0..<n_ps] ! p = N p"
+  unfolding N_def n_ps_def List.map_nth ..
+
+lemma N_split_simp[simp]:
+  assumes "A = (broadcast', N', B)"
+  shows "N' ! i = N i"
+  unfolding N_def unfolding assms by simp
+
+lemma state_preservation_updI:
+  assumes "l' \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})" "L \<in> states"
+  shows "L[p := l'] \<in> states"
+  using assms unfolding states_def by (fastforce simp: nth_list_update')
+
+lemma state_preservation_fold_updI:
+  assumes "\<forall>p \<in> set ps. ls' p \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})" "L \<in> states"
+  shows "fold (\<lambda>p L. L[p := ls' p]) ps L \<in> states"
+  using assms by (induction ps arbitrary: L) (auto intro: state_preservation_updI)
+
+lemma state_set_states:
+  "Simulation_Graphs_TA.state_set prod_ta \<subseteq> {(l, s). l \<in> states}"
+  unfolding prod_ta_def state_set_def
+  unfolding trans_of_def trans_prod_def
+  unfolding trans_int_def trans_bin_def trans_broad_def
+  by auto (auto intro: state_preservation_updI state_preservation_fold_updI)
+
+lemma trans_prod_bounded_inv:
+  \<open>bounded bounds s'\<close> if \<open>((L, s), g, a, r, (L', s')) \<in> trans_prod\<close>
+  using that unfolding bounds_def trans_prod_def trans_int_def trans_bin_def trans_broad_def
+  by (auto simp: bounds_def)
+
+lemma trans_prod_states_inv:
+  \<open>L' \<in> states\<close> if \<open>((L, s), g, a, r, (L', s')) \<in> trans_prod\<close> \<open>L \<in> states\<close>
+  using that
+  unfolding bounds_def trans_prod_def trans_int_def trans_bin_def trans_broad_def
+  apply clarsimp
+  apply safe
+         apply (force intro!: state_preservation_updI)
+        apply (force intro!: state_preservation_updI)
+       apply (force intro!: state_preservation_updI)
+      apply (force intro!: state_preservation_updI)
+     apply (force intro!: state_preservation_updI)
+    apply (fastforce intro!: state_preservation_updI state_preservation_fold_updI)
+  subgoal
+    apply (rule state_preservation_updI)
+    apply force
+    apply (force intro!: state_preservation_fold_updI)
+    done
+  apply (fastforce intro!: state_preservation_updI state_preservation_fold_updI)
+  done
+
+end (* Prod TA Defs *)
+
+
+locale Prod_TA_sem =
+  Prod_TA_Defs A for A :: "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta"
+begin
+
+lemma prod_invI[intro]:
+  \<open>u \<turnstile> prod_inv (L, s)\<close> if \<open>\<forall>p<n_ps. u \<turnstile> Simple_Network_Language.inv (N p) (L ! p)\<close>
+  using that unfolding prod_inv_def by (auto intro!: guard_concat)
+
+lemma prod_invD[dest]:
+  \<open>\<forall>p<n_ps. u \<turnstile> Simple_Network_Language.inv (N p) (L ! p)\<close> if \<open>u \<turnstile> prod_inv (L, s)\<close> \<open>L \<in> states\<close>
+  using that unfolding prod_inv_def by (auto elim: concat_guard)
+
+lemma delay_sound:
+  assumes "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>" "L \<in> states" "bounded bounds s"
+          "\<forall>N \<in> set (fst (snd A)). urgent N = {}"
+        shows "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+proof -
+  from assms(4) have "l \<notin> urgent (N p)" if "p < n_ps" for p l
+    using that unfolding N_def n_ps_def by auto
+  then show ?thesis
+    by (subst A_split) (use assms in \<open>cases, auto intro!: step_t simp: TAG_def\<close>)
+qed
+
+lemma action_sound:
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" if "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+  using that
+proof cases
+  case prems: (1 g r)
+  note [simp add] = map_N_n_ps_simp clock_val_append_iff clock_val_concat_iff
+  from \<open>prod_ta \<turnstile> (L, s) \<longrightarrow>g,a,r (L', s')\<close>[THEN state_setI2] have "L' \<in> states"
+    using state_set_states that by fast
+  from \<open>prod_ta \<turnstile> (L, s) \<longrightarrow>g,a,r (L', s')\<close> show ?thesis
+    unfolding trans_of_prod trans_prod_def
+  proof safe
+    assume "((L, s), g, a, r, L', s') \<in> trans_int"
+    then show "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      unfolding trans_int_def
+      apply clarsimp
+      using prems \<open>L' \<in> _\<close>
+      by (subst A_split) (intro step_int; simp add: TAG_def; elim prod_invD; assumption)
+  next
+    assume "((L, s), g, a, r, L', s') \<in> trans_bin"
+    then show "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      unfolding trans_bin_def
+      using prems \<open>L' \<in> _\<close>
+      apply clarsimp
+      apply (subst A_split, standard)
+                     apply (assumption | simp; elim prod_invD; assumption)+
+      done
+  next
+    assume "((L, s), g, a, r, L', s') \<in> trans_broad"
+    then show "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      using prems \<open>L' \<in> _\<close>
+      unfolding trans_broad_def inv_of_prod
+      apply clarsimp
+      apply (subst A_split)
+      apply standard
+      apply (simp; elim prod_invD; assumption)+
+      apply fastforce+
+      done
+  qed
+qed
+
+lemma bounded_inv:
+  \<open>bounded bounds s'\<close> if \<open>A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>\<close>
+  using that unfolding bounds_def by cases (simp add: TAG_def)+
+
+lemma states_inv:
+  \<open>L' \<in> states\<close> if \<open>A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>\<close> \<open>L \<in> states\<close>
+  using that unfolding bounds_def
+proof cases
+  case (step_t N d B broadcast)
+  with \<open>L \<in> states\<close> show ?thesis
+    by simp
+next
+  case prems[unfolded TAG_def]: (step_int l b g a f r l' N' p B broadcast)
+  from \<open>A = _\<close> prems(3) have "l' \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})"
+    by force
+  with \<open>L \<in> states\<close> show ?thesis
+    unfolding \<open>L' = _\<close> by (intro state_preservation_updI)
+next
+  case prems[unfolded TAG_def]:
+    (step_bin broadcast a l1 b1 g1 f1 r1 l1' N' p l2 b2 g2 f2 r2 l2' q s' B)
+  from \<open>A = _\<close> prems(4, 5) have
+    "l1' \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})"
+    "l2' \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N q). {l, l'})"
+    by force+
+  with \<open>L \<in> states\<close> show ?thesis
+    unfolding \<open>L' = _\<close> by (intro state_preservation_updI)
+next
+  case prems[unfolded TAG_def]: (step_broad a broadcast l b g f r l' N' p ps bs gs fs rs ls' s' B)
+  from \<open>A = _\<close> prems(4, 5) have
+    "l' \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})"
+    "\<forall> q \<in> set ps. ls' q \<in> (\<Union> (l, b, g, a, r, u, l') \<in> trans (N q). {l, l'})"
+    by force+
+  with \<open>L \<in> states\<close> show ?thesis
+    unfolding \<open>L' = _\<close> by (intro state_preservation_updI state_preservation_fold_updI)
+qed
+
+end (* Prod TA Defs on a time domain *)
+
+
+locale Prod_TA =
+  Prod_TA_sem A for A :: "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta" +
+  assumes broadcast_receivers_unguarded:
+    "\<forall>p < n_ps. \<forall>l b g a f r l'. (l, b, g, In a, f, r, l') \<in> trans (N p) \<and> a \<in> broadcast \<longrightarrow> g = []"
+begin
+
+lemma action_complete:
+  "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+  if "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" "a \<noteq> Del" "L \<in> states" "bounded bounds s"
+using that(1) proof cases
+  case (step_t N d B broadcast)
+  then show ?thesis
+    using that(2) by auto
+next
+  case prems[unfolded TAG_def]: (step_int l b g a' f r l' N' p B broadcast')
+  have [simp]:
+    "B = bounds" "broadcast' = broadcast" "length N' = n_ps"
+    unfolding bounds_def broadcast_def n_ps_def unfolding prems(1) by simp+
+  have [simp]:
+    "N' ! i = N i" for i
+    unfolding N_def unfolding prems(1) by simp
+  have "prod_ta \<turnstile> (L, s) \<longrightarrow>g,Internal a',r (L', s')"
+  proof -
+    from prems \<open>L \<in> states\<close> \<open>bounded _ s\<close> have "((L, s),g,Internal a',r,(L', s')) \<in> trans_int"
+      unfolding trans_int_def
+      by simp (rule exI conjI HOL.refl | assumption | (simp; fail))+
+    then show ?thesis
+      unfolding prod_ta_def trans_of_def trans_prod_def by simp
+  qed
+  moreover have "u \<turnstile> g"
+    by (rule prems)
+  moreover have "u' \<turnstile> inv_of prod_ta (L', s')"
+    using prems(7) by auto
+  moreover have "u' = [r\<rightarrow>0]u"
+    by (rule prems)
+  ultimately show ?thesis
+    unfolding \<open>a = _\<close> ..
+next
+  case prems[unfolded TAG_def]:
+    (step_bin a' broadcast' l1 b1 g1 f1 r1 l1' N' p l2 b2 g2 f2 r2 l2' q s'' B)
+  have [simp]:
+    "B = bounds" "broadcast' = broadcast" "length N' = n_ps"
+    unfolding bounds_def broadcast_def n_ps_def unfolding prems(1) by simp+
+  have [simp]:
+    "N' ! i = N i" for i
+    unfolding N_def unfolding prems(1) by simp
+  have "prod_ta \<turnstile> (L, s) \<longrightarrow>g1 @ g2,Bin a',r1 @ r2 (L', s')"
+  proof -
+    from prems \<open>L \<in> states\<close> \<open>bounded bounds s\<close> have
+      "((L, s),g1 @ g2,Bin a',r1 @ r2,(L', s')) \<in> trans_bin"
+      unfolding trans_bin_def
+      using [[simproc add: ex_reorder4]]
+      by simp (rule exI conjI HOL.refl | assumption | fast)+
+    then show ?thesis
+      unfolding prod_ta_def trans_of_def trans_prod_def by simp
+  qed
+  moreover have "u \<turnstile> g1 @ g2"
+    using \<open>u \<turnstile> g1\<close> \<open>u \<turnstile> g2\<close> by (rule guard_append)
+  moreover have "u' \<turnstile> inv_of prod_ta (L', s')"
+    using prems by auto
+  moreover have "u' = [r1@r2\<rightarrow>0]u"
+    by (rule prems)
+  ultimately show ?thesis
+    unfolding \<open>a = _\<close> ..
+next
+  case prems[unfolded TAG_def]: (step_broad a' broadcast' l b g f r l' N' p ps bs gs fs rs ls' s'' B)
+  have [simp]:
+    "B = bounds" "broadcast' = broadcast" "length N' = n_ps"
+    unfolding bounds_def broadcast_def n_ps_def unfolding prems(1) by simp+
+  have [simp]:
+    "N' ! i = N i" for i
+    unfolding N_def unfolding prems(1) by simp
+  let ?r = "r @ concat (map rs ps)" and ?g = "g @ concat (map gs ps)"
+  have "prod_ta \<turnstile> (L, s) \<longrightarrow>?g,Broad a',?r (L', s')"
+  proof -
+    have *: "\<not> u \<turnstile> g \<longleftrightarrow> False" if
+      "p < n_ps" "(l, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+      "a' \<in> broadcast"
+      for l b g a' f r l' p
+    proof -
+      from that broadcast_receivers_unguarded have \<open>g = []\<close>
+        by blast
+      then show ?thesis
+        by auto
+    qed
+    from prems \<open>L \<in> states\<close> \<open>bounded bounds s\<close> have "((L, s),?g,Broad a',?r,(L', s')) \<in> trans_broad"
+      unfolding trans_broad_def
+      by clarsimp
+         (intro exI conjI HOL.refl; (rule HOL.refl | assumption | fastforce simp: *))
+   then show ?thesis
+      unfolding prod_ta_def trans_of_def trans_prod_def by simp
+  qed
+  moreover have "u \<turnstile> ?g"
+    using prems by (auto intro!: guard_append guard_concat)
+  moreover have "u' \<turnstile> inv_of prod_ta (L', s')"
+    using prems by auto
+  moreover have "u' = [?r\<rightarrow>0]u"
+    by (rule prems)
+  ultimately show ?thesis
+    unfolding \<open>a = _\<close> ..
+qed
+
+lemma delay_complete:
+  assumes "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  obtains d where "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>"
+  using assms
+proof cases
+  case prems: (step_t N' d B broadcast)
+  have [simp]:
+    "length N' = n_ps"
+    unfolding n_ps_def unfolding prems(1) by simp+
+  have [simp]:
+    "N' ! i = N i" for i
+    unfolding N_def unfolding prems(1) by simp
+  from prems show ?thesis
+  by (intro that[of d]; unfold \<open>u' = _\<close> \<open>L' = L\<close> \<open>s' = s\<close> TAG_def)
+     (rule step_t.intros; (unfold inv_of_prod; rule prod_invI)?; simp)
+qed
+
+lemma step_iff:
+  "(\<exists> a. A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>) \<longleftrightarrow> prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+  if "bounded bounds s" "L \<in> states" "\<forall>N \<in> set (fst (snd A)). urgent N = {}"
+  using that(1,2)
+  apply safe
+  subgoal for a
+    by (cases a; blast dest: action_complete elim: delay_complete)
+  by (auto intro: action_sound delay_sound[OF _ _ _ that(3)])
+
+end (* Prod TA *)
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Deadlock_Checking.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Deadlock_Checking.thy
@@ -0,0 +1,112 @@
+theory Simple_Network_Language_Deadlock_Checking
+  imports Simple_Network_Language_Model_Checking Munta_Model_Checker.Deadlock_Checking
+begin
+
+context Simple_Network_Impl_nat_ceiling_start_state \<comment> \<open>slow: 120s\<close>
+begin
+
+context
+  fixes \<phi>
+  assumes formula_is_false: "formula = formula.EX (not sexp.true)"
+begin
+
+lemma F_is_FalseI:
+  shows "PR_CONST F = (\<lambda>_. False)"
+  by (subst formula_is_false) auto
+
+lemma final_fun_is_False:
+  "Fi a = False"
+  by (subst formula_is_false) auto
+
+lemmas deadlock_checker_hoare = impl.deadlock_checker_hoare[
+    OF final_fun_is_False F_is_FalseI, folded deadlock_start_iff has_deadlock_def]
+
+end
+
+schematic_goal deadlock_checker_alt_def:
+  "impl.deadlock_checker \<equiv> ?impl"
+  unfolding impl.deadlock_checker_def
+  unfolding impl.succs_impl_def
+  unfolding impl.check_deadlock_neg_impl_def impl.check_deadlock_impl_def
+  apply (abstract_let impl.is_start_in_states_impl is_start)
+  unfolding impl.is_start_in_states_impl_def
+  apply (abstract_let impl.E_op''_impl E_op''_impl)
+  unfolding impl.E_op''_impl_def fw_impl'_int
+  apply (abstract_let trans_impl trans_impl)
+  unfolding trans_impl_def
+  apply (abstract_let int_trans_impl int_trans_impl)
+  apply (abstract_let bin_trans_from_impl bin_trans_impl)
+  apply (abstract_let broad_trans_from_impl broad_trans_impl)
+  unfolding int_trans_impl_def bin_trans_from_impl_def broad_trans_from_impl_def
+  apply (abstract_let trans_in_broad_grouped trans_in_broad_grouped)
+  apply (abstract_let trans_out_broad_grouped trans_out_broad_grouped)
+  apply (abstract_let trans_in_map trans_in_map)
+  apply (abstract_let trans_out_map trans_out_map)
+  apply (abstract_let int_trans_from_all_impl int_trans_from_all_impl)
+  unfolding int_trans_from_all_impl_def
+  apply (abstract_let int_trans_from_vec_impl int_trans_from_vec_impl)
+  unfolding int_trans_from_vec_impl_def
+  apply (abstract_let int_trans_from_loc_impl int_trans_from_loc_impl)
+  unfolding int_trans_from_loc_impl_def
+  apply (abstract_let trans_i_map trans_i_map)
+  unfolding trans_out_broad_grouped_def trans_out_broad_map_def
+  unfolding trans_in_broad_grouped_def trans_in_broad_map_def
+  unfolding trans_in_map_def trans_out_map_def
+  unfolding trans_i_map_def
+  apply (abstract_let trans_map trans_map)
+  apply (abstract_let "inv_fun :: nat list \<times> int list \<Rightarrow> _" inv_fun)
+  unfolding inv_fun_alt_def
+  apply (abstract_let invs2 invs)
+  unfolding invs2_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i) (* Could be killed *)
+  apply (abstract_let n_ps n_ps)
+  (* These would need to be moved to a defs locale *)
+  unfolding k_i_def impl.state_copy_impl_def impl.a0_impl_def impl.abstr_repair_impl_def
+    impl.subsumes_impl_def impl.emptiness_check_impl_def impl.abstra_repair_impl_def
+    impl.init_dbm_impl_def
+  by (rule Pure.reflexive)
+
+end
+
+concrete_definition deadlock_checker uses
+  Simple_Network_Impl_nat_ceiling_start_state.deadlock_checker_alt_def
+
+definition precond_dc where
+  "precond_dc
+    show_clock show_state broadcast bounds' automata m num_states num_actions k L0 s0 formula \<equiv>
+    if Simple_Network_Impl_nat_ceiling_start_state
+      broadcast bounds' automata m num_states num_actions k L0 s0 formula
+    then
+      deadlock_checker
+        broadcast bounds' automata m num_states num_actions k L0 s0 show_clock show_state
+      \<bind> (\<lambda> x. return (Some x))
+    else return None"
+
+theorem deadlock_check:
+  "<emp> precond_dc
+    show_clock show_state broadcast bounds automata m num_states num_actions k L0 s0
+      (formula.EX (not sexp.true))
+    <\<lambda> Some r \<Rightarrow> \<up>(
+        Simple_Network_Impl_nat_ceiling_start_state
+          broadcast bounds automata m num_states num_actions k L0 s0 (formula.EX (not sexp.true)) \<and>
+          r = has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_ . 0)
+        )
+     | None \<Rightarrow> \<up>(\<not> Simple_Network_Impl_nat_ceiling_start_state
+        broadcast bounds automata m num_states num_actions k L0 s0 (formula.EX (not sexp.true)))
+    >t"
+proof -
+  define A where "A \<equiv> N broadcast automata bounds"
+  define check where "check \<equiv> A,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> (formula.EX (not sexp.true))"
+  note [sep_heap_rules] = Simple_Network_Impl_nat_ceiling_start_state.deadlock_checker_hoare[
+      OF _ HOL.refl,
+      of broadcast bounds automata m num_states num_actions k L0 s0 _,
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      folded A_def
+      ]
+  show ?thesis
+    unfolding A_def[symmetric] precond_dc_def
+    by (sep_auto simp: deadlock_checker.refine[symmetric] pure_def has_deadlock_def)
+qed
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Impl.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Impl.thy
@@ -0,0 +1,903 @@
+theory Simple_Network_Language_Impl
+  imports
+    Simple_Network_Language
+    Munta_Base.Normalized_Zone_Semantics_Impl_Refine
+    "HOL-Library.IArray" "HOL-Library.AList"
+    Munta_Base.More_Methods
+    Munta_Base.Bijective_Embedding
+    TA_Impl_Misc2
+    TA_More2
+    TA_Equivalences
+    "HOL-Eisbach.Eisbach_Tools"
+    Munta_Base.TA_Syntax_Bundles
+begin
+
+unbundle no_library_syntax
+
+paragraph \<open>Default maps\<close>
+
+definition default_map_of :: "'b \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b" where
+  "default_map_of a xs \<equiv> FinFun.map_default a (map_of xs)"
+
+lemma default_map_of_alt_def:
+  "default_map_of a xs x = (if x \<in> dom (map_of xs) then the (map_of xs x) else a)"
+  unfolding default_map_of_def map_default_def by (auto split: option.split_asm)
+
+lemma range_default_map_of:
+  "range (default_map_of x xs) \<subseteq> snd ` set xs \<union> {x}"
+  unfolding default_map_of_def map_default_def
+  by (auto split: option.split_asm) (meson img_snd map_of_SomeD)
+
+lemma finite_range_default_map_of:
+  "finite (range (default_map_of x m))"
+proof -
+  have "range (default_map_of x m) \<subseteq> the ` range (map_of m) \<union> {x}"
+    unfolding default_map_of_def FinFun.map_default_def
+    by (auto split: option.splits) (metis image_eqI option.sel rangeI)
+  also have "finite \<dots>"
+    by (blast intro: finite_range_map_of)
+  finally show ?thesis .
+qed
+
+lemma map_index'_inj:
+  "L = L'"
+  if "length L = length L'" "map_index' n f L = map_index' n g L'" "set L \<subseteq> S" "set L' \<subseteq> T"
+     "\<forall> i < length L + n. \<forall> x \<in> S. \<forall>y \<in> T. f i x = g i y \<longrightarrow> x = y"
+  using that
+proof (induction "length L" arbitrary: L L' n)
+  case 0
+  then show ?case
+    by simp
+next
+  case (Suc x)
+  from Suc.prems obtain a b L1 L1' where *:
+    "length L1 = x" "length L1' = x" "L = a # L1" "L' = b # L1'"
+    by (smt Suc.hyps(2) length_Suc_conv)
+  show ?case
+    unfolding \<open>L = _\<close> \<open>L' = _\<close>
+    apply (clarsimp, rule conjI)
+    subgoal
+      by (smt *(3,4) Suc.hyps(2) Suc.prems Suc_less_eq add_Suc_shift less_add_Suc2 list.set_intros(1) list_tail_coinc map_index'.simps(2) set_mp)
+    subgoal
+      apply (rule Suc.hyps)
+      using Suc.prems * by auto
+    done
+qed
+
+lemma map_index_inj:
+  "L = L'"
+  if "map_index f L = map_index g L'" "set L \<subseteq> S" "set L' \<subseteq> T"
+     "\<forall> i < length L. \<forall> x \<in> S. \<forall>y \<in> T. f i x = g i y \<longrightarrow> x = y"
+  using that by - (rule map_index'_inj, auto dest: map_index_eq_imp_length_eq)
+
+lemma map_index_inj1:
+  "L = L'"
+  if "map_index f L = map_index g L'"
+     "\<forall> i < length L. f i (L ! i) = g i (L' ! i) \<longrightarrow> L ! i = L' ! i"
+proof (intros add: nth_equalityI)
+  from that(1) show \<open>length L = length L'\<close>
+    by (rule map_index_eq_imp_length_eq)
+  fix i :: \<open>nat\<close>
+  assume \<open>i < length L\<close>
+  with that have "map_index f L ! i = map_index g L' ! i"
+    by auto
+  with \<open>i < length L\<close> \<open>length L = length L'\<close> have "f i (L ! i) = g i (L' ! i)"
+    by simp
+  with \<open>i < length L\<close> that(2) show \<open>L ! i = L' ! i\<close>
+    by simp
+qed
+
+lemma map_index_update:
+  "map_index f (xs[i := a]) = (map_index f xs)[i := f i a]"
+  by (rule nth_equalityI) (auto simp: nth_list_update')
+
+lemma map_trans_broad_aux1:
+  "map_index map_loc (fold (\<lambda>p L. L[p := ls' p]) ps L) =
+  fold (\<lambda>p L. L[p := map_loc p (ls' p)]) ps (map_index map_loc L)"
+  by (induction ps arbitrary: L) (auto simp: map_index_update)
+
+lemma InD2:
+  fixes map_action
+  assumes "inj map_action" "In (map_action a) = map_act map_action a'"
+  shows "a' = In a"
+  using assms(2) by (cases a')  (auto simp: injD[OF assms(1)])
+
+lemma OutD2:
+  fixes map_action
+  assumes "inj map_action" "Out (map_action a) = map_act map_action a'"
+  shows "a' = Out a"
+  using assms(2) by (cases a')  (auto simp: injD[OF assms(1)])
+
+lemma (in Prod_TA_Defs) trans_broad_alt_def:
+  "trans_broad =
+    {((L, s), g @ concat (map gs ps), Broad a, r @ concat (map rs ps), (L', s'')) |
+    L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps.
+      a \<in> broadcast  \<and>
+      (l, b, g, Out a, f, r, l') \<in> trans (N p) \<and>
+      (\<forall>p \<in> set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+      (l \<in> committed (N p) \<or> (\<exists>p \<in> set ps. L ! p \<in> committed (N p))
+      \<or> (\<forall>p < n_ps. L ! p \<notin> committed (N p))) \<and>
+      (\<forall>q < n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+        \<not> (\<exists>b g f r l'. (L!q, b, g, In a, f, r, l') \<in> trans (N q) \<and> check_bexp s b True)) \<and>
+      L!p = l \<and>
+      p < length L \<and> set ps \<subseteq> {0..<n_ps} \<and> p \<notin> set ps \<and> distinct ps \<and> sorted ps \<and>
+      check_bexp s b True \<and> (\<forall>p \<in> set ps. check_bexp s (bs p) True) \<and>
+      L' = (fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'] \<and>
+      is_upds s f s' \<and> is_upds s' (concat_map fs ps) s'' \<and>
+      L \<in> states \<and> bounded bounds s \<and> bounded bounds s'' \<and>
+      (\<forall>p. p\<notin>set ps \<longrightarrow> bs p = bexp.true) \<and> (\<forall>p. p\<notin>set ps \<longrightarrow> gs p = []) \<and>
+      (\<forall>p. p\<notin>set ps \<longrightarrow> fs p = []) \<and> (\<forall>p. p\<notin>set ps \<longrightarrow> rs p = [])
+    }"
+  unfolding trans_broad_def
+proof ((intro Collect_eqI iffI; elims add: more_elims), goal_cases)
+  case prems: (1 x L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps)
+  let ?f = "\<lambda>gs p. if p \<in> set ps then gs p else []"
+  let ?bs = "\<lambda>p. if p \<in> set ps then bs p else bexp.true"
+  let ?gs = "?f gs" let ?fs = "?f fs" let ?rs = "?f rs"
+  have [simp]: "map gs ps = map ?gs ps" "map rs ps = map ?rs ps" "map fs ps = map ?fs ps"
+    by (simp cong: map_cong)+
+  with prems show ?case
+    by (inst_existentials L s L' s' s'' a p l b g f r l' ?bs ?gs ?fs ?rs ls' ps; (assumption | simp))
+next
+  case (2 x L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps)
+  then show ?case
+    by blast
+qed
+
+
+definition
+  "conv_automaton \<equiv> \<lambda>(committed, urgent, trans, inv).
+    (committed,
+     urgent,
+     map (\<lambda>(l, b, g, a, f, r, l'). (l, b, conv_cc g, a, f, r, l')) trans,
+     map (\<lambda>(s, cc). (s, conv_cc cc)) inv)"
+
+definition
+  "automaton_of \<equiv>
+    \<lambda>(committed, urgent, trans, inv). (set committed, set urgent, set trans, default_map_of [] inv)"
+
+locale Simple_Network_Impl_Defs =
+  fixes automata ::
+    "('s list \<times> 's list \<times> ('a act, 's, 'c, 't, 'x, int) transition list
+      \<times> ('s \<times> ('c, 't) cconstraint) list) list"
+    and broadcast :: "'a list"
+    and bounds' :: "('x \<times> (int \<times> int)) list"
+begin
+
+definition \<comment>\<open>Number of state variables\<close>
+  "n_vs = length bounds'"
+
+definition
+  "B x \<equiv> if x \<in> dom (map_of bounds') then the (map_of bounds' x) else (0, 0)"
+
+sublocale Prod_TA_Defs
+  "(set broadcast, map automaton_of automata, map_of bounds')" .
+
+lemma L_len[intro, dest]:
+  "length L = n_ps" if "L \<in> states"
+  using that unfolding states_def by simp
+
+lemma N_eq:
+  \<open>N i = automaton_of (automata ! i)\<close> if \<open>i < n_ps\<close>
+  using that unfolding N_def n_ps_def fst_conv snd_conv by (intro nth_map; simp)
+
+end
+
+locale Simple_Network_Impl =
+  fixes automata ::
+    "('s list \<times> 's list \<times> ('a act, 's, 'c, int, 'x, int) transition list
+      \<times> ('s \<times> ('c, int) cconstraint) list) list"
+    and broadcast :: "'a list"
+    and bounds' :: "('x \<times> (int \<times> int)) list"
+begin
+
+sublocale Simple_Network_Impl_Defs automata broadcast bounds' .
+
+end
+
+
+paragraph \<open>Mapping through the product construction\<close>
+
+lemma f_the_inv_f:
+  "f (the_inv f x) = x" if "inj f" "x \<in> range f"
+  using that by (auto simp: the_inv_f_f)
+
+method fprem =
+  (match premises in R: _ \<Rightarrow> \<open>rule R[elim_format]\<close>, assumption)
+
+context Simple_Network_Impl
+begin
+
+paragraph \<open>Conversion from integers to reals commutes with product construction.\<close>
+
+sublocale conv: Prod_TA_Defs
+  "(set broadcast, map (Simple_Network_Language.conv_A o automaton_of) automata, map_of bounds')" .
+
+lemma (in -) conv_ac_inj:
+  "ac = ac'" if "conv_ac ac = conv_ac ac'"
+  using that by (cases ac; cases ac'; auto)
+
+lemma (in -) conv_cc_inj:
+  "cc = cc'" if "conv_cc cc = conv_cc cc'"
+  using that by (subst (asm) inj_map_eq_map) (auto simp add: conv_ac_inj intro: injI)
+
+context
+begin
+
+lemma conv_alt_def:
+  "conv (set broadcast, map automaton_of automata, map_of bounds') =
+    (set broadcast, map (Simple_Network_Language.conv_A o automaton_of) automata, map_of bounds')"
+  unfolding conv_def by simp
+
+private lemma 2:
+  "Simple_Network_Language.conv_A o automaton_of = (\<lambda>(committed, urgent, trans, inv).
+    (set committed,
+     set urgent,
+     set (map Simple_Network_Language.conv_t trans),
+     default_map_of [] (map (\<lambda> (l, g). (l, conv_cc g)) inv)))"
+  apply (rule ext)
+  apply clarsimp
+  unfolding Simple_Network_Language.conv_A_def automaton_of_def
+  apply (clarsimp split: prod.split)
+  unfolding default_map_of_def
+  unfolding FinFun.map_default_def
+  apply (rule ext)
+  subgoal for xs a
+    by (induction xs) auto
+  done
+
+lemma conv_n_ps_eq:
+  "conv.n_ps = n_ps"
+  by (simp add: Prod_TA_Defs.n_ps_def)
+
+lemma conv_N_eq:
+  "conv.N i = Simple_Network_Language.conv_A (N i)" if "i < n_ps"
+  using that unfolding n_ps_def Prod_TA_Defs.N_def fst_conv snd_conv by (subst nth_map | simp)+
+
+private lemma 5:
+  "inv (conv.N i) = conv_cc o (inv (N i))" if "i < n_ps"
+  unfolding conv_N_eq[OF that] unfolding Simple_Network_Language.conv_A_def
+  by (simp split: prod.split add: inv_def)
+
+lemma trans_conv_N_eq:
+  "trans (conv.N i) = Simple_Network_Language.conv_t ` (trans (N i))"  if "i < n_ps"
+  unfolding conv_N_eq[OF that] unfolding Simple_Network_Language.conv_A_def
+  by (simp split: prod.split add: trans_def)
+
+private lemma 71:
+  "(l, b, conv_cc g, a, r, u, l')\<in>Simple_Network_Language.trans (conv.N i)"
+  if "(l, b, g, a, r, u, l')\<in>Simple_Network_Language.trans (N i)" "i < n_ps"
+  using that by (force simp add: trans_conv_N_eq Simple_Network_Language.conv_t_def)
+
+private lemma 72:
+  "(l, b, conv_cc g, a, r, u, l')\<in>Simple_Network_Language.trans (conv.N i)
+\<longleftrightarrow> (l, b, g, a, r, u, l')\<in>Simple_Network_Language.trans (N i)" if "i < n_ps"
+  by (auto simp: trans_conv_N_eq[OF that] Simple_Network_Language.conv_t_def
+           dest: conv_cc_inj intro: image_eqI[rotated])
+
+private lemma 73:
+  "\<exists>g'. g = conv_cc g' \<and> (l, b, g', a, r, u, l')\<in>Simple_Network_Language.trans (N i)"
+  if "(l, b, g, a, r, u, l')\<in>Simple_Network_Language.trans (conv.N i)" "i < n_ps"
+  using that by (force simp: trans_conv_N_eq Simple_Network_Language.conv_t_def)
+
+lemma conv_bounds[simp]:
+  "conv.bounds = bounds"
+  unfolding conv.bounds_def bounds_def by simp
+
+lemma conv_states[simp]:
+  "conv.states = states"
+  unfolding conv.states_def states_def conv_n_ps_eq
+  by (auto simp add: trans_conv_N_eq Simple_Network_Language.conv_t_def) (fastforce, force)
+
+private lemma 9:
+  "committed (conv.N p) = committed (N p)" if \<open>p < n_ps\<close>
+  unfolding conv_N_eq[OF that] unfolding Simple_Network_Language.conv_A_def
+  by (simp split: prod.split add: committed_def)
+
+private lemma 10:
+  "conv.broadcast = set broadcast"
+  unfolding conv.broadcast_def by simp
+
+lemma conv_trans_int:
+  "conv.trans_int = (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` trans_int"
+  unfolding conv.trans_int_def trans_int_def
+  supply [simp] = L_len
+  apply standard
+  subgoal
+    by (clarsimp simp: Simple_Network_Language.conv_t_def conv_n_ps_eq trans_conv_N_eq 9)
+      (intros add: more_intros, solve_triv+)
+  subgoal
+    by (rule subsetI,
+        clarsimp simp: Simple_Network_Language.conv_t_def conv_n_ps_eq trans_conv_N_eq 9[symmetric])
+      ((drule (1) 71)+, intros add: more_intros, solve_triv+)
+  done
+
+lemma conv_trans_bin:
+  "conv.trans_bin = (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` trans_bin"
+  unfolding conv.trans_bin_def trans_bin_def 10 broadcast_def
+  supply [simp] = L_len
+  apply standard
+  subgoal
+    by (clarsimp simp add: Simple_Network_Language.conv_t_def conv_n_ps_eq trans_conv_N_eq 9)
+      (intros add: more_intros, solve_triv+)
+  subgoal
+    by (rule subsetI,
+        clarsimp simp: Simple_Network_Language.conv_t_def conv_n_ps_eq trans_conv_N_eq 9[symmetric])
+      ((drule (1) 71)+, intros add: more_intros, solve_triv+)
+  done
+
+lemma n_ps_rangeD:
+  "p < n_ps" if "p \<in> set ps" "set ps \<subseteq> {0..<n_ps}"
+  using that by auto
+
+lemma maximalD:
+  "(\<forall>g f r l'.
+       (L ! q, b, g, In a', f, r, l')
+       \<notin> (\<lambda>(l, b, g, a, f, r, l').
+              (l, b, map conv_ac g, a, f, r, l')) ` trans (N q)) \<or> \<not> check_bexp s b True" if
+  "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow> (\<forall>b. (\<forall>g f r l'.
+     (L ! q, b, g, In a', f, r, l') \<notin> trans (N q)) \<or> \<not> check_bexp s b True)"
+  "q<n_ps" "q \<notin> set ps" "p \<noteq> q"
+  for b ps p q L a' s using that by fastforce
+
+lemma conv_trans_broad:
+  "conv.trans_broad = (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` trans_broad"
+  unfolding conv.trans_broad_alt_def trans_broad_alt_def
+  apply standard
+   supply [simp] = L_len
+  subgoal
+  proof -
+    have **: "\<exists> gs'. gs = conv_cc o gs' \<and>
+          (\<forall>p\<in>set ps.(L ! p, bs p, gs' p, In a, fs p, rs p, ls' p) \<in> trans (N p))"
+      if assms:
+        "\<forall>p\<in>set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (conv.N p)"
+        "set ps \<subseteq> {0..<n_ps}" "\<forall>p. p \<notin> set ps \<longrightarrow> gs p = []"
+      for L ps bs gs a fs rs ls'
+    proof -
+      have *: "gs p = conv_cc (Hilbert_Choice.inv conv_cc (gs p))" if "p \<in> set ps" for p
+        using that assms by (auto 4 3 simp: f_inv_into_f dest!: 73)
+      show ?thesis
+        apply (inst_existentials "Hilbert_Choice.inv conv_cc o gs")
+        subgoal
+          apply (rule ext)
+          subgoal for p
+            apply (cases "p \<in> set ps")
+            subgoal
+              by (simp, erule *)
+            subgoal
+              using that(3) by (auto intro: injI inv_f_eq conv_ac_inj)
+            done
+          done
+        subgoal
+          using that * by (force dest!: conv_cc_inj 73)
+        done
+    qed
+    have *: "set ps \<subseteq> {0..<n_ps} \<longleftrightarrow> (\<forall>p \<in> set ps. p < n_ps)" for ps
+      by auto
+    have maximalI:
+      "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow> (\<forall>b. (\<forall>g f r l'.
+         (L ! q, b, g, In a', f, r, l') \<notin> trans (N q)) \<or> \<not> check_bexp s b True)" if
+      "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow> (\<forall>b. (\<forall>g f r l'.
+         (L ! q, b, g, In a', f, r, l') \<notin> trans (conv.N q)) \<or> \<not> check_bexp s b True)"
+    for ps p L a' s
+      using that by (blast dest!: 71)
+    show ?thesis
+      apply (rule subsetI)
+      apply (clarsimp simp add: conv_n_ps_eq 9 10 broadcast_def split: prod.split_asm)
+      apply (drule (2) **)
+      apply (drule (1) 73)+
+      apply elims
+      apply (intro image_eqI[rotated] CollectI exI conjI)
+                          apply solve_triv+
+      subgoal premises prems \<comment>\<open>Commited\<close>
+        using prems(2) \<open>set _ \<subseteq> {0..<n_ps}\<close> by (auto dest: n_ps_rangeD simp: 9)
+                        apply (erule maximalI; fail) \<comment>\<open>Maximal set\<close>
+      by solve_triv+ (simp split: prod.split add: map_concat)
+  qed
+  subgoal
+    apply (rule subsetI)
+    apply (clarsimp simp:
+        Simple_Network_Language.conv_t_def
+        conv_n_ps_eq trans_conv_N_eq 9[symmetric] 10 broadcast_def map_concat)
+    apply (drule (1) 71)
+    apply (intros add: more_intros)
+                        apply solve_triv+
+                    apply (subst comp_def; rule 71; fast elim: n_ps_rangeD; fail)
+    subgoal premises prems for L s s' s'' aj p g f r l' gs fs rs ls' ps
+      using prems(3,6) 9 by fastforce
+                  apply (erule maximalD)
+                    apply (solve_triv | blast)+
+    done
+  done
+
+lemma conv_prod_ta:
+  "conv.prod_ta = Normalized_Zone_Semantics_Impl.conv_A prod_ta"
+  unfolding conv.prod_ta_def prod_ta_def
+  unfolding conv.trans_prod_def trans_prod_def
+  unfolding conv.prod_inv_def prod_inv_def
+  unfolding conv.N_def N_def conv_n_ps_eq
+  unfolding conv_A_def
+  apply simp
+  apply (rule conjI)
+  subgoal
+    unfolding image_Un
+    by ((fo_rule arg_cong2)+; rule conv_trans_int conv_trans_bin conv_trans_broad)
+  subgoal \<comment>\<open>Invariant\<close>
+    unfolding conv.N_def N_def
+    by (rule ext) (auto simp: 5 map_concat intro!: map_cong arg_cong[where f = concat])
+  done
+
+end (* Anonymous context for private lemmas *)
+
+paragraph \<open>Fundamentals\<close>
+
+definition "clkp_set' \<equiv>
+    (\<Union> A \<in> set automata. \<Union> g \<in> set (snd (snd (snd A))). collect_clock_pairs (snd g))
+  \<union> (\<Union> A \<in> set automata. \<Union> (l, b, g, _) \<in> set (fst (snd (snd A))). collect_clock_pairs g)"
+
+definition clk_set'  where
+  \<open>clk_set'  =
+  fst ` clkp_set' \<union>
+  (\<Union> A \<in> set automata. \<Union> (_, _, _, _, _, r, _) \<in> set (fst (snd (snd A))). set r)\<close>
+
+lemma (in -) default_map_of_in_listD:
+  "x \<in> \<Union> (set ` snd ` set invs)" if "x \<in> set (default_map_of [] invs l)"
+proof -
+  have "[] \<noteq> default_map_of [] invs l"
+    using that by force
+  then have "default_map_of [] invs l \<in> snd ` set invs"
+    by (metis (no_types) UNIV_I Un_insert_right range_default_map_of[of "[]" "invs"]
+          image_eqI insertE subsetCE sup_bot.right_neutral)
+  with that show ?thesis
+    by blast
+qed
+
+lemma collect_clock_pairs_invsI:
+  "(a, b) \<in> \<Union> ((collect_clock_pairs o snd) ` set invs)"
+  if "(a, b) \<in> collect_clock_pairs (default_map_of [] invs l)"
+  using that unfolding collect_clock_pairs_def by (auto dest!: default_map_of_in_listD)
+
+lemma mem_trans_N_iff:
+  "t \<in> Simple_Network_Language.trans (N i) \<longleftrightarrow> t \<in> set (fst (snd (snd (automata ! i))))"
+  if "i < n_ps"
+  unfolding N_eq[OF that] by (auto split: prod.splits simp: automaton_of_def trans_def)
+
+lemma length_automata_eq_n_ps:
+  "length automata = n_ps"
+  unfolding n_ps_def by simp
+
+lemma clkp_set'_subs:
+  "Timed_Automata.clkp_set prod_ta \<subseteq> clkp_set'"
+  unfolding Timed_Automata.clkp_set_def clkp_set'_def
+proof (rule union_subsetI, goal_cases)
+  case 1
+  have *: False
+    if "i < n_ps" "L \<in> states"
+      "(a, b) \<in> collect_clock_pairs (Simple_Network_Language.inv (N i) (L ! i))"
+      "\<forall>x\<in>set automata. \<forall>x\<in>set (snd (snd (snd x))). (a, b) \<notin> collect_clock_pairs (snd x)"
+    for a :: 'c and b :: int and L :: "'s list" and i :: nat
+    using that
+    apply (subst (asm) N_eq)
+    apply assumption
+    apply (inst_existentials "automata ! i")
+    unfolding automaton_of_def
+    by (force dest!: nth_mem collect_clock_pairs_invsI
+        split: prod.split_asm simp: inv_def Prod_TA_Defs.n_ps_def)
+  from 1 show ?case
+    unfolding Timed_Automata.collect_clki_def inv_of_prod prod_inv_def
+    by (auto simp: collect_clock_pairs_concat intro: *)
+next
+  case 2
+  then show ?case
+    apply simp
+    unfolding trans_prod_def Timed_Automata.collect_clkt_def
+    apply safe
+    subgoal
+      unfolding trans_int_def by (fastforce simp: length_automata_eq_n_ps mem_trans_N_iff)
+    subgoal
+      unfolding trans_bin_def
+      by (fastforce
+          simp: length_automata_eq_n_ps mem_trans_N_iff
+          dest!: collect_clock_pairs_append_cases)
+    subgoal
+      unfolding trans_broad_def
+      apply (clarsimp simp: length_automata_eq_n_ps mem_trans_N_iff)
+      apply (drule collect_clock_pairs_append_cases)
+      unfolding collect_clock_pairs_concat
+      apply (clarsimp; safe)
+      by (fastforce simp: length_automata_eq_n_ps mem_trans_N_iff)+ \<comment> \<open>slow: 20s\<close>
+    done
+qed
+
+lemma collect_clkvt_subs:
+  "collect_clkvt (trans_of prod_ta) \<subseteq>
+    (\<Union> A \<in> set automata. \<Union> (_, _, _, _, _, r, _) \<in> set (fst (snd (snd A))). set r)"
+  apply simp
+  unfolding collect_clkvt_def
+  apply safe
+  unfolding trans_prod_def
+  subgoal
+    apply simp
+    unfolding trans_prod_def Timed_Automata.collect_clkt_def
+    apply safe
+    subgoal
+      unfolding trans_int_def
+      by (fastforce
+          simp: length_automata_eq_n_ps mem_trans_N_iff
+          dest!: collect_clock_pairs_append_cases)
+    subgoal
+      unfolding trans_bin_def
+      by (fastforce
+          simp: length_automata_eq_n_ps mem_trans_N_iff
+          dest!: collect_clock_pairs_append_cases)
+    subgoal
+      unfolding trans_broad_def
+      apply (clarsimp simp: length_automata_eq_n_ps mem_trans_N_iff)
+      unfolding collect_clock_pairs_concat
+      apply safe
+      by (fastforce simp: length_automata_eq_n_ps mem_trans_N_iff)+ \<comment> \<open>slow: 20s\<close>
+    done
+  done
+
+lemma clk_set'_subs: "clk_set prod_ta \<subseteq> clk_set'"
+  using collect_clkvt_subs clkp_set'_subs unfolding clk_set'_def by auto
+
+end (* Simple Network Impl *)
+
+
+lemma (in Prod_TA_Defs) finite_range_invI:
+  "finite (range prod_inv)" if assms: "\<forall> i < n_ps. finite (range (inv (N i)))"
+proof -
+  let ?N = "\<Union> (range ` inv ` N ` {0..<n_ps})"
+  let ?X = "{I. set I \<subseteq> ?N \<and> length I \<le> n_ps}"
+  have "finite ?N"
+    using assms by auto
+  then have "finite ?X"
+    by (rule finite_lists_length_le)
+  moreover have "range prod_inv \<subseteq> concat ` ?X \<union> {[]}"
+  proof
+    fix x assume "x \<in> range prod_inv"
+    then consider L where "x = concat (map (\<lambda>p. (inv (N p)) (L ! p)) [0..<n_ps])" | "x = []"
+      unfolding prod_inv_def by (auto split: if_split_asm)
+    then show "x \<in> concat ` ?X \<union> {[]}"
+      by (cases; force)
+  qed
+  ultimately show ?thesis by - (drule finite_subset; auto)
+qed
+
+definition (in Prod_TA_Defs)
+  "loc_set =
+  (\<Union> {fst ` trans (N p) | p. p < n_ps} \<union>
+        \<Union> {(snd o snd o snd \<circ> snd \<circ> snd \<circ> snd) ` trans (N p) | p. p < n_ps})"
+
+lemma (in Prod_TA_Defs) states_loc_set:
+  "states \<subseteq> {L. set L \<subseteq> loc_set \<and> length L = n_ps}"
+  unfolding states_def loc_set_def
+  apply (intros add: more_intros)
+   apply (elims add: more_elims)
+   apply (drule mem_nth)
+   apply simp
+   apply (elims add: allE, assumption)
+   apply (simp split: prod.split_asm)
+   apply (erule disjE; (intros add: disjI1 disjI2 more_intros, solve_triv+); fail)
+  by (elims add: more_elims)
+
+lemma (in Prod_TA_Defs) finite_states:
+  assumes finite_trans: "\<forall>p < n_ps. finite (Simple_Network_Language.trans (N p))"
+  shows "finite states"
+proof -
+  have "states \<subseteq> {L. set L \<subseteq> loc_set \<and> length L = n_ps}"
+    by (rule states_loc_set)
+  also from finite_trans have "finite \<dots>"
+    unfolding loc_set_def by (intro finite_intros) auto
+  finally show ?thesis .
+qed
+
+context Simple_Network_Impl
+begin
+
+lemma trans_N_finite:
+  assumes "p < n_ps"
+  shows "finite (Simple_Network_Language.trans (N p))"
+  using assms by (subst N_eq) (auto simp: automaton_of_def trans_def split: prod.split)
+
+lemma states_finite:
+  "finite states"
+  by (intros add: finite_states trans_N_finite)
+
+lemma bounded_finite:
+  "finite {s. bounded bounds s}"
+proof -
+  let ?l = "Min {fst (the (bounds x)) | x. x \<in> dom bounds}"
+  let ?u = "Max {snd (the (bounds x)) | x. x \<in> dom bounds}"
+  have "finite (dom bounds)"
+    unfolding bounds_def by simp
+  then have "{s. bounded bounds s} \<subseteq> {s. dom s = dom bounds \<and> ran s \<subseteq> {?l..?u}}"
+    unfolding bounded_def
+    apply clarsimp
+    apply (rule conjI)
+    subgoal for s v
+      unfolding ran_is_image
+      apply clarsimp
+      subgoal for x l u
+        by (rule order.trans[where b = "fst (the (bounds x))"]; (rule Min_le)?; force)
+      done
+    subgoal for s v
+      unfolding ran_is_image
+      apply clarsimp
+      subgoal for x l u
+        by (rule order.trans[where b = "snd (the (bounds x))"]; (rule Max_ge)?; force)
+      done
+    done
+  also from \<open>finite (dom bounds)\<close> have "finite \<dots>"
+    by (rule finite_set_of_finite_maps) blast
+  finally show ?thesis .
+qed
+
+lemma trans_prod_finite:
+  "finite trans_prod"
+proof -
+  have "finite trans_int"
+  proof -
+    have "trans_int \<subseteq>
+      {((L, s), g, Internal a, r, (L', s')) | L s p l b g a f r l' s' L'.
+        L \<in> states \<and> bounded bounds s \<and> p < n_ps \<and>
+        (l, b, g, Sil a, f, r, l') \<in> trans (N p) \<and>
+        bounded bounds s'
+        \<and> L' = L[p := l']
+      }"
+      unfolding trans_int_def by (force simp: L_len)
+    also have "finite \<dots>"
+    proof -
+      have "finite {(a, b, c, d, e, f, g). (a, b, c, Sil d, e, f, g) \<in> trans (N p)}"
+        if "p < n_ps" for p
+        using [[simproc add: finite_Collect]] that
+        by (auto intro: trans_N_finite finite_vimageI injI)
+      with states_finite bounded_finite show ?thesis
+        by defer_ex
+    qed
+    finally show ?thesis .
+  qed
+  moreover have "finite trans_bin"
+  proof -
+    have "trans_bin \<subseteq>
+      {((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+        L s p q l1 b1 g1 a f1 r1 l1' l2 b2 g2 f2 r2 l2' s'' L'.
+          L \<in> states \<and> bounded bounds s \<and>
+          p < n_ps \<and> q < n_ps \<and>
+          (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+          (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+          bounded bounds s'' \<and>
+          L' = L[p := l1', q := l2']
+    }"
+      unfolding trans_bin_def by (fastforce simp: L_len) (* slow *)
+    also have "finite \<dots>"
+    proof -
+      have "finite {(a, b, c, d, e, f, g). (a, b, c, In d, e, f, g) \<in> trans (N p)}"
+        if "p < n_ps" for p
+        using [[simproc add: finite_Collect]] that
+        by (auto intro: trans_N_finite finite_vimageI injI)
+      moreover have "finite {(a, b, c, e, f, g). (a, b, c, Out d, e, f, g) \<in> trans (N p)}"
+        if "p < n_ps" for p d
+        using [[simproc add: finite_Collect]] that
+        by (auto intro: trans_N_finite finite_vimageI injI)
+      ultimately show ?thesis
+        using states_finite bounded_finite by defer_ex
+    qed
+    finally show ?thesis .
+  qed
+  moreover have "finite trans_broad"
+  proof -
+    define P where "P ps \<equiv> set ps \<subseteq> {0..<n_ps} \<and> distinct ps" for ps
+    define Q where "Q a n bs gs fs rs \<equiv>
+      (\<forall>p < n. \<exists> q < n_ps. \<exists> l l'. (l, bs ! p, gs ! p, In a, fs ! p, rs ! p, l') \<in> trans (N q)) \<and>
+              length bs = n \<and> length gs = n \<and> length fs = n \<and> length rs = n" for a n bs gs fs rs
+    define tag where "tag x = True" for x :: 's
+    have Q_I: "Q a (length ps) (map bs ps) (map gs ps) (map fs ps) (map rs ps)"
+      if "set ps \<subseteq> {0..<n_ps}"
+         "\<forall>p\<in>set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N p)"
+      for ps :: "nat list" and L a bs gs fs rs ls'
+      using that unfolding Q_def by (auto 4 4 dest!: nth_mem)
+    have "trans_broad \<subseteq>
+      {((L, s), g @ concat gs, Broad a, r @ concat rs, (L', s'')) |
+      L s a p l b g f r l' ps bs gs fs rs L' s''.
+        L \<in> states \<and> bounded bounds s \<and> a \<in> set broadcast \<and>
+        p < n_ps \<and>
+        (l, b, g, Out a, f, r, l') \<in> trans (N p) \<and>
+        P ps \<and>
+        Q a (length ps) bs gs fs rs \<and>
+        L' \<in> states \<and>
+        bounded bounds s'' \<and>
+        tag l'
+    }"
+      unfolding trans_broad_def broadcast_def
+      apply (rule subsetI)
+      apply (elims add: more_elims)
+      apply (intros add: more_intros)
+                apply solve_triv+
+            apply (simp add: L_len; fail)
+           apply assumption
+          apply (unfold P_def; intros; assumption)
+         apply (rule Q_I; assumption)
+      subgoal
+        by (blast intro: state_preservation_updI state_preservation_fold_updI)
+       apply assumption
+      unfolding tag_def ..
+    also have "finite \<dots>"
+    proof -
+      have "finite {(a, b, c, e, f, g). (a, b, c, Out d, e, f, g) \<in> trans (N p)}"
+        if "p < n_ps" for p d
+        using [[simproc add: finite_Collect]] that
+        by (auto intro: trans_N_finite finite_vimageI injI)
+      moreover have "finite {ps. P ps}"
+        unfolding P_def by (simp add: finite_intros)
+      moreover have "finite {(bs, gs, fs, rs). Q a n bs gs fs rs}" (is "finite ?S") for a n
+      proof -
+        let ?T = "\<Union> (trans ` N ` {0..<n_ps})"
+        have "?S \<subseteq> {(bs, gs, fs, rs).
+          (set bs \<subseteq> (\<lambda>(_,b,_). b) ` ?T \<and> length bs = n) \<and>
+          (set gs \<subseteq> (\<lambda>(_,_,g,_). g) ` ?T \<and> length gs = n) \<and>
+          (set fs \<subseteq> (\<lambda>(_,_,_,_,f,_). f) ` ?T \<and> length fs = n) \<and>
+          (set rs \<subseteq> (\<lambda>(_,_,_,_,_,r,_). r) ` ?T \<and> length rs = n)
+        }"
+          unfolding Q_def
+          apply safe
+          apply (all \<open>drule mem_nth; elims; drule spec; elims\<close>)
+          apply force+
+          done
+        also have "finite \<dots>"
+          using trans_N_finite by (intro finite_intros more_finite_intros) auto
+        finally show ?thesis .
+      qed
+      ultimately show ?thesis
+        using states_finite bounded_finite by defer_ex
+    qed
+    finally show ?thesis .
+  qed
+  ultimately show ?thesis
+    by (simp add: trans_prod_def)
+qed
+
+lemma prod_inv_finite:
+  "finite (range prod_inv)"
+  apply (intros add: finite_range_invI)
+  unfolding length_automata_eq_n_ps[symmetric]
+  unfolding N_def
+  unfolding automaton_of_def
+  by (auto intro: finite_range_default_map_of split: prod.split simp: inv_def)
+
+end (* Simple Network Impl *)
+
+paragraph \<open>Collecting variables from expressions.\<close>
+
+fun vars_of_bexp and vars_of_exp where
+  "vars_of_bexp (not e) = vars_of_bexp e"
+| "vars_of_bexp (and e1 e2) = (vars_of_bexp e1 \<union> vars_of_bexp e2)"
+| "vars_of_bexp (bexp.or e1 e2) = (vars_of_bexp e1 \<union> vars_of_bexp e2)"
+| "vars_of_bexp (imply e1 e2) = (vars_of_bexp e1 \<union> vars_of_bexp e2)"
+| "vars_of_bexp (eq i x) = vars_of_exp i \<union> vars_of_exp x"
+| "vars_of_bexp (le i x) = vars_of_exp i \<union> vars_of_exp x"
+| "vars_of_bexp (lt i x) = vars_of_exp i \<union> vars_of_exp x"
+| "vars_of_bexp (ge i x) = vars_of_exp i \<union> vars_of_exp x"
+| "vars_of_bexp (gt i x) = vars_of_exp i \<union> vars_of_exp x"
+| "vars_of_bexp bexp.true = {}"
+| "vars_of_exp (const c) = {}"
+| "vars_of_exp (var x) = {x}"
+| "vars_of_exp (if_then_else b e1 e2) = vars_of_bexp b \<union> vars_of_exp e1 \<union> vars_of_exp e2"
+| "vars_of_exp (binop _ e1 e2) = vars_of_exp e1 \<union> vars_of_exp e2"
+| "vars_of_exp (unop _ e) = vars_of_exp e"
+
+definition (in Prod_TA_Defs)
+  "var_set =
+  (\<Union>S \<in> {(fst \<circ> snd) ` trans (N p) | p. p < n_ps}. \<Union>b \<in> S. vars_of_bexp b) \<union>
+  (\<Union>S \<in> {(fst \<circ> snd \<circ> snd \<circ> snd \<circ> snd) ` trans (N p) | p. p < n_ps}.
+    \<Union>f \<in> S. \<Union> (x, e) \<in> set f. {x} \<union> vars_of_exp e)"
+
+locale Simple_Network_Impl_nat_defs =
+  Simple_Network_Impl automata
+  for automata ::
+    "(nat list \<times> nat list \<times> (nat act, nat, nat, int, nat, int) transition list
+      \<times> (nat \<times> (nat, int) cconstraint) list) list" +
+  fixes m :: nat and num_states :: "nat \<Rightarrow> nat" and num_actions :: nat
+
+locale Simple_Network_Impl_nat =
+  Simple_Network_Impl_nat_defs +
+  assumes has_clock: "m > 0"
+  assumes non_empty: "0 < length automata"
+    (* assumes "length automata = length state_nums" *)
+  assumes trans_num_states:
+    "\<forall>i < n_ps. let (_, _, trans, _) = (automata ! i) in \<forall> (l, _, _, _, _, _, l') \<in> set trans.
+      l < num_states i \<and> l' < num_states i"
+    and inv_num_states:
+    "\<forall>i < n_ps. let (_, _, _, inv) = (automata ! i) in \<forall> (x, _) \<in> set inv. x < num_states i"
+  assumes var_set:
+    "\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, _, _, f, _, _) \<in> set trans.
+      \<forall>(x, upd) \<in> set f. x < n_vs \<and> (\<forall>i \<in> vars_of_exp upd. i < n_vs)"
+    "\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, b, _, _, _, _, _) \<in> set trans.
+      \<forall>i \<in> vars_of_bexp b. i < n_vs"
+  assumes bounds:
+    "\<forall> i < n_vs. fst (bounds' ! i) = i"
+  assumes action_set:
+    "\<forall>a \<in> set broadcast. a < num_actions"
+    "\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, _, a, _, _, _) \<in> set trans.
+        pred_act (\<lambda>a. a < num_actions) a"
+  assumes clock_set:
+    "\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, g, _, _, r, _) \<in> set trans.
+      (\<forall>c \<in> set r. 0 < c \<and> c \<le> m) \<and>
+      (\<forall> (c, x) \<in> collect_clock_pairs g. 0 < c \<and> c \<le> m \<and> x \<in> \<nat>)
+      "
+    "\<forall>(_, _, _, inv) \<in> set automata. \<forall>(l, g) \<in> set inv.
+      (\<forall> (c, x) \<in> collect_clock_pairs g. 0 < c \<and> c \<le> m \<and> x \<in> \<nat>)
+      "
+  assumes broadcast_receivers:
+  "\<forall>(_, _, trans, _) \<in> set automata. \<forall>(_, _, g, a, _, _, _) \<in> set trans.
+      case a of In a \<Rightarrow> a \<in> set broadcast \<longrightarrow> g = [] | _ \<Rightarrow> True"
+begin
+
+lemma broadcast_receivers_unguarded:
+  "\<forall>p<n_ps. \<forall>l b g a f r l'.
+    (l, b, g, In a, f, r, l') \<in> Simple_Network_Language.trans (N p) \<and> a \<in> set broadcast \<longrightarrow> g = []"
+  using broadcast_receivers by (fastforce dest: nth_mem simp: n_ps_def mem_trans_N_iff)
+
+sublocale conv: Prod_TA
+  "(set broadcast, map (Simple_Network_Language.conv_A o automaton_of) automata, map_of bounds')"
+  using broadcast_receivers_unguarded
+  by - (standard,
+ fastforce simp: conv.broadcast_def Simple_Network_Language.conv_t_def conv_n_ps_eq trans_conv_N_eq)
+
+sublocale TA_Start_No_Ceiling prod_ta init m
+proof standard
+  show "finite (trans_of prod_ta)"
+    using trans_prod_finite by simp
+next
+  show "finite (range (inv_of prod_ta))"
+    using prod_inv_finite by simp
+next
+  from clk_set'_subs have "clk_set prod_ta \<subseteq> clk_set'" .
+  also have "\<dots> \<subseteq> {1..m}"
+    using clock_set unfolding clk_set'_def clkp_set'_def by force
+  finally show "clk_set prod_ta \<subseteq> {1..m}" .
+next
+  from clock_set have "\<forall>(_, d)\<in>clkp_set'. d \<in> \<nat>"
+    unfolding clkp_set'_def by force
+  then show "\<forall>(_, d)\<in>Timed_Automata.clkp_set prod_ta. d \<in> \<nat>"
+    by (auto dest!: subsetD[OF clkp_set'_subs])
+next
+  show "0 < m"
+    by (rule has_clock)
+qed
+
+end (* Simple Network Impl nat *)
+
+
+context Simple_Network_Impl
+begin
+
+definition "sem \<equiv> (set broadcast, map (automaton_of o conv_automaton) automata, map_of bounds')"
+
+sublocale sem: Prod_TA_sem sem .
+
+lemma sem_N_eq:
+  "sem.N p = automaton_of (conv_automaton (automata ! p))" if \<open>p < n_ps\<close>
+  using that unfolding sem.N_def n_ps_def unfolding sem_def fst_conv snd_conv
+  by (subst nth_map) auto
+
+end (* Simple Network Impl *)
+
+inductive_cases step_u_elims:
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Internal a \<langle>L', s', u'\<rangle>"
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Bin a \<langle>L', s'', u'\<rangle>"
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Broad a \<langle>L', s'', u'\<rangle>"
+
+inductive_cases step_u_elims':
+  "(broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  "(broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Internal a \<langle>L', s', u'\<rangle>"
+  "(broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Bin a \<langle>L', s'', u'\<rangle>"
+  "(broadcast, N, B) \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Broad a \<langle>L', s'', u'\<rangle>"
+
+lemma (in Prod_TA_Defs) states_lengthD:
+  "length L = n_ps" if "L \<in> states"
+  using that unfolding states_def by simp
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Impl_Refine.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Impl_Refine.thy
@@ -0,0 +1,3438 @@
+theory Simple_Network_Language_Impl_Refine
+  imports Simple_Network_Language_Impl
+begin
+
+unbundle no_library_syntax
+notation fun_rel_syn (infixr "\<rightarrow>" 60)
+hide_const (open) upd
+
+paragraph \<open>Expression evaluation\<close>
+
+fun bval :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 'b :: linorder) bexp \<Rightarrow> bool" and eval where
+  "bval _ bexp.true \<longleftrightarrow> True" |
+  "bval s (not e) \<longleftrightarrow> \<not> bval s e" |
+  "bval s (and e1 e2) \<longleftrightarrow> bval s e1 \<and> bval s e2" |
+  "bval s (bexp.or e1 e2) \<longleftrightarrow> bval s e1 \<or> bval s e2" |
+  "bval s (imply e1 e2) \<longleftrightarrow> bval s e1 \<longrightarrow> bval s e2" |
+  "bval s (eq i x) \<longleftrightarrow> eval s i = eval s x" |
+  "bval s (le i x) \<longleftrightarrow> eval s i \<le> eval s x" |
+  "bval s (lt i x) \<longleftrightarrow> eval s i < eval s x" |
+  "bval s (ge i x) \<longleftrightarrow> eval s i \<ge> eval s x" |
+  "bval s (gt i x) \<longleftrightarrow> eval s i > eval s x"
+| "eval s (const c) = c"
+| "eval s (var x)   = s x"
+| "eval s (if_then_else b e1 e2) = (if bval s b then eval s e1 else eval s e2)"
+| "eval s (binop f e1 e2) = f (eval s e1) (eval s e2)"
+| "eval s (unop f e) = f (eval s e)"
+
+lemma check_bexp_determ:
+  "check_bexp s b b1 \<Longrightarrow> check_bexp s b b2 \<Longrightarrow> b1 = b2"
+  and is_val_determ: "is_val s e v1 \<Longrightarrow> is_val s e v2 \<Longrightarrow> v1 = v2"
+  by (induction b and e arbitrary: b1 b2 and v1 v2)
+     (auto dest:  elim!: is_val_elims check_bexp_elims)+
+
+lemma is_upd_determ:
+  "s1 = s2" if "is_upd s f s1" "is_upd s f s2"
+  using that unfolding is_upd_def
+  by (clarsimp, fo_rule arg_cong)
+     (smt
+      case_prodE case_prodE' case_prod_conv is_val_determ list.rel_eq list_all2_swap list_all2_trans
+     )
+
+lemma is_upds_determ:
+  "s1 = s2" if "is_upds s fs s1" "is_upds s fs s2"
+  using that
+  by (induction fs arbitrary: s) (auto 4 4 elim: is_upds.cases dest: is_upd_determ)
+
+lemma check_bexp_bval:
+  "\<forall>x \<in> vars_of_bexp b. x \<in> dom s \<Longrightarrow> check_bexp s b (bval (the o s) b)"
+and is_val_eval:
+  "\<forall>x \<in> vars_of_exp e. x \<in> dom s \<Longrightarrow> is_val s e (eval (the o s) e)"
+  apply (induction b and e)
+                apply (simp; (subst eq_commute)?; rule check_bexp_is_val.intros; simp add: dom_def)+
+    apply ((subst eq_commute eval.simps)?; rule check_bexp_is_val.intros; simp add: dom_def)+
+  done
+
+lemma is_upd_dom:
+  "dom s' = dom s" if "is_upd s (x, e) s'" "x \<in> dom s"
+  using that unfolding is_upd_def by (auto split: if_split_asm)
+
+lemma is_upds_dom:
+  "dom s' = dom s" if "is_upds s upds s'" "\<forall>(x, e) \<in> set upds. x \<in> dom s"
+  using that
+  by (induction upds arbitrary: s)
+     (erule is_upds.cases; auto simp: is_upd_dom split: prod.split_asm)+
+
+definition
+  "mk_upd \<equiv> \<lambda>(x, e) s. s(x \<mapsto> eval (the o s) e)"
+
+definition mk_upds ::
+  "('a \<Rightarrow> ('b :: linorder) option) \<Rightarrow> ('a \<times> ('a, 'b) exp) list \<Rightarrow> ('a \<Rightarrow> 'b option)" where
+  "mk_upds s upds = fold mk_upd upds s"
+
+lemma is_upd_make_updI:
+  "is_upd s upd (mk_upd upd s)" if "upd = (x, e)" "\<forall>x \<in> vars_of_exp e. x \<in> dom s"
+  using that(1) is_val_eval[OF that(2)] unfolding is_upd_def mk_upd_def by auto
+
+lemma dom_mk_upd:
+  "dom s \<subseteq> dom (mk_upd upd s)"
+  unfolding mk_upd_def by (auto split: prod.split)
+
+lemma is_upds_make_updsI:
+  "is_upds s upds (mk_upds s upds)" if "\<forall> (_, e) \<in> set upds. \<forall>x \<in> vars_of_exp e. x \<in> dom s"
+  using that
+proof (induction upds arbitrary: s)
+  case Nil
+  then show ?case
+    by (auto intro!: is_upds.intros simp: mk_upds_def)
+next
+  case (Cons upd upds)
+  let ?s = "mk_upd upd s"
+  from Cons(2) have "is_upd s upd ?s"
+    by (auto simp: dom_def intro: is_upd_make_updI)
+  moreover have "is_upds ?s upds (mk_upds ?s upds)"
+    using Cons.prems dom_mk_upd by (intro Cons.IH) force
+  ultimately show ?case
+    using dom_mk_upd by (auto intro!: is_upds.intros simp: mk_upds_def)
+qed
+
+
+paragraph \<open>Implementation auxiliaries\<close>
+
+definition
+  "union_map_of xs =
+    fold (\<lambda> (x, y) m. case m x of None \<Rightarrow> m(x \<mapsto> [y]) | Some ys \<Rightarrow> m(x \<mapsto> y # ys)) xs Map.empty"
+
+lemma union_map_of_alt_def:
+  "union_map_of xs x = (let
+    ys = rev (map snd (filter (\<lambda> (x', y). x' = x) xs))
+   in if ys = [] then None else Some ys)"
+proof -
+  have "fold (\<lambda> (x, y) m. case m x of None \<Rightarrow> m(x \<mapsto> [y]) | Some ys \<Rightarrow> m(x \<mapsto> y # ys)) xs m x
+  = (let
+      ys = rev (map snd (filter (\<lambda> (x', y). x' = x) xs))
+    in
+    case m x of None \<Rightarrow> if ys = [] then None else Some ys | Some zs \<Rightarrow> Some (ys @ zs))" for x m
+    by (induction xs arbitrary: m; clarsimp split: option.split)
+  then show ?thesis
+    unfolding union_map_of_def by simp
+qed
+
+lemma in_union_map_ofI:
+  "\<exists> ys. union_map_of xs x = Some ys \<and> y \<in> set ys" if "(x, y) \<in> set xs"
+  unfolding union_map_of_alt_def Let_def using that  set_filter[of "\<lambda>(x', y). x' = x" xs] by auto+
+
+lemma in_union_map_ofD:
+  "(x, y) \<in> set xs" if "union_map_of xs x = Some ys" "y \<in> set ys"
+  using that unfolding union_map_of_alt_def by (auto split: if_split_asm)
+
+
+paragraph \<open>Implementation of state set\<close>
+
+context Simple_Network_Impl_nat_defs
+begin
+
+definition
+  "states_i i = (\<Union>(l, e, g, a, r, u, l')\<in>set (fst (snd (snd (automata ! i)))). {l, l'})"
+
+lemma states_mem_compute[code]:
+  "L \<in> states \<longleftrightarrow> length L = n_ps \<and> (\<forall>i<n_ps. L ! i \<in> states_i i)"
+  unfolding states_def states_i_def by simp (metis mem_trans_N_iff)
+
+lemma states_mem_compute':
+  "L \<in> states \<longleftrightarrow> length L = n_ps \<and> (\<forall>i<n_ps. L ! i \<in> map states_i [0..<n_ps] ! i)"
+  unfolding states_mem_compute by simp
+
+end
+
+
+context Simple_Network_Impl_nat
+begin
+
+paragraph \<open>Fundamentals\<close>
+
+lemma mem_trans_N_iff:
+  \<open>t \<in> Simple_Network_Language.trans (N i) \<longleftrightarrow> t \<in> set (fst (snd (snd (automata ! i))))\<close>
+  if "i < n_ps"
+  unfolding N_def fst_conv snd_conv
+  unfolding automaton_of_def
+  unfolding trans_def
+  using that by (cases "automata ! i") (auto simp: length_automata_eq_n_ps)
+
+lemma L_i_len:
+  \<open>L ! i < num_states i\<close> if "i < n_ps" "L \<in> states"
+  using trans_num_states that
+  unfolding states_def by (auto 4 4 simp: mem_trans_N_iff)
+
+lemma L_i_simp:
+  \<open>[0..<num_states i] ! (L ! i) = L ! i\<close>
+  if "i < n_ps" "L \<in> states"
+  using L_i_len[OF that] by simp
+
+lemma action_setD:
+  \<open>pred_act (\<lambda>a'. a' < num_actions) a\<close>
+  if \<open>(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)\<close> \<open>p < n_ps\<close>
+  using that action_set
+  by (cases "automata ! p")
+    (subst (asm) mem_trans_N_iff; force dest!: nth_mem simp flip: length_automata_eq_n_ps)
+
+paragraph \<open>More precise state sets\<close>
+
+definition
+  "states' \<equiv> {(L, s). L \<in> states \<and> dom s = {0..<n_vs} \<and> bounded bounds s}"
+
+lemma states'_states[intro, dest]:
+  "L \<in> states" if "(L, s) \<in> states'"
+  using that unfolding states'_def by auto
+
+lemma states'_dom[intro, dest]:
+  "dom s = {0..<n_vs}" if "(L, s) \<in> states'"
+  using that unfolding states'_def by auto
+
+lemma states'_bounded[intro, dest]:
+  "bounded bounds s" if "(L, s) \<in> states'"
+  using that unfolding states'_def by auto
+
+paragraph \<open>Implementation of invariants\<close>
+
+definition (in Simple_Network_Impl_nat_defs)
+  "invs i \<equiv> let m = default_map_of [] (snd (snd (snd (automata ! i))));
+    m' = map (\<lambda> j. m j) [0..<num_states i]
+  in m'"
+
+definition (in Simple_Network_Impl_nat_defs)
+  "invs1 \<equiv> map (\<lambda> i. let m = default_map_of [] (snd (snd (snd (automata ! i))));
+    m' = map (\<lambda> j. m j) [0..<num_states i]
+  in m') [0..<n_ps]"
+
+definition (in Simple_Network_Impl_nat_defs)
+  "invs2 \<equiv> IArray (map (\<lambda> i. let m = default_map_of [] (snd (snd (snd (automata ! i))));
+    m' = IArray (map (\<lambda> j. m j) [0..<num_states i])
+  in m') [0..<n_ps])"
+
+lemma refine_invs2:
+  "invs2 !! i !! j = invs1 ! i ! j" if "i < n_ps"
+  using that unfolding invs2_def invs1_def by simp
+
+definition (in Simple_Network_Impl_nat_defs)
+  "inv_fun \<equiv> \<lambda>(L, _).
+    concat (map (\<lambda>i. invs1 ! i ! (L ! i)) [0..<n_ps])"
+
+lemma refine_invs1:
+  \<open>invs1 ! i = invs i\<close> if \<open>i < n_ps\<close>
+  using that unfolding invs_def invs1_def by simp
+
+lemma invs_simp:
+  "invs1 ! i ! (L ! i) = Simple_Network_Language.inv (N i) (L ! i)"
+  if "i < n_ps" "L \<in> states"
+  using that unfolding refine_invs1[OF \<open>i < _\<close>] invs_def N_def fst_conv snd_conv
+  unfolding inv_def
+  by (subst nth_map;
+      clarsimp split: prod.split simp: automaton_of_def length_automata_eq_n_ps L_i_len)
+
+lemma inv_fun_inv_of':
+  "(inv_fun, inv_of prod_ta) \<in> inv_rel R states'" if "R \<subseteq> Id \<times>r S"
+  using that
+  unfolding inv_rel_def
+  unfolding inv_fun_def
+  unfolding inv_of_prod prod_inv_def
+  apply (clarsimp simp: states'_def)
+  apply (rule arg_cong[where f = concat])
+  apply (auto simp add: invs_simp cong: map_cong)
+  done
+
+lemma inv_fun_alt_def:
+  "inv_fun = (\<lambda>(L, _). concat (map (\<lambda>i. invs2 !! i !! (L ! i)) [0..<n_ps]))"
+  unfolding inv_fun_def
+  apply (intro ext)
+  apply (clarsimp simp del: IArray.sub_def)
+  apply (fo_rule arg_cong)
+  apply (simp add: refine_invs2 del: IArray.sub_def)
+  done
+
+end (* Simple Network Impl nat *)
+
+paragraph \<open>Implementation of transitions\<close>
+
+context Simple_Network_Impl_nat_defs
+begin
+
+definition
+  "bounds_map \<equiv> the o map_of bounds'"
+
+definition
+  "check_bounded s =
+    (\<forall>x \<in> dom s. fst (bounds_map x) \<le> the (s x) \<and> the (s x) \<le> snd (bounds_map x))"
+
+text \<open>Compute pairs of processes with committed initial locations from location vector.\<close>
+definition
+  "get_committed L =
+    List.map_filter (\<lambda>p.
+    let l = L ! p in
+    if l \<in> set (fst (automata ! p)) then Some (p, l) else None) [0..<n_ps]"
+
+text \<open>Given a process and a location, return the corresponding transitions.\<close>
+definition
+  "trans_map i \<equiv>
+    let m = union_map_of (fst (snd (snd (automata ! i)))) in (\<lambda>j.
+      case m j of None \<Rightarrow> [] | Some xs \<Rightarrow> xs)"
+
+text \<open>Filter for internal transitions.\<close>
+definition
+  "trans_i_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (b, g, a, m, l'). case a of Sil a \<Rightarrow> Some (b, g, a, m, l') | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+text \<open>Compute valid internal successors given:
+  item a process \<open>p\<close>,
+  item initial location \<open>l\<close>,
+  item location vector \<open>L\<close>,
+  item and initial state \<open>s\<close>.
+\<close>
+definition
+  "int_trans_from_loc p l L s \<equiv>
+    let trans = trans_i_map p l
+    in
+    List.map_filter (\<lambda> (b, g, a, f, r, l').
+      let s' = mk_upds s f in
+        if bval (the o s) b \<and> check_bounded s' then Some (g, Internal a, r, (L[p := l'], s'))
+        else None
+    ) trans"
+
+
+definition
+  "int_trans_from_vec pairs L s \<equiv>
+    concat (map (\<lambda>(p, l). int_trans_from_loc p l L s) pairs)"
+
+definition
+  "int_trans_from_all L s \<equiv>
+    concat (map (\<lambda>p. int_trans_from_loc p (L ! p) L s) [0..<n_ps])"
+
+definition
+  "int_trans_from \<equiv> \<lambda> (L, s).
+    let pairs = get_committed L in
+    if pairs = []
+    then int_trans_from_all L s
+    else int_trans_from_vec pairs L s
+  "
+
+definition
+  "actions_by_state i \<equiv>
+   fold (\<lambda> t acc. acc[fst (snd (snd t)) := (i, t) # (acc ! fst (snd (snd t)))])"
+
+definition
+  "all_actions_by_state t L \<equiv>
+    fold (\<lambda> i. actions_by_state i (t i (L ! i))) [0..<n_ps] (repeat [] num_actions)"
+
+definition
+  "all_actions_from_vec t vec \<equiv>
+    fold (\<lambda>(p, l). actions_by_state p (t p l)) vec (repeat [] num_actions)"
+
+
+definition
+  "pairs_by_action L s OUT IN \<equiv>
+  concat (
+    map (\<lambda> (p, b1, g1, a1, f1, r1, l1).
+      List.map_filter (\<lambda> (q, b2, g2, a2, f2, r2, l2).
+        if p = q then None else
+          let s' = mk_upds (mk_upds s f1) f2 in
+          if bval (the o s) b1 \<and> bval (the o s) b2 \<and> check_bounded s'
+          then Some (g1 @ g2, Bin a1, r1 @ r2, (L[p := l1, q := l2], s'))
+          else None
+    ) OUT) IN)
+  "
+
+definition
+  "trans_in_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (b, g, a, m, l'). case a of In a \<Rightarrow> Some (b, g, a, m, l') | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+definition
+  "trans_out_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (b, g, a, m, l'). case a of Out a \<Rightarrow> Some (b, g, a, m, l') | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+definition
+  "bin_actions = filter (\<lambda>a. a \<notin> set broadcast) [0..<num_actions]"
+
+lemma mem_bin_actions_iff:
+  "a \<in> set bin_actions \<longleftrightarrow> a \<notin> local.broadcast \<and> a < num_actions "
+  unfolding bin_actions_def broadcast_def by auto
+
+definition
+  "bin_trans_from \<equiv> \<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In =  all_actions_by_state trans_in_map L;
+      Out = all_actions_by_state trans_out_map L
+    in
+    if pairs = [] then
+      concat (map (\<lambda>a. pairs_by_action L s (Out ! a) (In ! a)) bin_actions)
+    else
+      let
+        In2  = all_actions_from_vec trans_in_map pairs;
+        Out2 = all_actions_from_vec trans_out_map pairs
+      in
+        concat (map (\<lambda>a. pairs_by_action L s (Out ! a) (In2 ! a)) bin_actions)
+      @ concat (map (\<lambda>a. pairs_by_action L s (Out2 ! a) (In ! a)) bin_actions)
+    "
+
+definition
+  "trans_in_broad_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (b, g, a, m, l').
+      case a of In a \<Rightarrow> if a \<in> set broadcast then Some (b, g, a, m, l') else None | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+definition
+  "trans_out_broad_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (b, g, a, m, l').
+      case a of Out a \<Rightarrow> if a \<in> set broadcast then Some (b, g, a, m, l') else None | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+definition
+  "actions_by_state' xs \<equiv>
+    fold (\<lambda> t acc. acc[fst (snd (snd t)) := t # (acc ! fst (snd (snd t)))])
+      xs (repeat [] num_actions)"
+
+definition
+  "trans_out_broad_grouped i j \<equiv> actions_by_state' (trans_out_broad_map i j)"
+
+definition
+  "trans_in_broad_grouped i j \<equiv> actions_by_state' (trans_in_broad_map i j)"
+
+definition
+  "pair_by_action OUT IN \<equiv>
+  concat (
+    map (\<lambda>(g1, a, r1, (L, s)).
+      List.map (\<lambda>(q, g2, a2, f2, r2, l2).
+          (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+    ) OUT) IN)
+  "
+
+definition make_combs where
+  "make_combs p a xs \<equiv>
+    let
+      ys = List.map_filter
+        (\<lambda>i.
+          if i = p then None
+          else if xs ! i ! a = [] then None
+          else Some (map (\<lambda>t. (i, t)) (xs ! i ! a))
+        )
+        [0..<n_ps]
+    in if ys = [] then [] else product_lists ys
+  "
+
+definition make_combs_from_pairs where
+  "make_combs_from_pairs p a pairs xs \<equiv>
+    let
+      ys = List.map_filter
+        (\<lambda>i.
+          if i = p then None
+          else if xs ! i ! a = [] then None
+          else Some (map (\<lambda>t. (i, t)) (xs ! i ! a))
+        )
+        [0..<n_ps]
+    in if ys = [] then [] else product_lists ys
+  "
+
+definition
+  "broad_trans_from \<equiv> \<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+      Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps];
+      In  = map (map (filter (\<lambda> (b, _). bval (the o s) b))) In;
+      Out = map (map (filter (\<lambda> (b, _). bval (the o s) b))) Out
+    in
+    if pairs = [] then
+      concat (
+        map (\<lambda>a.
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+              List.map_filter (\<lambda>comb.
+                let (g, a, r, L', s) =
+                  fold
+                    (\<lambda>(q, _, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+                      (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+                    )
+                    comb
+                    init
+                in if check_bounded s then Some (g, a, r, L', s) else None
+              ) combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    else
+      concat (
+        map (\<lambda>a.
+          let
+            ins_committed =
+              List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs;
+            always_committed = (length ins_committed > 1)
+          in
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else if
+              \<not> always_committed \<and> (ins_committed = [p] \<or> ins_committed = [])
+              \<and> \<not> list_ex (\<lambda> (q, _). q = p) pairs
+            then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+              List.map_filter (\<lambda>comb.
+                let (g, a, r, L', s) =
+                  fold
+                    (\<lambda>(q, _, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+                      (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+                    )
+                    comb
+                    init
+                in if check_bounded s then Some (g, a, r, L', s) else None
+              ) combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    "
+
+end (* Simple Network Impl nat defs *)
+
+lemma product_lists_empty: "product_lists xss = [] \<longleftrightarrow> (\<exists>xs \<in> set xss. xs = [])" for xss
+  by (induction xss) auto
+
+context Simple_Network_Impl_nat
+begin
+
+lemma broad_trans_from_alt_def:
+  "broad_trans_from \<equiv> \<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+      Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps];
+      In  = map (map (filter (\<lambda> (b, _). bval (the o s) b))) In;
+      Out = map (map (filter (\<lambda> (b, _). bval (the o s) b))) Out
+    in
+    if pairs = [] then
+      concat (
+        map (\<lambda>a.
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+              filter (\<lambda> (g, a, r, L, s). check_bounded s) (
+                map (\<lambda>comb.
+                    fold
+                      (\<lambda>(q, _, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+                        (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+                      )
+                      comb
+                      init
+                ) combs)
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    else
+      concat (
+        map (\<lambda>a.
+          let
+            ins_committed =
+              List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs
+          in
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else if
+              (ins_committed = [p] \<or> ins_committed = []) \<and> \<not> list_ex (\<lambda>(q, _). q = p) pairs
+            then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+              filter (\<lambda> (g, a, r, L, s). check_bounded s) (
+                map (\<lambda>comb.
+                    fold
+                      (\<lambda>(q, _, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+                        (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+                      )
+                      comb
+                      init
+                ) combs)
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    "
+  apply (rule eq_reflection)
+  unfolding broad_trans_from_def
+  unfolding filter_map_map_filter
+  unfolding Let_def
+  by (fo_rule
+      arg_cong2[where f = map] arg_cong2[where f = List.map_filter] arg_cong HOL.refl |
+      rule if_cong ext | auto split: if_split_asm)+
+
+paragraph \<open>Correctness for implementations of primitives\<close>
+
+lemma dom_bounds_eq:
+  "dom bounds = {0..<n_vs}"
+  using bounds unfolding bounds_def
+  apply simp
+  unfolding n_vs_def dom_map_of_conv_image_fst
+  apply safe
+   apply (force dest: mem_nth; fail)
+  apply (force dest: nth_mem; fail)
+  done
+
+lemma check_bounded_iff:
+  "Simple_Network_Language.bounded bounds s \<longleftrightarrow> check_bounded s" if "dom s = {0..<n_vs}"
+  using that
+  unfolding dom_bounds_eq[symmetric]
+  unfolding check_bounded_def Simple_Network_Language.bounded_def bounds_map_def bounds_def
+  by auto
+
+lemma get_committed_mem_iff:
+  "(p, l) \<in> set (get_committed L) \<longleftrightarrow> (l = L ! p \<and> l \<in> committed (N p) \<and> p < n_ps)"
+  unfolding get_committed_def
+  unfolding set_map_filter Let_def
+  apply clarsimp
+  unfolding N_def fst_conv snd_conv
+  unfolding committed_def
+  by safe
+    ((subst nth_map | subst (asm) nth_map);
+      auto split: prod.splits simp: automaton_of_def length_automata_eq_n_ps
+      )+
+
+lemma get_committed_empty_iff:
+  "(\<forall>p < n_ps. L ! p \<notin> committed (N p)) \<longleftrightarrow> get_committed L = []"
+  apply safe
+  subgoal
+  proof (rule ccontr)
+    assume prems:
+      "\<forall>p<n_ps. L ! p \<notin> committed (N p)" and
+      "get_committed L \<noteq> []"
+    then obtain p l where "(p, l) \<in> set (get_committed L)"
+      by (metis length_greater_0_conv nth_mem old.prod.exhaust)
+    from this[unfolded get_committed_mem_iff] prems(1)
+    show "False"
+      by auto
+  qed
+  subgoal for p
+    using get_committed_mem_iff[of p "L ! p" L] by auto
+  done
+
+lemma get_committed_distinct: "distinct (get_committed L)"
+  unfolding get_committed_def by (rule distinct_map_filterI) (auto simp: Let_def)
+
+lemma is_upds_make_updsI2:
+  "is_upds s upds (mk_upds s upds)"
+  if "(l, b, g, a, upds, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+    "dom s = {0..<n_vs}"
+  using that var_set
+  by (intro is_upds_make_updsI, subst (asm) mem_trans_N_iff)
+     (auto 4 5 simp flip: length_automata_eq_n_ps dest!: nth_mem)
+
+lemma var_setD:
+  "\<forall>(x, upd)\<in>set f. x < n_vs \<and> (\<forall>i\<in>vars_of_exp upd. i < n_vs)"
+  if "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  using var_set that
+  by (force dest: nth_mem simp flip: length_automata_eq_n_ps simp: mem_trans_N_iff)+
+
+lemma var_setD2:
+  "\<forall>i\<in>vars_of_bexp b. i < n_vs"
+  if "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  using var_set that
+  by (force dest: nth_mem simp flip: length_automata_eq_n_ps simp: mem_trans_N_iff)+
+
+lemma is_upds_dom2:
+  "dom s' = {0..<n_vs}" if "is_upds s f s'"
+  "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  "dom s = {0..<n_vs}"
+  unfolding that(4)[symmetric] using that by - (drule (1) var_setD, erule is_upds_dom, auto)
+
+lemma is_updsD:
+  "s' = mk_upds s f" if "is_upds s f s'"
+  "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  "dom s = {0..<n_vs}"
+proof -
+  from is_upds_make_updsI2[OF that(2-)] have "is_upds s f (mk_upds s f)" .
+  from is_upds_determ[OF that(1) this] show ?thesis .
+qed
+
+lemma is_upds_make_upds_concatI2:
+  "is_upds s (concat upds) (mk_upds s (concat upds))"
+  if "dom s = {0..<n_vs}" 
+    "\<forall>upd \<in> set upds. \<exists>p < n_ps. \<exists>l b g a r l'.
+        (l, b, g, a, upd, r, l') \<in> Simple_Network_Language.trans (N p)"
+  using that var_set
+  by (intro is_upds_make_updsI, clarsimp)
+     (smt (z3) atLeastLessThan_iff case_prod_conv domD var_setD zero_le)
+
+lemma is_upds_concat_dom2:
+  assumes "is_upds s (concat upds) s'"
+    and "\<forall>f \<in> set upds. \<exists> p < n_ps. \<exists> l b g a r l'.
+      (l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)"
+    and "dom s = {0..<n_vs}"
+  shows "dom s' = dom s"
+  using assms by (elim is_upds_dom) (auto dest!: var_setD)
+
+lemma is_upds_dom3:
+  assumes "is_upds s (concat_map fs ps) s'"
+    and "\<forall>p\<in>set ps. (L ! p, bs p, gs p, a, fs p, rs p, ls' p) \<in> trans (N p)"
+    and "set ps \<subseteq> {0..<n_ps}"
+    and "dom s = {0..<n_vs}"
+  shows "dom s' = dom s"
+  using assms by (elim is_upds_concat_dom2; force)
+
+lemma is_upds_make_updsI3:
+  "is_upds s (concat_map fs ps) (mk_upds s (concat_map fs ps))"
+  if "dom s = {0..<n_vs}" 
+    and "\<forall>p\<in>set ps. (L ! p, bs p, gs p, a, fs p, rs p, ls' p) \<in> trans (N p)"
+    and "set ps \<subseteq> {0..<n_ps}"
+  for s :: "nat \<Rightarrow> int option"
+  using that by (elim is_upds_make_upds_concatI2) force
+
+lemma is_upds_concatD:
+  assumes 
+    "dom s = {0..<n_vs}" and
+    "\<forall>p\<in>set ps.
+        (ls p, bs p, gs p, a, fs p, rs p, ls' p)
+        \<in> Simple_Network_Language.trans (N p)" and
+    "set ps \<subseteq> {0..<n_ps}" and
+    "is_upds s (concat_map fs ps) s'"
+  shows "s' = mk_upds s (concat_map fs ps)"
+proof -
+  let ?upds = "concat_map fs ps"
+  have "is_upds s ?upds (mk_upds s ?upds)"
+    using assms(1-3) by (intro is_upds_make_upds_concatI2; force)
+  from is_upds_determ[OF assms(4) this] show ?thesis
+    by (simp add: fold_map comp_def)
+qed
+
+
+
+context
+  notes [simp] = length_automata_eq_n_ps
+begin
+
+lemma trans_mapI:
+  assumes
+    "(L ! p, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)"
+    "p < n_ps"
+  shows
+    "(g, a, f, r, l') \<in> set (trans_map p (L ! p))"
+  using assms unfolding trans_map_def N_def fst_conv snd_conv trans_def
+  by (subst (asm) nth_map) (auto dest: in_union_map_ofI split: prod.split_asm simp: automaton_of_def)
+
+lemma trans_i_mapI:
+  assumes
+    "(L ! p, b, g, Sil a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+    "p < n_ps"
+  shows
+    "(b, g, a', f, r, l') \<in> set (trans_i_map p (L ! p))"
+  using assms unfolding trans_i_map_def set_map_filter by (fastforce dest: trans_mapI)
+
+lemma trans_mapI':
+  assumes
+    "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)"
+    "p < n_ps"
+  shows
+    "(b, g, a, f, r, l') \<in> set (trans_map p l)"
+  using assms unfolding trans_map_def N_def fst_conv snd_conv trans_def
+  by (subst (asm) nth_map) (auto dest: in_union_map_ofI split: prod.split_asm simp: automaton_of_def)
+
+lemma trans_mapD:
+  assumes
+    "(b, g, a, f, r, l') \<in> set (trans_map p l)"
+    "p < n_ps"
+  shows
+    "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)"
+  using assms unfolding trans_map_def N_def fst_conv snd_conv trans_def
+  by (subst nth_map) (auto split: prod.split elim: in_union_map_ofD[rotated] simp: automaton_of_def)
+
+lemma trans_map_iff:
+  assumes
+    "p < n_ps"
+  shows
+    "(b, g, a, f, r, l') \<in> set (trans_map p l)
+ \<longleftrightarrow> (l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)"
+  using trans_mapD trans_mapI' \<open>p < n_ps\<close> by auto
+
+lemma trans_i_mapD:
+  assumes
+    "(b, g, a', f, r, l') \<in> set (trans_i_map p (L ! p))"
+    "p < n_ps"
+  shows
+    "(L ! p, b, g, Sil a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+  using assms unfolding trans_i_map_def set_map_filter
+  by (force split: act.split_asm intro: trans_mapD)
+
+paragraph \<open>An additional brute force method for forward-chaining of facts\<close>
+
+method frules_all =
+  repeat_rotate \<open>frules\<close>, dedup_prems
+
+paragraph \<open>Internal transitions\<close>
+
+lemma get_committed_memI:
+  "(p, L ! p) \<in> set (get_committed L)" if "L ! p  \<in> committed (N p)" "p < n_ps"
+  using that unfolding get_committed_mem_iff by simp
+
+lemma check_bexp_bvalI:
+  "bval (the o s) b" if "check_bexp s b True"
+  "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  "dom s = {0..<n_vs}"
+proof -
+  from var_setD2[OF that(2,3)] \<open>dom s = {0..<n_vs}\<close> have "check_bexp s b (bval (the o s) b)"
+    by (intro check_bexp_bval, simp)
+  with check_bexp_determ that(1) show ?thesis
+    by auto
+qed
+
+lemma check_bexp_bvalD:
+  "check_bexp s b True" if "bval (the o s) b"
+  "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  "dom s = {0..<n_vs}"
+proof -
+  from var_setD2[OF that(2,3)] \<open>dom s = {0..<n_vs}\<close> have "check_bexp s b (bval (the o s) b)"
+    by (intro check_bexp_bval, simp)
+  with check_bexp_determ that(1) show ?thesis
+    by auto
+qed
+
+lemmas [forward2] =
+  trans_i_mapD
+  trans_i_mapI
+  get_committed_memI
+lemmas [forward3] =
+  is_upds_make_updsI2
+lemmas [forward4] =
+  is_updsD
+  is_upds_dom2
+  check_bexp_bvalI
+  check_bexp_bvalD
+
+lemma int_trans_from_correct:
+  "(int_trans_from, trans_int) \<in> transition_rel states'"
+  unfolding transition_rel_def
+proof (safe del: iffI)
+  note [more_elims] = allE
+  fix L s g a r L' s' assume "(L, s) \<in> states'"
+  then have "L \<in> states" "dom s = {0..<n_vs}" "bounded bounds s"
+    by auto
+  then have [simp]: "length L = n_ps" "check_bounded s"
+    by (auto simp: check_bounded_iff)
+  show "(((L, s), g, a, r, L', s') \<in> trans_int)
+    \<longleftrightarrow> ((g, a, r, L', s') \<in> set (int_trans_from (L, s)))"
+  proof (cases "get_committed L = []")
+    case True
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_int \<longleftrightarrow>
+      ((L, s), g, a, r, L', s') \<in> {((L, s), g, Internal a, r, (L', s')) | L s l b g f p a r l' L' s'.
+        (l, b, g, Sil a, f, r, l') \<in> trans (N p) \<and>
+        (\<forall>p < n_ps. L ! p \<notin> committed (N p)) \<and>
+        check_bexp s b True \<and>
+        L!p = l \<and> p < length L \<and> L' = L[p := l'] \<and> is_upds s f s' \<and>
+        L \<in> states \<and> bounded bounds s \<and> bounded bounds s'
+      }"
+      unfolding get_committed_empty_iff[symmetric] trans_int_def by blast
+    from True have **: "int_trans_from (L, s) = int_trans_from_all L s"
+      unfolding int_trans_from_def by simp
+    from \<open>dom s = _\<close> show ?thesis
+      unfolding * ** int_trans_from_all_def
+      apply clarsimp
+      unfolding int_trans_from_loc_def Let_def set_map_filter
+      apply clarsimp
+      apply safe
+      subgoal for b f p a' l'
+        apply frules
+        unfolding check_bounded_iff by (intros; solve_triv)
+      subgoal for p _ a' upds l'
+        apply simp
+        apply frules
+        using \<open>L \<in> states\<close> \<open>check_bounded s\<close> True[folded get_committed_empty_iff]
+        unfolding check_bounded_iff by (intros; solve_triv)
+      done
+  next
+    case False
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_int \<longleftrightarrow>
+      ((L, s), g, a, r, L', s') \<in> {((L, s), g, Internal a, r, (L', s')) | L s l b g f p a r l' L' s'.
+        (l, b, g, Sil a, f, r, l') \<in> trans (N p) \<and>
+        l \<in> committed (N p) \<and>
+        L!p = l \<and> p < length L \<and> check_bexp s b True \<and> L' = L[p := l'] \<and> is_upds s f s' \<and>
+        L \<in> states \<and> bounded bounds s \<and> bounded bounds s'
+      }"
+      unfolding get_committed_empty_iff[symmetric] trans_int_def by blast
+    from False have **: "int_trans_from (L, s) = int_trans_from_vec (get_committed L) L s"
+      unfolding int_trans_from_def by simp
+    from \<open>dom s = _\<close> \<open>L \<in> states\<close> show ?thesis
+      unfolding * ** int_trans_from_vec_def
+      apply clarsimp
+      unfolding int_trans_from_loc_def Let_def set_map_filter
+      apply clarsimp
+      apply safe
+      subgoal for f p a' l'
+        apply frules
+        unfolding check_bounded_iff by (intros; solve_triv)
+      subgoal for p _ a' upds l'
+        unfolding get_committed_mem_iff
+        apply (elims; simp)
+        apply frules
+        unfolding check_bounded_iff by (intros; solve_triv)
+      done
+  qed
+qed
+
+
+paragraph \<open>Mostly copied\<close>
+
+lemma in_actions_by_stateI:
+  assumes
+    "(b1, g1, a, r1) \<in> set xs" "a < length acc"
+  shows
+    "(q, b1, g1, a, r1) \<in> set (actions_by_state q xs acc ! a)
+    \<and> a < length (actions_by_state q xs acc)"
+  using assms unfolding actions_by_state_def
+  apply (induction xs arbitrary: acc)
+   apply (simp; fail)
+  apply simp
+  apply (erule disjE)
+   apply (rule fold_acc_preserv
+      [where P = "\<lambda> acc. (q, b1, g1, a, r1) \<in> set (acc ! a) \<and> a < length acc"]
+      )
+    apply (subst list_update_nth_split; auto)
+  by auto
+
+lemma in_actions_by_state'I:
+  assumes
+    "(b1, g1, a, r1) \<in> set xs" "a < num_actions"
+  shows
+    "(b1, g1, a, r1) \<in> set (actions_by_state' xs ! a)
+    \<and> a < length (actions_by_state' xs)"
+proof -
+  let ?f = "(\<lambda>t acc. acc[fst (snd (snd t)) := t # acc ! fst (snd (snd t))])"
+  have "(b1, g1, a, r1) \<in> set (fold ?f xs acc ! a)
+    \<and> a < length (fold ?f xs acc)" if "a < length acc" for acc
+    using assms(1) that
+    apply (induction xs arbitrary: acc)
+     apply (simp; fail)
+    apply simp
+    apply (erule disjE)
+     apply (rule fold_acc_preserv
+        [where P = "\<lambda> acc. (b1, g1, a, r1) \<in> set (acc ! a) \<and> a < length acc"]
+        )
+      apply (subst list_update_nth_split; auto)
+    by auto
+  with \<open>a < _\<close> show ?thesis
+    unfolding actions_by_state'_def by simp
+qed
+
+lemma in_actions_by_state_preserv:
+  assumes
+    "(q, b1, g1, a, r1) \<in> set (acc ! a)" "a < length acc"
+  shows
+    "(q, b1, g1, a, r1) \<in> set (actions_by_state y xs acc ! a)
+    \<and> a < length (actions_by_state y xs acc)"
+  using assms unfolding actions_by_state_def
+  apply -
+  apply (rule fold_acc_preserv
+      [where P = "\<lambda> acc. (q, b1, g1, a, r1) \<in> set (acc ! a) \<and> a < length acc"]
+      )
+   apply (subst list_update_nth_split; auto)
+  by auto
+
+lemma length_actions_by_state_preserv[simp]:
+  shows "length (actions_by_state y xs acc) = length acc"
+  unfolding actions_by_state_def by (auto intro: fold_acc_preserv simp: list_update_nth_split)
+
+lemma in_all_actions_by_stateI:
+  assumes
+    "a < num_actions" "q < n_ps" "(b1, g1, a, r1) \<in> set (M q (L ! q))"
+  shows
+    "(q, b1, g1, a, r1) \<in> set (all_actions_by_state M L ! a)"
+  unfolding all_actions_by_state_def
+  apply (rule fold_acc_ev_preserv
+      [where P = "\<lambda> acc. (q, b1, g1, a, r1) \<in> set (acc ! a)" and Q = "\<lambda> acc. a < length acc",
+        THEN conjunct1]
+      )
+      apply (rule in_actions_by_state_preserv[THEN conjunct1])
+  using assms by (auto intro: in_actions_by_stateI[THEN conjunct1])
+
+lemma in_all_actions_from_vecI:
+  assumes
+    "a < num_actions" "(b1, g1, a, r1) \<in> set (M q l)" "(q, l) \<in> set pairs"
+  shows
+    "(q, b1, g1, a, r1) \<in> set (all_actions_from_vec M pairs ! a)"
+  unfolding all_actions_from_vec_def using assms
+  by (intro fold_acc_ev_preserv
+      [where P = "\<lambda> acc. (q, b1, g1, a, r1) \<in> set (acc ! a)" and Q = "\<lambda> acc. a < length acc",
+        THEN conjunct1])
+    (auto intro: in_actions_by_stateI[THEN conjunct1] in_actions_by_state_preserv[THEN conjunct1])
+
+lemma actions_by_state_inj:
+  assumes "j < length acc"
+  shows "\<forall> (q, a) \<in> set (actions_by_state i xs acc ! j). (q, a) \<notin> set (acc ! j) \<longrightarrow> i = q"
+  unfolding actions_by_state_def
+  apply (rule fold_acc_preserv
+      [where P =
+        "\<lambda> acc'. (\<forall> (q, a) \<in> set (acc' ! j). (q, a) \<notin> set (acc ! j) \<longrightarrow> i = q) \<and> j < length acc'",
+        THEN conjunct1])
+  subgoal for x acc
+    by (cases "fst (snd (snd x)) = j"; simp)
+  using assms by auto
+
+lemma actions_by_state_inj':
+  assumes "j < length acc" "(q, a) \<notin> set (acc ! j)" "(q, a) \<in> set (actions_by_state i xs acc ! j)"
+  shows "i = q"
+  using actions_by_state_inj[OF assms(1)] assms(2-) by fast
+
+lemma in_actions_by_stateD:
+  assumes
+    "(q, b, g, a, t) \<in> set (actions_by_state i xs acc ! j)" "(q, b, g, a, t) \<notin> set (acc ! j)"
+    "j < length acc"
+  shows
+    "(b, g, a, t) \<in> set xs \<and> j = a"
+  using assms unfolding actions_by_state_def
+  apply -
+  apply (drule fold_evD
+      [where y = "(b, g, a, t)" and Q = "\<lambda> acc'. length acc' = length acc"
+        and R = "\<lambda> (_, _, a', t). a' = j"]
+      )
+       apply assumption
+      apply (subst (asm) list_update_nth_split[of j]; force)
+     apply simp+
+   apply (subst (asm) list_update_nth_split[of j]; force)
+  by auto
+
+lemma in_actions_by_state'D:
+  assumes
+    "(b, g, a, r) \<in> set (actions_by_state' xs ! a')" "a' < num_actions"
+  shows
+    "(b, g, a, r) \<in> set xs \<and> a' = a"
+proof -
+  let ?f = "(\<lambda>t acc. acc[fst (snd (snd t)) := t # acc ! fst (snd (snd t))])"
+  have "(b, g, a, r) \<in> set xs \<and> a' = a"
+    if "(b, g, a, r) \<in> set (fold ?f xs acc ! a')" "(b, g, a, r) \<notin> set (acc ! a')" "a' < length acc"
+    for acc
+    using that
+    apply -
+    apply (drule fold_evD
+        [where y = "(b, g, a, r)" and Q = "\<lambda> acc'. length acc' = length acc"
+          and R = "\<lambda> (_, _, a, t). a = a'"]
+        )
+         apply ((subst (asm) list_update_nth_split[where j = a'])?; solves auto)+
+    done
+  with assms show ?thesis
+    unfolding actions_by_state'_def by auto
+qed
+
+lemma in_all_actions_by_stateD:
+  assumes
+    "(q, b1, g1, a, r1) \<in> set (all_actions_by_state M L ! a')" "a' < num_actions"
+  shows
+    "(b1, g1, a, r1) \<in> set (M q (L ! q)) \<and> q < n_ps \<and> a' = a"
+  using assms
+  unfolding all_actions_by_state_def
+  apply -
+  apply (drule fold_evD''[where y = q and Q = "\<lambda> acc. length acc = num_actions"])
+      apply (simp; fail)
+     apply (drule actions_by_state_inj'[rotated])
+       apply (simp; fail)+
+  apply safe
+    apply (drule in_actions_by_stateD)
+      apply assumption
+     apply (rule fold_acc_preserv)
+      apply (simp; fail)+
+  subgoal premises prems
+  proof -
+    from prems(2) have "q \<in> set [0..<n_ps]" by auto
+    then show ?thesis by simp
+  qed
+  by (auto intro: fold_acc_preserv dest!: in_actions_by_stateD)
+
+lemma length_all_actions_by_state_preserv:
+  "length (all_actions_by_state M L) = num_actions"
+  unfolding all_actions_by_state_def by (auto intro: fold_acc_preserv)
+
+lemma length_actions_by_state'_preserv:
+  "length (actions_by_state' xs) = num_actions"
+  unfolding actions_by_state'_def by (rule fold_acc_preserv; simp)
+
+lemma all_actions_from_vecD:
+  assumes
+    "(q, b1, g1, a, r1) \<in> set (all_actions_from_vec M pairs ! a')" "a' < num_actions"
+    "distinct (map fst pairs)"
+  shows
+    "\<exists> l. (q, l) \<in> set pairs \<and> (b1, g1, a, r1) \<in> set (M q l) \<and> a' = a"
+  using assms(1,2)
+  unfolding all_actions_from_vec_def
+  apply -
+  apply (drule fold_evD2'[where
+        y = "(q, SOME l. (q, l) \<in> set pairs)"
+        and Q = "\<lambda> acc. length acc = num_actions"])
+      apply (clarsimp; fail)
+     apply clarsimp
+     apply (drule actions_by_state_inj'[rotated])
+       apply assumption
+      apply assumption
+     apply simp
+  subgoal
+    using assms(3) by (meson distinct_map_fstD someI_ex)
+  apply (solves auto)+
+  apply clarsimp
+  apply (drule in_actions_by_stateD)
+  apply (simp; fail)
+   apply (rule fold_acc_preserv)
+  apply (solves auto)+
+  subgoal premises prems for ys zs
+    using prems(4) by (subst \<open>pairs = _\<close>) auto
+  done
+
+lemma all_actions_from_vecD2:
+  assumes
+    "(q, b1, g1, a, r1) \<in> set (all_actions_from_vec M pairs ! a')" "a' < num_actions"
+    "(q, l) \<in> set pairs" "distinct (map fst pairs)"
+  shows
+    "(b1, g1, a, r1) \<in> set (M q l) \<and> a' = a"
+  using assms(1,2,3)
+  unfolding all_actions_from_vec_def
+  apply -
+  apply (drule fold_evD2'[where y = "(q, l)" and Q = "\<lambda> acc. length acc = num_actions"])
+      apply (clarsimp; fail)
+     apply clarsimp
+     apply (drule actions_by_state_inj'[rotated])
+       apply assumption
+      apply assumption
+     apply simp
+  subgoal
+    using assms(4) by (meson distinct_map_fstD)
+  by (auto intro: fold_acc_preserv dest!: in_actions_by_stateD)
+
+paragraph \<open>Binary transitions\<close>
+
+lemma bin_trans_from_correct:
+  "(bin_trans_from, trans_bin) \<in> transition_rel states'"
+  unfolding transition_rel_def
+proof (safe del: iffI)
+  fix L s g a r L' s' assume "(L, s) \<in> states'"
+  then have "L \<in> states" "dom s = {0..<n_vs}" "bounded bounds s"
+    by auto
+  then have [simp]: "length L = n_ps"
+    by auto
+  define IN where  "IN  = all_actions_by_state trans_in_map L"
+  define OUT where "OUT = all_actions_by_state trans_out_map L"
+  have IN_I:
+    "(p, b, g, a', f, r, l') \<in> set (IN ! a')"
+    if "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+      "p < n_ps" "a' < num_actions"
+    for p b g a' f r l'
+  proof -
+    from trans_mapI[OF that(1,2)] have "(b, g, In a', f, r, l') \<in> set (trans_map p (L ! p))"
+      by auto
+    then have "(b, g, a', f, r, l') \<in> set (trans_in_map p (L ! p))"
+      unfolding trans_in_map_def set_map_filter by (auto 4 7)
+    with \<open>p < _\<close> \<open>a' < _\<close> show ?thesis
+      unfolding IN_def by (intro in_all_actions_by_stateI)
+  qed
+  have OUT_I:
+    "(p, b, g, a', f, r, l') \<in> set (OUT ! a')"
+    if "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+      "p < n_ps" "a' < num_actions"
+    for p b g a' f r l'
+  proof -
+    from trans_mapI[OF that(1,2)] have "(b, g, Out a', f, r, l') \<in> set (trans_map p (L ! p))"
+      by auto
+    then have "(b, g, a', f, r, l') \<in> set (trans_out_map p (L ! p))"
+      unfolding trans_out_map_def set_map_filter by (auto 4 7)
+    with \<open>p < _\<close> \<open>a' < _\<close> show ?thesis
+      unfolding OUT_def by (intro in_all_actions_by_stateI)
+  qed
+  have IN_D:
+    "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p) \<and> p < n_ps \<and> a' = a1"
+    if "(p, b, g, a', f, r, l') \<in> set (IN ! a1)" "a1 < num_actions"
+    for p b g a' f r l' a1
+    using that
+    unfolding IN_def
+    apply -
+    apply (drule in_all_actions_by_stateD)
+     apply assumption
+    unfolding trans_in_map_def set_map_filter
+    apply (clarsimp split: option.split_asm)
+    apply (auto split: act.split_asm dest: trans_mapD)
+    done
+  have OUT_D:
+    "(L ! p, b, g, Out a1, f, r, l') \<in> Simple_Network_Language.trans (N p) \<and> p < n_ps"
+    if "(p, b, g, a', f, r, l') \<in> set (OUT ! a1)" "a1 < num_actions"
+    for p b g a' f r l' a1
+    using that
+    unfolding OUT_def
+    apply -
+    apply (drule in_all_actions_by_stateD)
+     apply assumption
+    unfolding trans_out_map_def set_map_filter
+    apply (clarsimp split: option.split_asm)
+    apply (auto split: act.split_asm dest: trans_mapD)
+    done
+  have IN_p_num:
+    "p < n_ps"
+    if "(p, b, g, a', f, r, l') \<in> set (IN ! a1)" "a1 < num_actions"
+    for p b g a' f r l' a1
+    using that unfolding IN_def by (auto dest: in_all_actions_by_stateD split: option.split_asm)
+  have OUT_p_num:
+    "p < n_ps"
+    if "(p, b, g, a', f, r, l') \<in> set (OUT ! a1)" "a1 < num_actions"
+    for p b g a' f r l' a1
+    using that unfolding OUT_def by (auto dest: in_all_actions_by_stateD split: option.split_asm)
+  have actions_unique:
+    "a' = a1"
+    if "(p, b, g, a', f, r, l') \<in> set (IN ! a1)" "a1 < num_actions"
+    for p b g a' f r l' a1
+    using that unfolding IN_def trans_in_map_def set_map_filter
+    by (auto dest: in_all_actions_by_stateD split: option.split_asm)
+  note [forward3] = OUT_I IN_I
+  note [forward2] = action_setD IN_D OUT_D
+  show "(((L, s), g, a, r, L', s') \<in> trans_bin) =
+        ((g, a, r, L', s') \<in> set (bin_trans_from (L, s)))"
+  proof (cases "get_committed L = []")
+    case True
+    with get_committed_empty_iff[of L] have "\<forall>p<n_ps. L ! p \<notin> committed (N p)"
+      by simp
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_bin \<longleftrightarrow> ((L, s), g, a, r, L', s') \<in>
+      {((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+        L s L' s' s'' a p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'.
+        a \<notin> local.broadcast \<and>
+        (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+        (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+        L!p = l1 \<and> L!q = l2 \<and> p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+        check_bexp s b1 True \<and> check_bexp s b2 True \<and>
+        L' = L[p := l1', q := l2'] \<and> is_upds s f1 s' \<and> is_upds s' f2 s''
+        \<and> L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }
+    "
+      unfolding trans_bin_def by blast
+    from True have **:
+      "bin_trans_from (L, s)
+      = concat (map (\<lambda>a. pairs_by_action L s (OUT ! a) (IN ! a)) bin_actions)"
+      unfolding bin_trans_from_def IN_def OUT_def by simp
+    from \<open>dom s = _\<close> \<open>L \<in> _\<close> show ?thesis
+      unfolding * **
+      apply clarsimp
+      unfolding pairs_by_action_def
+      apply (clarsimp simp: set_map_filter Let_def)
+      apply safe
+      subgoal for _ s'' _ a' p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'
+        apply clarsimp
+        apply (inst_existentials a')
+        subgoal
+          apply frules
+          apply simp
+          apply frules_all
+          unfolding check_bounded_iff by (intros; solve_triv)
+        subgoal
+          by (simp add: mem_bin_actions_iff; frules; simp)
+        done
+      subgoal
+        unfolding mem_bin_actions_iff
+        apply simp
+        apply (erule conjE)
+        apply frules
+        apply elims
+        apply frules_all
+        apply frules_all
+        apply elims
+        apply intros
+        using \<open>bounded bounds s\<close> unfolding check_bounded_iff[symmetric] by solve_triv+
+      done
+  next
+    case False
+    with get_committed_empty_iff[of L] have "\<exists>p<n_ps. L ! p \<in> committed (N p)"
+      by simp
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_bin \<longleftrightarrow> ((L, s), g, a, r, L', s') \<in>
+      {((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+        L s L' s' s'' a p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'.
+        a \<notin> local.broadcast \<and>
+        (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+        (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+        (l1 \<in> committed (N p) \<or> l2 \<in> committed (N q)) \<and>
+        L!p = l1 \<and> L!q = l2 \<and> p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+        check_bexp s b1 True \<and> check_bexp s b2 True \<and>
+        L' = L[p := l1', q := l2'] \<and> is_upds s f1 s' \<and> is_upds s' f2 s'' \<and>
+        L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }"
+      unfolding trans_bin_def
+      by - (rule iffI; elims add: CollectE; intros add: CollectI; blast)
+    let ?S1 =
+      "{((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+        L s L' s' s'' a p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'.
+          a \<notin> local.broadcast \<and>
+          (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+          (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+          l1 \<in> committed (N p) \<and>
+          L!p = l1 \<and> L!q = l2 \<and> p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+          check_bexp s b1 True \<and> check_bexp s b2 True \<and>
+          L' = L[p := l1', q := l2'] \<and> is_upds s f1 s' \<and> is_upds s' f2 s'' \<and>
+          L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }"
+    let ?S2 =
+      "{((L, s), g1 @ g2, Bin a, r1 @ r2, (L', s'')) |
+        L s L' s' s'' a p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'.
+          a \<notin> local.broadcast \<and>
+          (l1, b1, g1, In a,  f1, r1, l1') \<in> trans (N p) \<and>
+          (l2, b2, g2, Out a, f2, r2, l2') \<in> trans (N q) \<and>
+          l2 \<in> committed (N q) \<and>
+          L!p = l1 \<and> L!q = l2 \<and>
+          p < length L \<and> q < length L \<and> p \<noteq> q \<and>
+          check_bexp s b1 True \<and> check_bexp s b2 True \<and>
+          L' = L[p := l1', q := l2'] \<and>
+          is_upds s f1 s' \<and> is_upds s' f2 s'' \<and>
+          L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }"
+    have *: "((L, s), g, a, r, L', s') \<in> trans_bin \<longleftrightarrow>
+      ((L, s), g, a, r, L', s') \<in> ?S1 \<or> ((L, s), g, a, r, L', s') \<in> ?S2"
+      unfolding * by clarsimp (rule iffI; elims add: disjE; intros add: disjI1 disjI2 HOL.refl)
+    define pairs where "pairs = get_committed L"
+    define In2 where "In2  = all_actions_from_vec trans_in_map pairs"
+    define Out2 where "Out2 = all_actions_from_vec trans_out_map pairs"
+    have In2_I:
+      "(p, b, g, a', f, r, l') \<in> set (In2 ! a')"
+      if "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+        "p < n_ps" "a' < num_actions" "L ! p \<in> committed (N p)"
+      for p b g a' f r l'
+    proof -
+      from \<open>L ! p \<in> committed (N p)\<close> \<open>p < n_ps\<close> have "(p, L ! p) \<in> set pairs"
+        unfolding pairs_def get_committed_mem_iff by blast
+      from trans_mapI[OF that(1,2)] have "(b, g, In a', f, r, l') \<in> set (trans_map p (L ! p))"
+        by auto
+      then have "(b, g, a', f, r, l') \<in> set (trans_in_map p (L ! p))"
+        unfolding trans_in_map_def set_map_filter by (auto 4 7)
+      with \<open>p < _\<close> \<open>a' < _\<close> \<open>_ \<in> set pairs\<close> show ?thesis
+        unfolding In2_def by (intro in_all_actions_from_vecI)
+    qed
+    have Out2_I:
+      "(p, b, g, a', f, r, l') \<in> set (Out2 ! a')"
+      if "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+        "p < n_ps" "a' < num_actions" "L ! p \<in> committed (N p)"
+      for p b g a' f r l'
+    proof -
+      from \<open>L ! p \<in> committed (N p)\<close> \<open>p < n_ps\<close> have "(p, L ! p) \<in> set pairs"
+        unfolding pairs_def get_committed_mem_iff by blast
+      from trans_mapI[OF that(1,2)] have "(b, g, Out a', f, r, l') \<in> set (trans_map p (L ! p))"
+        by auto
+      then have "(b, g, a', f, r, l') \<in> set (trans_out_map p (L ! p))"
+        unfolding trans_out_map_def set_map_filter by (auto 4 7)
+      with \<open>p < _\<close> \<open>a' < _\<close> \<open>_ \<in> set pairs\<close> show ?thesis
+        unfolding Out2_def by (intro in_all_actions_from_vecI)
+    qed
+    have "distinct (map fst pairs)"
+      unfolding pairs_def get_committed_def distinct_map inj_on_def Let_def
+      by (auto simp: set_map_filter intro!: distinct_map_filterI split: if_split_asm)
+    have in_pairsD: "p < n_ps" "l = L ! p" "L ! p \<in> committed (N p)"
+      if "(p, l) \<in> set pairs" for p l
+      using that using get_committed_mem_iff pairs_def by auto
+    have In2_D:
+      "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p) \<and>
+      p < n_ps \<and> a' = a1 \<and> L ! p \<in> committed (N p)"
+      if "(p, b, g, a', f, r, l') \<in> set (In2 ! a1)" "a1 < num_actions"
+      for p b g a' f r l' a1
+      using that
+      unfolding In2_def
+      apply -
+      apply (drule all_actions_from_vecD)
+        apply assumption
+       apply (rule \<open>distinct _\<close>)
+      unfolding trans_in_map_def set_map_filter
+      apply (clarsimp split: option.split_asm)
+        apply (auto dest: in_pairsD trans_mapD split: act.split_asm)
+      done
+    have Out2_D:
+      "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)
+      \<and> p < n_ps \<and> a' = a1 \<and> L ! p \<in> committed (N p)"
+      if "(p, b, g, a', f, r, l') \<in> set (Out2 ! a1)" "a1 < num_actions"
+      for p b g a' f r l' a1
+      using that
+      unfolding Out2_def
+      apply -
+      apply (drule all_actions_from_vecD)
+        apply assumption
+       apply (rule \<open>distinct _\<close>)
+      unfolding trans_out_map_def set_map_filter
+      apply (clarsimp split: option.split_asm)
+        apply (auto dest: in_pairsD trans_mapD split: act.split_asm)
+      done
+    from False have **: "bin_trans_from (L, s) =
+        concat (map (\<lambda>a. pairs_by_action L s (OUT ! a) (In2 ! a)) bin_actions)
+      @ concat (map (\<lambda>a. pairs_by_action L s (Out2 ! a) (IN ! a)) bin_actions)"
+      unfolding bin_trans_from_def IN_def OUT_def In2_def Out2_def pairs_def
+      by (simp add: Let_def)
+    from \<open>dom s = _\<close> \<open>L \<in> _\<close> have "
+      ((L, s), g, a, r, L', s') \<in> ?S1 \<longleftrightarrow> (g, a, r, L', s') \<in>
+      set (concat (map (\<lambda>a. pairs_by_action L s (OUT ! a) (In2 ! a)) bin_actions))"
+      apply clarsimp
+      unfolding pairs_by_action_def
+      apply (clarsimp simp: set_map_filter Let_def)
+      apply safe
+      subgoal for _ s'' _ a' p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'
+        apply clarsimp
+        apply (inst_existentials a')
+        subgoal
+          supply [forward4] = In2_I
+          apply frules_all
+          apply simp
+          apply frules_all
+          unfolding check_bounded_iff by (intros; solve_triv)
+        subgoal
+          by (simp add: mem_bin_actions_iff; frules; simp)
+        done
+      subgoal
+        supply [forward2] = In2_D OUT_D
+        apply simp
+        unfolding mem_bin_actions_iff
+        apply (erule conjE)
+        apply frules_all
+        apply elims
+        apply frules_all
+        apply elims
+        apply simp
+        apply frules_all
+        using \<open>bounded bounds s\<close> unfolding check_bounded_iff[symmetric] by (intros; solve_triv)
+      done
+    moreover from \<open>dom s = _\<close> \<open>L \<in> _\<close> have "
+      ((L, s), g, a, r, L', s') \<in> ?S2 \<longleftrightarrow> (g, a, r, L', s')
+      \<in> set (concat (map (\<lambda>a. pairs_by_action L s (Out2 ! a) (IN ! a)) bin_actions))"
+      supply [forward2] = Out2_D In2_D
+      supply [forward4] = Out2_I
+      apply clarsimp
+      unfolding pairs_by_action_def
+      apply (clarsimp simp: set_map_filter Let_def)
+      apply safe
+      subgoal for _ s'' _ a' p q l1 b1 g1 f1 r1 l1' l2 b2 g2 f2 r2 l2'
+        apply clarsimp
+        apply (inst_existentials a')
+        subgoal
+          apply frules_all
+          apply simp
+          apply frules_all
+          unfolding check_bounded_iff by (intros; solve_triv)
+        subgoal
+          unfolding mem_bin_actions_iff by frules_all simp
+        done
+      subgoal
+        apply simp
+        unfolding mem_bin_actions_iff
+        apply (erule conjE)
+        apply frules_all
+        apply elims
+        apply frules_all
+        apply elims
+        apply simp
+        apply frules_all
+        using \<open>bounded bounds s\<close> unfolding check_bounded_iff[symmetric] by (intros; solve_triv)
+      done
+    ultimately show ?thesis
+      unfolding * ** by simp
+  qed
+qed
+
+paragraph \<open>Broadcast transitions\<close>
+
+lemma make_combs_alt_def:
+  "make_combs p a xs \<equiv>
+    let
+      ys = 
+        map (\<lambda> i. map (\<lambda>t. (i, t)) (xs ! i ! a))
+        (filter
+          (\<lambda>i. xs ! i ! a \<noteq> [] \<and> i \<noteq> p)
+          [0..<n_ps])
+    in if ys = [] then [] else product_lists ys"
+  apply (rule eq_reflection)
+  unfolding make_combs_def
+  apply (simp add: map_filter_def comp_def if_distrib[where f = the])
+  apply intros
+  apply (fo_rule arg_cong)
+  apply auto
+  done
+
+lemma list_all2_fst_aux:
+  "map fst xs = ys" if "list_all2 (\<lambda>x y. fst x = y) xs ys"
+  using that by (induction) auto
+
+lemma broad_trans_from_correct:
+  "(broad_trans_from, trans_broad) \<in> transition_rel states'"
+  unfolding transition_rel_def
+proof (safe del: iffI)
+  fix L s g a r L' s' assume "(L, s) \<in> states'"
+  then have "L \<in> states" "dom s = {0..<n_vs}" "bounded bounds s"
+    by auto
+  then have [simp]: "length L = n_ps"
+    by auto
+  define IN  where "IN  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps]"
+  define OUT where "OUT = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps]"
+  define IN' where "IN' = map (map (filter (\<lambda> (b, _). bval (the o s) b))) IN"
+  define OUT' where "OUT' = map (map (filter (\<lambda> (b, _). bval (the o s) b))) OUT"
+  have IN_I:
+    "(b, g, a', f, r, l') \<in> set (IN ! p ! a')"
+    if "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+      "p < n_ps" "a' < num_actions" "a' \<in> set broadcast"
+    for p b g a' f r l'
+  proof -
+    from trans_mapI[OF that(1,2)] \<open>a' \<in> _\<close> have
+      "(b, g, a', f, r, l') \<in> set (trans_in_broad_map p (L ! p))"
+      unfolding trans_in_broad_map_def set_map_filter by (auto 4 7)
+    with \<open>a' < _\<close> have "(b, g, a', f, r, l') \<in> set (trans_in_broad_grouped p (L ! p) ! a')"
+      unfolding trans_in_broad_grouped_def by (auto dest: in_actions_by_state'I)
+    with \<open>p < _\<close> show ?thesis
+      unfolding IN_def by auto
+  qed
+  have OUT_I:
+    "(b, g, a', f, r, l') \<in> set (OUT ! p ! a')"
+    if "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)"
+      "p < n_ps" "a' < num_actions" "a' \<in> set broadcast"
+    for p b g a' f r l'
+  proof -
+    from trans_mapI[OF that(1,2)] \<open>a' \<in> _\<close> have
+      "(b, g, a', f, r, l') \<in> set (trans_out_broad_map p (L ! p))"
+      unfolding trans_out_broad_map_def set_map_filter by (auto 4 7)
+    with \<open>a' < _\<close> have "(b, g, a', f, r, l') \<in> set (trans_out_broad_grouped p (L ! p) ! a')"
+      unfolding trans_out_broad_grouped_def by (auto dest: in_actions_by_state'I)
+    with \<open>p < _\<close> show ?thesis
+      unfolding OUT_def by auto
+  qed
+  have IN_D:
+    "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)
+     \<and> a' = a1 \<and> a1 \<in> set broadcast"
+    if "(b, g, a', f, r, l') \<in> set (IN ! p ! a1)" "a1 < num_actions" "p < n_ps"
+    for p b g a' f r l' a1
+    using that unfolding IN_def trans_in_broad_grouped_def
+    apply simp
+    apply (drule in_actions_by_state'D)
+    unfolding trans_in_broad_map_def set_map_filter
+    by (auto split: option.split_asm) (auto split: act.split_asm if_split_asm dest: trans_mapD)
+  have [simp]: "length IN = n_ps" "length OUT = n_ps"
+    unfolding IN_def OUT_def by simp+
+  have [simp]: "length (IN ! p) = num_actions" "length (OUT ! p) = num_actions" if "p < n_ps" for p
+    using that by (simp add:
+      length_actions_by_state'_preserv trans_in_broad_grouped_def trans_out_broad_grouped_def
+      IN_def OUT_def)+
+  have OUT_D:
+    "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)
+     \<and> a' = a1 \<and> a1 \<in> set broadcast"
+    if "(b, g, a', f, r, l') \<in> set (OUT ! p ! a1)" "a1 < num_actions" "p < n_ps"
+    for p b g a' f r l' a1
+    using that unfolding OUT_def trans_out_broad_grouped_def
+    apply simp
+    apply (drule in_actions_by_state'D)
+    unfolding trans_out_broad_map_def set_map_filter
+    by (auto split: option.split_asm) (auto split: act.split_asm if_split_asm dest: trans_mapD)
+  have IN'_I:
+    "(b, g, a', f, r, l') \<in> set (IN' ! p ! a')"
+    if "(b, g, a', f, r, l') \<in> set (IN ! p ! a')" "p < n_ps" "a' < num_actions"
+       "check_bexp s b True" for p b g a' f r l'
+    using that \<open>dom s = {0..<n_vs}\<close> unfolding IN'_def
+    by (auto dest!: IN_D[THEN conjunct1] elim: check_bexp_bvalI)
+  have IN'_D:
+    "(L ! p, b, g, In a', f, r, l') \<in> Simple_Network_Language.trans (N p)
+     \<and> a' = a1 \<and> a1 \<in> set broadcast \<and> check_bexp s b True"
+    if "(b, g, a', f, r, l') \<in> set (IN' ! p ! a1)" "a1 < num_actions" "p < n_ps"
+    for p b g a' f r l' a1
+    using that \<open>dom s = {0..<n_vs}\<close> unfolding IN'_def by (auto dest!: IN_D elim: check_bexp_bvalD)
+  have OUT'_I:
+    "(b, g, a', f, r, l') \<in> set (OUT' ! p ! a')"
+    if "(b, g, a', f, r, l') \<in> set (OUT ! p ! a')" "p < n_ps" "a' < num_actions"
+       "check_bexp s b True" for p b g a' f r l'
+    using that \<open>dom s = {0..<n_vs}\<close> unfolding OUT'_def
+    by (auto dest!: OUT_D[THEN conjunct1] elim: check_bexp_bvalI)
+  have OUT'_D:
+    "(L ! p, b, g, Out a', f, r, l') \<in> Simple_Network_Language.trans (N p)
+     \<and> a' = a1 \<and> a1 \<in> set broadcast \<and> check_bexp s b True"
+    if "(b, g, a', f, r, l') \<in> set (OUT' ! p ! a1)" "a1 < num_actions" "p < n_ps"
+    for p b g a' f r l' a1
+    using that \<open>dom s = {0..<n_vs}\<close> unfolding OUT'_def by (auto dest!: OUT_D elim: check_bexp_bvalD)
+  define make_trans where "make_trans a p \<equiv>
+      let
+        outs = OUT' ! p ! a
+      in if outs = [] then []
+      else
+        let
+          combs = make_combs p a IN';
+          outs = map (\<lambda>t. (p, t)) outs;
+          combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+          init = ([], Broad a, [], (L, s))
+        in
+        filter (\<lambda> (g, a, r, L, s). check_bounded s) (
+          map (\<lambda>comb.
+              fold
+                (\<lambda>(q, b2, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+                  (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+                )
+                comb
+                init
+          ) combs)" for a p
+  {
+    fix p ps bs gs a' fs rs ls'
+    assume assms:
+      "\<forall>p\<in>set ps.
+          (L ! p, bs p, gs p, In a', fs p, rs p, ls' p) \<in> trans (N p)"
+      "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+          (\<forall>b g f r l'. (L ! q, b, g, In a', f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)"
+      "p < n_ps" "set ps \<subseteq> {0..<n_ps}" "p \<notin> set ps" "distinct ps" "sorted ps"
+      "a' < num_actions" "a' \<in> set broadcast" "\<forall>p \<in> set ps. check_bexp s (bs p) True"
+    define ys where "ys = List.map_filter
+        (\<lambda> i.
+          if i = p then None
+          else if IN' ! i ! a' = [] then None
+          else Some (map (\<lambda>t. (i, t)) (IN' ! i ! a'))
+        )
+        [0..<n_ps]"
+    have "filter (\<lambda>i. IN' ! i ! a' \<noteq> [] \<and> i \<noteq> p) [0..<n_ps] = ps"
+      apply (rule filter_distinct_eqI)
+      subgoal
+        using assms(4-) by (simp add: sorted_distinct_subseq_iff)
+      subgoal
+        using assms(1,3-) by (auto 4 3 dest!: IN_I IN'_I)
+      subgoal
+        using assms(1,2,4,8-)
+        apply -
+        apply (rule ccontr)
+        apply simp
+        apply elims
+        apply (drule hd_in_set)
+        subgoal for x
+          by (cases "hd (IN' ! x ! a')") (fastforce dest!: IN'_D)
+        done
+      by auto
+    then have "length ys = length ps"
+      unfolding ys_def map_filter_def by simp
+    have non_empty: "ys \<noteq> []" if "ps \<noteq> []"
+      using \<open>_ = ps\<close> \<open>ps \<noteq> []\<close> unfolding ys_def map_filter_def by simp
+    have make_combsD:
+      "map (\<lambda>p. (p, bs p, gs p, a', fs p, rs p, ls' p)) ps \<in> set (make_combs p a' IN')"
+      if "ps \<noteq> []"
+    proof -
+      from assms(1,3-) have
+        "\<forall>i<length ps. let p = ps ! i in (p, bs p, gs p, a', fs p, rs p, ls' p) \<in> set (ys ! i)"
+        unfolding ys_def Let_def map_filter_def
+        apply (simp add: comp_def if_distrib[where f = the])
+        apply (subst (2) map_cong)
+          apply (rule HOL.refl)
+         apply (simp; fail)
+        apply (simp add: \<open>_ = ps\<close>)
+        by (intros add: image_eqI[OF HOL.refl] IN_I IN'_I; simp add: subset_code(1))
+      with non_empty[OF that] show ?thesis
+        unfolding make_combs_def ys_def[symmetric] Let_def
+        by (auto simp: \<open>length ys = _\<close> product_lists_set intro:list_all2_all_nthI)
+    qed
+    have make_combs_empty:
+      "make_combs p a' IN' = [] \<longleftrightarrow> ps = []"
+    proof (cases "ps = []")
+      case True
+      then show ?thesis
+        using \<open>length ys = length ps\<close> unfolding make_combs_def ys_def[symmetric] Let_def by auto
+    next
+      case False
+      then show ?thesis
+        using make_combsD by auto
+    qed
+    note make_combsD make_combs_empty
+  } note make_combsD = this(1) and make_combs_empty = this(2)
+  have make_combs_emptyD: "filter (\<lambda>i. IN' ! i ! a' \<noteq> [] \<and> i \<noteq> p) [0..<n_ps] = []"
+    if "make_combs p a' IN' = []" for p a'
+    apply (rule filter_distinct_eqI)
+    subgoal
+      by auto
+      subgoal
+        by auto
+      subgoal
+        using that unfolding make_combs_alt_def
+        by (auto simp: filter_empty_conv product_lists_empty split: if_split_asm)
+      subgoal
+        by simp
+      done
+  have make_combsI:
+    "\<exists> ps bs gs fs rs ls'.
+      (\<forall>p\<in>set ps.
+       (L ! p, bs p, gs p, In a', fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+      (\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+       (\<forall>b g f r l'. (L ! q, b, g, In a', f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)) \<and>
+      set ps \<subseteq> {0..<n_ps} \<and> p \<notin> set ps \<and> distinct ps \<and> sorted ps \<and> ps \<noteq> [] \<and> a' \<in> set broadcast
+      \<and> (\<forall>p \<in> set ps. check_bexp s (bs p) True)
+      \<and> xs = map (\<lambda> p. (p, bs p, gs p, a', fs p, rs p, ls' p)) ps
+      \<and> filter (\<lambda>i. IN' ! i ! a' \<noteq> [] \<and> i \<noteq> p) [0..<n_ps] = ps"
+    if "xs \<in> set (make_combs p a' IN')" "p < n_ps" "a' < num_actions"
+    for xs p a'
+  proof -
+    define ps bs gs fs rs ls' where defs:
+      "ps  = map fst xs"
+      "bs  = (\<lambda>i. case the (map_of xs i) of (b, g, a, f, r, l') \<Rightarrow> b)"
+      "gs  = (\<lambda>i. case the (map_of xs i) of (b, g, a, f, r, l') \<Rightarrow> g)"
+      "fs  = (\<lambda>i. case the (map_of xs i) of (b, g, a, f, r, l') \<Rightarrow> f)"
+      "rs  = (\<lambda>i. case the (map_of xs i) of (b, g, a, f, r, l') \<Rightarrow> r)"
+      "ls' = (\<lambda>i. case the (map_of xs i) of (b, g, a, f, r, l') \<Rightarrow> l')"
+    have "filter (\<lambda>i. IN' ! i ! a' \<noteq> [] \<and> i \<noteq> p) [0..<n_ps] = ps"
+      apply (rule filter_distinct_eqI)
+      subgoal
+        using that
+        unfolding defs make_combs_alt_def Let_def
+        by (auto simp: set_map_filter product_lists_set list_all2_map2 list.rel_eq
+            dest: list_all2_map_fst_aux split: if_split_asm)
+      subgoal
+        using that unfolding defs make_combs_alt_def Let_def
+        by (auto simp: set_map_filter product_lists_set dest!: list_all2_set1 split: if_split_asm)
+      subgoal
+        using that
+        unfolding defs make_combs_alt_def Let_def
+        by (auto
+            simp: set_map_filter product_lists_set list_all2_map2 list.rel_eq
+            dest!: map_eq_imageD list_all2_map_fst_aux split: if_split_asm)
+      subgoal
+        by simp
+      done
+    then have "set ps \<subseteq> {0..<n_ps}" "p \<notin> set ps" "distinct ps" "sorted ps"
+      by (auto intro: sorted_filter')
+    have to_map: "a' = a" "the (map_of xs q) = (b, g, a, r, f, l')"
+      if "(q, b, g, a, r, f, l') \<in> set xs" for b q g a r f l'
+      using that
+       apply -
+      subgoal
+        using \<open>set ps \<subseteq> _\<close> \<open>a' < _\<close> \<open>xs \<in> _\<close>
+        by
+          (simp 
+            add: make_combs_alt_def product_lists_set \<open>_ = ps\<close> list_all2_map2
+            split: if_split_asm
+            ) (auto 4 3 dest!: IN'_D list_all2_set1)
+      subgoal
+        using \<open>distinct ps\<close>
+        by (subst (asm) map_of_eq_Some_iff[of xs q, symmetric]) (auto simp: defs)
+      done
+    from that have *: "\<forall>p\<in>set ps.
+       (L ! p, bs p, gs p, In a', fs p, rs p, ls' p) \<in> Simple_Network_Language.trans (N p)"
+      unfolding make_combs_alt_def \<open>_ = ps\<close>
+      apply (clarsimp simp: set_map_filter product_lists_set split: if_split_asm)
+      using \<open>set ps \<subseteq> _\<close> unfolding defs
+      by (auto simp: comp_def list_all2_map2 list_all2_same to_map
+          elim!: IN'_D[THEN conjunct1, rotated]
+          )
+    from that have "ps \<noteq> []" "a' \<in> set broadcast"
+      apply (simp_all add: make_combs_alt_def set_map_filter product_lists_set split: if_split_asm)
+      using \<open>_ = ps\<close> \<open>set ps \<subseteq> {0..<n_ps}\<close>
+      by (cases xs; auto dest: IN'_D simp: list_all2_Cons1)+
+    with that have "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+       (\<forall>b g f r l'. (L ! q, b, g, In a', f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)"
+      unfolding make_combs_alt_def
+      by (auto 4 3 dest!: list_all2_set2 IN_I IN'_I
+          simp: defs set_map_filter product_lists_set split: if_split_asm)
+    have "xs = map (\<lambda> p. (p, bs p, gs p, a', fs p, rs p, ls' p)) ps"
+      apply (intro nth_equalityI)
+       apply (simp add: defs; fail)
+      subgoal for i
+        by (cases "xs ! i") (auto 4 4 simp: defs dest: to_map nth_mem)
+      done
+    have "\<forall>p\<in>set ps. check_bexp s (bs p) True"
+    proof
+      fix q assume "q \<in> set ps"
+      from \<open>xs \<in> _\<close> have "distinct (map fst xs)"
+        unfolding make_combs_alt_def
+        by (auto simp: product_lists_set list_all2_map2 to_map
+                 dest!: list_all2_fst_aux[OF list_all2_mono]
+                 split: if_split_asm)
+      from \<open>q \<in> set ps\<close> \<open>_ = ps\<close> have "IN' ! q ! a' \<noteq> []" "q \<noteq> p" "q < n_ps"
+        by auto
+      with that have "the (map_of xs q) \<in> set (IN' ! q ! a')"
+        using set_map_of_compr[OF \<open>distinct (map _ _)\<close>] unfolding make_combs_alt_def
+        apply (clarsimp simp: set_map_filter product_lists_set split: if_split_asm)
+        apply (drule list_all2_set2)
+        apply auto
+        done
+      then obtain a where "(bs q, gs q, a, fs q, rs q, ls' q) \<in> set (IN' ! q ! a')"
+        unfolding defs by atomize_elim (auto split!: prod.splits)
+      then show "check_bexp s (bs q) True"
+        using IN'_D \<open>a' <  _\<close> \<open>set ps \<subseteq> _\<close> \<open>q \<in> set ps\<close> by auto
+    qed
+    show ?thesis
+      by (inst_existentials ps bs gs fs rs ls'; fact)
+  qed
+  let ?f = "\<lambda>(q, b2, g2, a2, f2, r2, l2) (g1, a, r1, L, s).
+                  (g1 @ g2, a, r1 @ r2, L[q := l2], mk_upds s f2)"
+  have ***: "
+    fold ?f (map (\<lambda>p. (p, bs p, gs p, a', fs p, rs p, ls' p)) ps) (g, a, r, L, s)
+    = (g @ concat (map gs ps), a, r @ concat (map rs ps),
+        fold (\<lambda>p L. L[p := ls' p]) ps L, mk_upds s (concat_map fs ps))
+    " for ps bs gs a' fs rs ls' g a r L s
+    by (induction ps arbitrary: g a r L s; simp add: mk_upds_def)
+  have upd_swap:
+    "(fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'] = fold (\<lambda>p L . L[p := ls' p]) ps (L[p := l'])"
+    if "p \<notin> set ps" for ps ls' p l'
+    using that by (induction ps arbitrary: L) (auto simp: list_update_swap)
+  have make_transI:
+    "a' < num_actions \<and> p < n_ps \<and>
+       (g1 @ concat (map gs ps), Broad a', r1 @ concat (map rs ps),
+        (fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l'], s') \<in> set (make_trans a' p)"
+    if 
+      "L \<in> states" and
+      "g = g1 @ concat (map gs ps)" and
+      "a = Broad a'" and
+      "r = r1 @ concat (map rs ps)" and
+      "L' = (fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l']" and
+      "a' \<in> set broadcast" and
+      "(L ! p, b1, g1, Out a', f1, r1, l') \<in> trans (N p)" and
+      "\<forall>p\<in>set ps. (L ! p, bs p, gs p, In a', fs p, rs p, ls' p) \<in> trans (N p)"
+      and
+      "\<forall>q<n_ps. q \<notin> set ps \<and> p \<noteq> q
+        \<longrightarrow> (\<forall>b g f r l'. (L ! q, b, g, In a', f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)" and
+      "p < n_ps" and
+      "set ps \<subseteq> {0..<n_ps}" and
+      "p \<notin> set ps" and
+      "distinct ps" and
+      "sorted ps" and
+      "check_bexp s b1 True" and
+      "\<forall>p\<in>set ps. check_bexp s (bs p) True" and
+      "is_upds s f1 s''" and
+      "is_upds s'' (concat_map fs ps) s'" and
+      "Simple_Network_Language.bounded bounds s'"
+    for a' p b1 g1 bs gs ps r1 rs ls' l' f1 s'' fs
+    using that \<open>dom s = {0..<n_vs}\<close>
+    unfolding make_trans_def
+    apply (clarsimp simp: set_map_filter Let_def split: prod.split)
+    supply [forward2] = action_setD
+    supply [forward4] = is_upds_dom2 is_upds_dom3 is_upds_concatD[rotated 3] OUT_I OUT'_I
+    apply frule2
+    apply simp
+    apply (rule conjI, rule impI)
+    subgoal
+      apply (subst (asm) make_combs_empty, (assumption | simp)+)
+      apply frules_all
+      apply (intro conjI)
+        apply (solves auto)
+       apply (intros add: more_intros)
+         apply (solve_triv | intros add: more_intros UN_I)+
+      subgoal
+        unfolding comp_def by (auto elim!: is_upds.cases)
+      by (auto simp: check_bounded_iff[symmetric] elim: is_upds_NilE)
+    apply (rule impI)
+    apply (frule make_combsD, simp, assumption+)
+    subgoal
+      by (subst (asm) make_combs_empty) (assumption | simp)+
+    apply frules_all
+    apply simp
+    apply (intro conjI)
+      apply (solves auto)
+     apply (intros add: more_intros)
+      apply (solve_triv | intros add: more_intros UN_I)+
+     apply (simp add: *** upd_swap; fail)
+    unfolding check_bounded_iff[symmetric] .
+  have make_transD:
+    "\<exists>s'' b ga f ra l' bs gs fs rs ls' ps.
+         g = ga @ concat (map gs ps) \<and>
+         a = Broad a' \<and>
+         r = ra @ concat (map rs ps) \<and>
+         L' = (fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l'] \<and>
+         a' \<in> set broadcast \<and>
+         (L ! p, b, ga, Out a', f, ra, l') \<in> trans (N p) \<and>
+         (\<forall>p\<in>set ps. (L ! p, bs p, gs p, In a', fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+         (\<forall>q<n_ps.
+             q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+             (\<forall>b g f r l'. (L ! q, b, g, In a', f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)) \<and>
+         p < n_ps \<and>
+         set ps \<subseteq> {0..<n_ps} \<and>
+         p \<notin> set ps \<and>
+         distinct ps \<and>
+         sorted ps \<and>
+         check_bexp s b True \<and> (\<forall>p\<in>set ps. check_bexp s (bs p) True) \<and>
+         is_upds s f s'' \<and> is_upds s'' (concat_map fs ps) s' \<and>
+         bounded bounds s' \<and>
+         filter (\<lambda>i. IN' ! i ! a' \<noteq> [] \<and> i \<noteq> p) [0..<n_ps] = ps"
+    if
+      "L \<in> states" and
+      "a' < num_actions" and
+      "p < n_ps" and
+      "(g, a, r, L', s') \<in> set (make_trans a' p)"
+    for a' p
+    supply [forward2] = action_setD
+    supply [forward3] = is_upds_make_updsI3[rotated] OUT'_D
+    supply [forward4] = is_upds_dom2 is_upds_dom3 is_upds_concatD[rotated 3] OUT_I OUT'_I
+    using that \<open>dom s = {0..<n_vs}\<close>
+    unfolding make_trans_def
+    apply mini_ex
+    apply (clarsimp simp: set_map_filter Let_def split: prod.split if_split_asm)
+    subgoal for b a1 f l'
+      apply (inst_existentials
+      "b :: (nat, int) bexp"
+      "g :: (nat, int) acconstraint list"
+      "f :: (nat \<times> (nat, int) exp) list"
+      "r :: nat list"
+      l'
+      "undefined :: nat \<Rightarrow> (nat, int) bexp"
+      "undefined :: nat \<Rightarrow> (nat, int) acconstraint list"
+      "undefined :: nat \<Rightarrow> (nat \<times> (nat, int) exp) list"
+      "undefined :: nat \<Rightarrow> nat list"
+      "undefined :: nat \<Rightarrow> nat"
+      "[] :: nat list")
+                   apply (solves \<open>auto dest: OUT'_D\<close>)+
+      subgoal
+        by (auto 4 3 simp: filter_empty_conv dest: bspec dest!: make_combs_emptyD OUT'_D IN_I IN'_I)
+            apply (solves \<open>auto dest: OUT'_D\<close>)+
+      subgoal
+        apply (inst_existentials s')
+        subgoal is_upd
+          by (auto intro: is_upds_make_updsI2 dest: OUT'_D)
+        subgoal
+          by simp (rule is_upds.intros)
+        subgoal
+          by (subst check_bounded_iff) (metis OUT'_D is_upds_dom2 is_upd)+
+        subgoal
+          by (rule make_combs_emptyD)
+        done
+      done
+    subgoal for b1 g1 a1 r1 f1 l1' xs
+      apply (drule make_combsI, assumption+)
+      apply frules
+      apply elims
+      apply dedup_prems
+      apply frules_all
+      apply (simp add: *** )
+      apply intros
+                    apply solve_triv+
+                  apply (erule upd_swap[symmetric]; fail)
+                 apply solve_triv+
+                apply (erule bspec; assumption)
+               apply (elims add: allE; intros?; assumption)
+              apply solve_triv+
+      subgoal for ps gs fs rs ls'
+        apply (subst check_bounded_iff)
+        subgoal
+          apply (subst is_upds_dom3)
+             apply (simp add: fold_map comp_def; fail)
+            apply assumption+
+          done
+        subgoal
+          by simp
+        done
+      apply solve_triv
+      done
+    done
+  have make_trans_iff: "
+      (\<exists>s'' aa p b ga f ra l' bs gs fs rs ls' ps.
+          g = ga @ concat (map gs ps) \<and>
+          a = Broad aa \<and>
+          r = ra @ concat (map rs ps) \<and>
+          L' = (fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l'] \<and>
+          aa \<in> set broadcast \<and>
+          (L ! p, b, ga, Out aa, f, ra, l') \<in> Simple_Network_Language.trans (N p) \<and>
+          (\<forall>p\<in>set ps.
+              (L ! p, bs p, gs p, In aa, fs p, rs p, ls' p)
+              \<in> Simple_Network_Language.trans (N p)) \<and>
+          (\<forall>q<n_ps.
+              q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+              (\<forall>b g f r l'.
+                  (L ! q, b, g, In aa, f, r, l') \<notin> trans (N q) \<or> \<not> check_bexp s b True)) \<and>
+          p < n_ps \<and>
+          set ps \<subseteq> {0..<n_ps} \<and>
+          p \<notin> set ps \<and>
+          distinct ps \<and>
+          sorted ps \<and>
+          check_bexp s b True \<and> (\<forall>p\<in>set ps. check_bexp s (bs p) True) \<and>
+          is_upds s f s'' \<and>
+          is_upds s'' (concat_map fs ps) s' \<and>
+          bounded bounds s') =
+      (\<exists>a'\<in>{0..<num_actions}.
+          \<exists>p\<in>{0..<n_ps}. (g, a, r, L', s') \<in> set (make_trans a' p))" (is "?l \<longleftrightarrow> ?r")
+    if "dom s = {0..<n_vs}" "L \<in> states"
+  proof (intro iffI)
+    assume ?l
+    with that show ?r
+      by elims (drule make_transI; (elims; intros)?; solve_triv)
+  next
+    assume ?r
+    with that show ?l
+      apply elims
+      subgoal for a' p
+        apply simp
+        apply (drule make_transD)
+            apply assumption+
+        apply elims
+        apply intros
+                        apply assumption+
+        apply blast+
+        done
+      done
+  qed
+  show "(((L, s), g, a, r, L', s') \<in> trans_broad) =
+        ((g, a, r, L', s') \<in> set (broad_trans_from (L, s)))"
+  proof (cases "get_committed L = []")
+    case True
+    with get_committed_empty_iff[of L] have "\<forall>p<n_ps. L ! p \<notin> committed (N p)"
+      by simp
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_broad \<longleftrightarrow> ((L, s), g, a, r, L', s') \<in>
+      {((L, s), g @ concat (map gs ps), Broad a, r @ concat (map rs ps), (L', s'')) |
+        L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps.
+        a \<in> set broadcast \<and>
+        (l, b, g, Out a, f, r, l') \<in> trans (N p) \<and>
+        (\<forall>p \<in> set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+        (\<forall>q < n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+          \<not> (\<exists>b g f r l'. (L ! q, b, g, In a, f, r, l') \<in> trans (N q) \<and> check_bexp s b True)) \<and>
+        L!p = l \<and>
+        p < length L \<and> set ps \<subseteq> {0..<n_ps} \<and> p \<notin> set ps \<and> distinct ps \<and> sorted ps \<and>
+        check_bexp s b True \<and> (\<forall>p \<in> set ps. check_bexp s (bs p) True) \<and>
+        L' = (fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'] \<and>
+        is_upds s f s' \<and> is_upds s' (concat_map fs ps) s'' \<and>
+        L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }
+    "
+      unfolding trans_broad_def broadcast_def[simplified]
+      by (intro iffI; elims add: CollectE; intros add: CollectI) blast+
+    from True have **:
+      "broad_trans_from (L, s)
+      =  concat (
+           map (\<lambda>a.
+             concat (map (\<lambda>p. make_trans a p) [0..<n_ps])
+           )
+           [0..<num_actions]
+         )"
+      unfolding broad_trans_from_alt_def IN_def OUT_def IN'_def OUT'_def Let_def make_trans_def
+      by simp
+    from \<open>dom s = _\<close> \<open>L \<in> _\<close> \<open>bounded bounds s\<close> show ?thesis
+      unfolding * **
+      apply simp
+      apply (subst make_trans_iff[symmetric])
+        apply simp+
+      apply (intro iffI; elims; intros)
+                          apply (solve_triv | blast)+
+      done
+  next
+    case False
+    with get_committed_empty_iff[of L] have "\<not> (\<forall>p<n_ps. L ! p \<notin> committed (N p))"
+      by simp
+    then have *: "((L, s), g, a, r, L', s') \<in> trans_broad \<longleftrightarrow> ((L, s), g, a, r, L', s') \<in>
+      {((L, s), g @ concat (map gs ps), Broad a, r @ concat (map rs ps), (L', s'')) |
+        L s L' s' s'' a p l b g f r l' bs gs fs rs ls' ps.
+        a \<in> set broadcast \<and>
+        (l, b, g, Out a, f, r, l') \<in> trans (N p) \<and>
+        (\<forall>p \<in> set ps. (L ! p, bs p, gs p, In a, fs p, rs p, ls' p) \<in> trans (N p)) \<and>
+        (l \<in> committed (N p) \<or> (\<exists>p \<in> set ps. L ! p \<in> committed (N p))) \<and>
+        (\<forall>q < n_ps. q \<notin> set ps \<and> p \<noteq> q \<longrightarrow>
+          \<not> (\<exists>b g f r l'. (L ! q, b, g, In a, f, r, l') \<in> trans (N q) \<and> check_bexp s b True)) \<and>
+        L!p = l \<and>
+        p < length L \<and> set ps \<subseteq> {0..<n_ps} \<and> p \<notin> set ps \<and> distinct ps \<and> sorted ps \<and>
+        check_bexp s b True \<and> (\<forall>p \<in> set ps. check_bexp s (bs p) True) \<and>
+        L' = (fold (\<lambda>p L . L[p := ls' p]) ps L)[p := l'] \<and>
+        is_upds s f s' \<and> is_upds s' (concat_map fs ps) s'' \<and>
+        L \<in> states \<and> bounded bounds s \<and> bounded bounds s''
+      }"
+      unfolding trans_broad_def broadcast_def[simplified]
+      by (intro iffI; elims add: CollectE; intros add: CollectI) blast+
+    have committed_iff: "
+      List.map_filter (\<lambda>(p, _). if IN' ! p ! a' = [] then None else Some p) (get_committed L) \<noteq> [p] \<and>
+      List.map_filter (\<lambda>(p, _). if IN' ! p ! a' = [] then None else Some p) (get_committed L) \<noteq> []
+    \<longleftrightarrow> (\<exists>q<n_ps. IN' ! q ! a' \<noteq> [] \<and> q \<noteq> p \<and> L ! q \<in> committed (N q))"
+      for p a'
+    proof -
+      have *: "xs \<noteq> [p] \<and> xs \<noteq> [] \<longleftrightarrow> (\<exists>x \<in> set xs. x \<noteq> p)" if "distinct xs" for xs
+        using that by auto (metis distinct.simps(2) distinct_length_2_or_more revg.elims)
+      show ?thesis
+        by (subst *)
+          (auto
+            intro: distinct_map_filterI get_committed_distinct
+            simp: set_map_filter get_committed_mem_iff split: if_split_asm
+            )
+    qed
+    from False have **:
+      "broad_trans_from (L, s)
+      = concat (
+        map (\<lambda>a.
+          let
+            ins_committed =
+              List.map_filter (\<lambda>(p, _). if IN' ! p ! a \<noteq> [] then Some p else None) (get_committed L)
+          in
+          concat (map (\<lambda>p.
+            if
+              (ins_committed = [p] \<or> ins_committed = [])
+              \<and> \<not> list_ex (\<lambda> (q, _). q = p) (get_committed L)
+            then []
+            else
+              make_trans a p
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+      "
+      unfolding broad_trans_from_alt_def IN_def OUT_def IN'_def OUT'_def make_trans_def
+      unfolding Let_def if_contract
+      apply simp
+      apply (fo_rule if_cong arg_cong2[where f = map] arg_cong[where f = concat] | rule ext)+
+          apply blast+
+      done
+    from \<open>dom s = _\<close> \<open>L \<in> _\<close> \<open>bounded bounds s\<close> show ?thesis
+      unfolding * **
+      apply (simp add: make_trans_iff)
+      apply (intro iffI; elims)
+      subgoal for s'a aa p ga f ra l' gs fs rs ls' ps \<comment> \<open>?l \<open>\<longrightarrow>\<close> ?r\<close>
+        apply (frule make_transI)
+                            apply (assumption | blast)+
+        apply elims
+        apply intros
+         apply (simp; fail)
+        apply (simp add: Let_def)
+        apply (erule disjE)
+        subgoal \<comment> \<open>The process with the outgoing action label is committed\<close>
+          using get_committed_mem_iff[of p "L ! p" L, simplified, symmetric]
+          by (inst_existentials p) (auto simp add: list_ex_iff)
+        apply (erule bexE)
+        subgoal for q \<comment> \<open>One of the processes with an ingoing action label is committed\<close>
+          apply (inst_existentials p)
+           apply assumption
+          apply (rule IntI)
+           apply (simp; fail)
+          apply simp
+          unfolding committed_iff
+          apply (rule disjI1; inst_existentials q; force dest!: IN_I IN'_I)
+          done
+        done
+      subgoal for a' \<comment> \<open>?r \<open>\<longrightarrow>\<close> ?l\<close>
+        apply (clarsimp split: if_split_asm simp: Let_def)
+        apply (drule make_transD[rotated 4])
+            apply assumption+
+        apply elims
+        apply intros
+                         apply assumption+
+                   apply (erule bspec; assumption)
+        subgoal for p s'' g' f r' l' gs fs rs ls' ps
+          unfolding committed_iff by (auto simp: get_committed_mem_iff list_ex_iff)
+        apply blast+
+        done
+      done
+  qed
+qed
+
+
+paragraph \<open>Refinement of the State Implementation\<close>
+
+definition state_rel :: "(nat \<rightharpoonup> int) \<Rightarrow> int list \<Rightarrow> bool"
+  where
+  "state_rel s xs \<equiv> length xs = n_vs \<and> dom s = {0..<n_vs} \<and> (\<forall>i < n_vs. xs ! i = the (s i))"
+
+definition loc_rel where
+  "loc_rel \<equiv> {((L', s'), (L, s)) | L s L' s'. L' = L \<and> length L = n_ps \<and> state_rel s s'}"
+
+lemma state_impl_abstract:
+  "\<exists>L s. ((Li, si), (L, s)) \<in> loc_rel" if "length Li = n_ps" "length si = n_vs"
+  using that unfolding loc_rel_def state_rel_def
+  by (inst_existentials Li "\<lambda>i. if i < n_vs then Some (si ! i) else None")(auto split: if_split_asm)
+
+lemma state_rel_left_unique:
+  "l \<in> states' \<Longrightarrow> (li, l) \<in> loc_rel \<Longrightarrow> (li', l) \<in> loc_rel \<Longrightarrow> li' = li"
+  unfolding loc_rel_def state_rel_def by (auto intro: nth_equalityI)
+
+lemma state_rel_right_unique:
+  "l \<in> states' \<Longrightarrow> l' \<in> states' \<Longrightarrow> (li, l) \<in> loc_rel \<Longrightarrow> (li, l') \<in> loc_rel \<Longrightarrow> l' = l"
+  unfolding loc_rel_def state_rel_def
+  apply clarsimp
+  apply (rule ext)
+  subgoal premises prems for L s s'1 s'2 x
+  proof -
+    show "s'1 x = s x"
+    proof (cases "x < n_vs")
+      case True
+      then have "x \<in> dom s'1" "x \<in> dom s"
+        using prems by auto
+      with \<open>x < n_vs\<close> show ?thesis
+        using prems(9) by auto
+    next
+      case False
+      then have "x \<notin> dom s'1" "x \<notin> dom s"
+        using prems by auto
+      then show ?thesis
+        by (auto simp: dom_def)
+    qed
+  qed
+  done
+
+end (* Anonymous context for simp setup *)
+
+end (* Simple Network Impl nat *)
+
+fun bvali :: "_ \<Rightarrow> (nat, 'b :: linorder) bexp \<Rightarrow> bool" and evali where
+  "bvali s bexp.true = True" |
+  "bvali s (not e) \<longleftrightarrow> \<not> bvali s e" |
+  "bvali s (and e1 e2) \<longleftrightarrow> bvali s e1 \<and> bvali s e2" |
+  "bvali s (bexp.or e1 e2) \<longleftrightarrow> bvali s e1 \<or> bvali s e2" |
+  "bvali s (imply e1 e2) \<longleftrightarrow> bvali s e1 \<longrightarrow> bvali s e2" |
+  "bvali s (eq i x) \<longleftrightarrow> evali s i = evali s x" |
+  "bvali s (le i x) \<longleftrightarrow> evali s i \<le> evali s x" |
+  "bvali s (lt i x) \<longleftrightarrow> evali s i < evali s x" |
+  "bvali s (ge i x) \<longleftrightarrow> evali s i \<ge> evali s x" |
+  "bvali s (gt i x) \<longleftrightarrow> evali s i > evali s x"
+| "evali s (const c) = c"
+| "evali s (var x)   = s ! x"
+| "evali s (if_then_else b e1 e2) = (if bvali s b then evali s e1 else evali s e2)"
+| "evali s (binop f e1 e2) = f (evali s e1) (evali s e2)"
+| "evali s (unop f e) = f (evali s e)"
+
+definition mk_updsi ::
+  "int list \<Rightarrow> (nat \<times> (nat, int) exp) list \<Rightarrow> int list" where
+  "mk_updsi s upds = fold (\<lambda>(x, upd) s. s[x := evali s upd]) upds s"
+
+context Simple_Network_Impl_nat_defs
+begin
+
+definition
+  "check_boundedi s =
+    (\<forall>x < length s. fst (bounds_map x) \<le> s ! x \<and> s ! x \<le> snd (bounds_map x))"
+
+definition
+  "states'_memi \<equiv> \<lambda>(L, s). L \<in> states \<and> length s = n_vs \<and> check_boundedi s"
+
+definition
+  "int_trans_from_loc_impl p l L s \<equiv>
+    let trans = trans_i_map p l
+    in
+    List.map_filter (\<lambda> (b, g, a, f, r, l').
+      let s' = mk_updsi s f in
+        if bvali s b \<and> check_boundedi s' then Some (g, Internal a, r, (L[p := l'], s'))
+        else None
+    ) trans"
+
+definition
+  "int_trans_from_vec_impl pairs L s \<equiv>
+    concat (map (\<lambda>(p, l). int_trans_from_loc_impl p l L s) pairs)"
+
+definition
+  "int_trans_from_all_impl L s \<equiv>
+    concat (map (\<lambda>p. int_trans_from_loc_impl p (L ! p) L s) [0..<n_ps])"
+
+definition
+  "int_trans_impl \<equiv> \<lambda> (L, s).
+    let pairs = get_committed L in
+    if pairs = []
+    then int_trans_from_all_impl L s
+    else int_trans_from_vec_impl pairs L s
+  "
+
+definition
+  "pairs_by_action_impl L s OUT IN \<equiv>
+  concat (
+    map (\<lambda> (p, b1, g1, a1, f1, r1, l1).
+      List.map_filter (\<lambda> (q, b2, g2, a2, f2, r2, l2).
+        if p = q then None else
+          let s' = mk_updsi (mk_updsi s f1) f2 in
+          if bvali s b1 \<and> bvali s b2 \<and> check_boundedi s'
+          then Some (g1 @ g2, Bin a1, r1 @ r2, (L[p := l1, q := l2], s'))
+          else None
+    ) OUT) IN)
+  "
+
+definition
+  "bin_trans_from_impl \<equiv> \<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In =  all_actions_by_state trans_in_map L;
+      Out = all_actions_by_state trans_out_map L
+    in
+    if pairs = [] then
+      concat (map (\<lambda>a. pairs_by_action_impl L s (Out ! a) (In ! a)) bin_actions)
+    else
+      let
+        In2  = all_actions_from_vec trans_in_map pairs;
+        Out2 = all_actions_from_vec trans_out_map pairs
+      in
+        concat (map (\<lambda>a. pairs_by_action_impl L s (Out ! a) (In2 ! a)) bin_actions)
+      @ concat (map (\<lambda>a. pairs_by_action_impl L s (Out2 ! a) (In ! a)) bin_actions)
+    "
+
+definition
+  "compute_upds init \<equiv> List.map_filter (\<lambda>comb.
+  let (g, a, r, L', s) =
+    fold
+      (\<lambda>(q, b2, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+        (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_upds s f2))
+      )
+      comb
+      init
+  in if check_bounded s then Some (g, a, r, L', s) else None
+)"
+
+definition
+  "compute_upds_impl init \<equiv> List.map_filter (\<lambda>comb.
+  let (g, a, r, L', s) =
+    fold
+      (\<lambda>(q, b2, g2, a2, f2, r2, l2) (g1, a, r1, (L, s)).
+        (g1 @ g2, a, r1 @ r2, (L[q := l2], mk_updsi s f2))
+      )
+      comb
+      init
+  in if check_boundedi s then Some (g, a, r, L', s) else None
+)"
+
+definition trans_from where
+  "trans_from st = int_trans_from st @ bin_trans_from st @ broad_trans_from st"
+
+definition
+  "broad_trans_from_impl \<equiv> \<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+      Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps];
+      In  = map (map (filter (\<lambda>(b, _). bvali s b))) In;
+      Out = map (map (filter (\<lambda>(b, _). bvali s b))) Out
+    in
+    if pairs = [] then
+      concat (
+        map (\<lambda>a.
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+                compute_upds_impl init combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    else
+      concat (
+        map (\<lambda>a.
+          let
+            ins_committed =
+              List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs;
+            always_committed = (length ins_committed > 1)
+          in
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else if
+              \<not> always_committed \<and> (ins_committed = [p] \<or> ins_committed = [])
+              \<and> \<not> list_ex (\<lambda> (q, _). q = p) pairs
+            then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+                compute_upds_impl init combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    "
+
+definition trans_impl where
+  "trans_impl st = int_trans_impl st @ bin_trans_from_impl st @ broad_trans_from_impl st"
+
+end (* Simple Network Impl nat defs *)
+
+
+context Simple_Network_Impl_nat
+begin
+
+lemma bval_bvali:
+  "state_rel s si \<Longrightarrow> \<forall>x \<in> vars_of_bexp b. x \<in> dom s \<Longrightarrow> bval (the o s) b = bvali si b"
+and eval_evali:
+  "state_rel s si \<Longrightarrow> \<forall>x \<in> vars_of_exp e. x \<in> dom s \<Longrightarrow> eval (the o s) e = evali si e"
+  by (induction b and e) (auto simp: state_rel_def)
+
+lemma mk_upds_mk_updsi:
+  "state_rel (mk_upds s upds) (mk_updsi si upds)"
+  if assms: "state_rel s si" "\<forall> (_, e) \<in> set upds. \<forall>x \<in> vars_of_exp e. x < n_vs"
+    "\<forall> (x, e) \<in> set upds. x < n_vs"
+proof -
+  have upd_stepI: "state_rel (mk_upd (x, e) s') (si'[x := evali si' e])"
+    if "state_rel s' si'" "\<forall>x \<in> vars_of_exp e. x < n_vs" "x < n_vs"
+    for s' si' x e
+    using that assms unfolding mk_upd_def state_rel_def by (auto simp: state_rel_def eval_evali)
+  from assms show ?thesis
+  proof (induction upds arbitrary: s si)
+    case Nil
+    then show ?case
+      by (simp add: mk_upds_def mk_updsi_def)
+  next
+    case (Cons upd upds)
+    then show ?case
+      by (simp add: mk_upds_def mk_updsi_def split: prod.splits) (rprem, auto intro!: upd_stepI)
+  qed
+qed
+
+lemma check_bounded_check_boundedi:
+  "check_bounded s = check_boundedi si" if "state_rel s si"
+  using that unfolding check_bounded_def check_boundedi_def state_rel_def by auto
+
+definition
+  "valid_upd \<equiv> \<lambda>(x, e). x < n_vs \<and> (\<forall>x \<in> vars_of_exp e. x < n_vs)"
+
+definition
+  "valid_check b \<equiv> (\<forall>x \<in> vars_of_bexp b. x < n_vs)"
+
+context includes lifting_syntax begin
+notation rel_prod (infixr "\<times>R" 56)
+
+definition is_at_least_equality where
+  "is_at_least_equality R \<equiv> \<forall>x y. R x y \<longrightarrow> x = y" for R
+
+named_theorems is_at_least_equality
+
+lemma [is_at_least_equality]:
+  "is_at_least_equality (=)"
+  by (simp add: is_at_least_equality_def)
+
+lemma [is_at_least_equality]:
+  "is_at_least_equality R" if "is_equality R" for R
+  using that by (simp add: is_at_least_equality_def is_equality_def)
+
+lemma [is_at_least_equality]:
+  "is_at_least_equality (eq_onp P)"
+  by (simp add: is_at_least_equality_def eq_onp_def)
+
+lemma is_at_least_equality_list_all2[is_at_least_equality]:
+  "is_at_least_equality (list_all2 R)" if "is_at_least_equality R" for R
+  using that unfolding is_at_least_equality_def
+  by (auto simp: list.rel_eq dest: list_all2_mono[where Q = "(=)"])
+
+lemma is_at_least_equality_rel_prod[is_at_least_equality]:
+  "is_at_least_equality (R1 \<times>R R2)"
+  if "is_at_least_equality R1" "is_at_least_equality R2" for R1 R2
+  using that unfolding is_at_least_equality_def by auto
+
+lemma is_at_least_equality_cong1:
+  "(S1 ===> (=)) f f" if "is_at_least_equality S1" "is_at_least_equality S2" for S1 f
+  using that unfolding is_at_least_equality_def by (intro rel_funI) auto
+
+lemma is_at_least_equality_cong2:
+  "(S1 ===> S2 ===> (=)) f f" if "is_at_least_equality S1" "is_at_least_equality S2" for S1 S2 f
+  using that unfolding is_at_least_equality_def by (intro rel_funI) auto
+
+lemma is_at_least_equality_cong3:
+  "(S1 ===> S2 ===> S3 ===> (=)) f f"
+  if "is_at_least_equality S1" "is_at_least_equality S2" "is_at_least_equality S3" for S1 S2 S3 f
+  using that unfolding is_at_least_equality_def by (intro rel_funI) force
+
+lemma is_at_least_equality_Let:
+  "(S ===> ((=) ===> R) ===> R) Let Let" if "is_at_least_equality S" for R
+  using that unfolding is_at_least_equality_def
+  by (intro rel_funI) (erule Let_transfer[THEN rel_funD, THEN rel_funD, rotated], auto)
+
+lemma map_transfer_length:
+  fixes R S n
+  shows
+    "((R ===> S)
+      ===> (\<lambda>x y. list_all2 R x y \<and> length x = n)
+      ===> (\<lambda>x y. list_all2 S x y \<and> length x = n))
+    map map"
+  apply (intro rel_funI conjI)
+  apply (erule list.map_transfer[THEN rel_funD, THEN rel_funD], erule conjunct1)
+  apply simp
+  done
+
+lemma upt_0_transfer:
+  "(eq_onp (\<lambda>x. x = 0) ===> eq_onp (\<lambda>x. x = n) ===> list_all2 (eq_onp (\<lambda>x. x < n))) upt upt" for n
+  apply (intro rel_funI upt_transfer_upper_bound[THEN rel_funD, THEN rel_funD])
+   apply (assumption | erule eq_onp_to_eq)+
+  done
+
+lemma upt_length_transfer:
+  "(eq_onp (\<lambda>x. x = 0) ===> eq_onp (\<lambda>x. x = n)
+    ===> (\<lambda> x y. list_all2 (eq_onp (\<lambda>x. x < n)) x y \<and> length x = n)) upt upt" for n
+  apply (intro rel_funI conjI upt_0_transfer[THEN rel_funD, THEN rel_funD], assumption+)
+  apply (simp add: eq_onp_def)
+  done
+
+lemma case_prod_transfer_strong:
+  fixes A B C
+  assumes "\<And> x y. A1 x y \<Longrightarrow> A x y" "\<And> x y. B1 x y \<Longrightarrow> B x y"
+  shows "((A ===> B ===> C) ===> A1 \<times>R B1 ===> C) case_prod case_prod"
+  apply (intro rel_funI)
+  apply clarsimp
+  apply (drule assms)+
+  apply (drule (1) rel_funD)+
+  apply assumption
+  done
+
+lemma concat_transfer_strong:
+  fixes A B C
+  assumes "\<And>x y. A x y \<Longrightarrow> B x y" "\<And> x y. C x y \<Longrightarrow> list_all2 (list_all2 A) x y"
+  shows "(C ===> list_all2 B) concat concat"
+  apply (intro rel_funI concat_transfer[THEN rel_funD])
+  apply (drule assms)
+  apply (erule list_all2_mono)
+  apply (erule list_all2_mono)
+  apply (erule assms)
+  done
+
+lemma map_transfer_strong:
+  fixes A B C
+  assumes "\<And>xs ys. C xs ys \<Longrightarrow> list_all2 A xs ys"
+  shows "((A ===> B) ===> C ===> list_all2 B) map map"
+  apply (intro rel_funI)
+  apply (erule list.map_transfer[THEN rel_funD, THEN rel_funD])
+  apply (erule assms)
+  done
+
+lemma list_update_transfer':
+  fixes A :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
+  shows "(list_all2 A ===> eq_onp (\<lambda>i. i<n_ps) ===> A ===> list_all2 A) list_update list_update"
+  apply (intro rel_funI)
+  apply (rule list_update_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD])
+    apply (auto simp: eq_onp_def)
+  done
+
+lemma list_update_transfer'':
+  fixes A :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and n
+  shows "((\<lambda> x y. list_all2 A x y \<and> length x = n) ===> eq_onp (\<lambda>i. i<n) ===> A
+    ===> (\<lambda> x y. list_all2 A x y \<and> length x = n)) list_update list_update"
+  apply (intro rel_funI conjI)
+  subgoal
+    apply (erule conjE)
+    apply (rule List.list_update_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD])
+    apply (assumption | elim eq_onp_to_eq)+
+    done
+  apply simp
+  done
+
+lemma list_update_transfer''':
+  fixes A :: "'a \<Rightarrow> 'b \<Rightarrow> bool" and n
+  shows "((\<lambda> x y. list_all2 A x y \<and> length x = n) ===> (=) ===> A
+    ===> (\<lambda> x y. list_all2 A x y \<and> length x = n)) list_update list_update"
+  apply (intro rel_funI conjI)
+  subgoal
+    apply (erule conjE)
+    apply (rule List.list_update_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD])
+    apply (assumption | elim eq_onp_to_eq)+
+    done
+  apply simp
+  done
+
+lemma fold_transfer_strong:
+  fixes A B
+  assumes "\<And>x y. A1 x y \<Longrightarrow> A x y" "\<And>x y. B1 x y \<Longrightarrow> B x y" "\<And>x y. B x y \<Longrightarrow> B2 x y"
+    "\<And>x y. B x y \<Longrightarrow> B3 x y"
+  shows "((A ===> B2 ===> B1) ===> list_all2 A1 ===> B ===> B3) fold fold"
+  apply (intro rel_funI, rule assms)
+  apply (rule fold_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD])
+    apply (intro rel_funI)
+    apply (drule rel_funD, erule assms)
+    apply (drule rel_funD, erule assms, erule assms)
+   apply assumption+
+  done
+
+lemma bval_bvali_transfer[transfer_rule]:
+  "(state_rel ===> eq_onp valid_check ===> (=)) (\<lambda> s. bval (the o s)) bvali"
+  by (intro rel_funI) (auto simp: eq_onp_def valid_check_def state_rel_def intro!: bval_bvali)
+
+lemma mk_upds_mk_updsi_transfer[transfer_rule]:
+  "(state_rel ===> list_all2 (eq_onp valid_upd) ===> state_rel) mk_upds mk_updsi"
+  apply (intro rel_funI)
+  subgoal for x y upds upds'
+    apply (subgoal_tac "upds' = upds")
+     apply simp
+     apply (rule mk_upds_mk_updsi)
+       apply assumption
+    subgoal
+      by (smt case_prodI2 case_prod_conv eq_onp_def list_all2_same valid_upd_def)
+    subgoal
+      by (smt case_prodE case_prod_conv eq_onp_def list_all2_same valid_upd_def)
+    subgoal
+      by (metis eq_onp_to_eq list.rel_eq_onp)
+    done
+  done
+
+lemma check_bounded_transfer[transfer_rule]:
+  "(state_rel ===> (=)) check_bounded check_boundedi"
+  by (simp add: check_bounded_check_boundedi rel_funI)
+
+lemma trans_map_transfer:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===>
+    list_all2 (
+      eq_onp valid_check \<times>R (=) \<times>R eq_onp (pred_act (\<lambda>x. x < num_actions))
+      \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)
+   )) trans_map trans_map"
+  apply (intro rel_funI, simp add: eq_onp_def, intro list.rel_refl_strong)
+  apply clarsimp
+  apply (auto 4 4 dest!: trans_mapD dest: action_setD var_setD var_setD2 intro: list.rel_refl_strong
+                  simp: valid_upd_def valid_check_def
+        )
+  done
+
+lemma trans_map_transfer':
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===>
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))
+  ) trans_map trans_map"
+  apply (intro rel_funI, simp add: eq_onp_def, intro list.rel_refl_strong)
+  apply clarsimp
+  apply (intro conjI list.rel_refl_strong)
+    apply (auto 4 4 dest: var_setD trans_mapD var_setD2 simp: valid_upd_def valid_check_def)
+  done
+
+lemma map_filter_transfer[transfer_rule]:
+  "((S ===> rel_option R) ===> list_all2 S ===> list_all2 R) List.map_filter List.map_filter"
+  unfolding map_filter_def
+  apply (clarsimp intro!: rel_funI)
+  subgoal for f g xs ys
+  apply (rule list.map_transfer[THEN rel_funD, THEN rel_funD, of "\<lambda> x y. f x \<noteq> None \<and> S x y"])
+   apply (rule rel_funI)
+  subgoal for a b
+    apply (cases "f a")
+    apply (auto simp: option_rel_Some1 option_rel_Some2 dest!: rel_funD)
+    done
+  subgoal
+    apply rotate_tac
+    apply (induction rule: list_all2_induct)
+     apply (auto dest: rel_funD)
+    done
+  done
+  done
+
+lemma trans_i_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===>
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))
+   ) trans_i_map trans_i_map"
+  supply [transfer_rule] = trans_map_transfer'
+  unfolding trans_i_map_def by transfer_prover
+
+lemma int_trans_from_loc_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) ===> state_rel
+    ===> list_all2((=) \<times>R (=) \<times>R (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel))
+  int_trans_from_loc int_trans_from_loc_impl"
+  supply [transfer_rule] = list_update_transfer''
+  unfolding int_trans_from_loc_def int_trans_from_loc_impl_def Let_def by transfer_prover
+
+lemma n_ps_transfer:
+  "eq_onp (\<lambda>x. x = n_ps) n_ps n_ps"
+  by (simp add: eq_onp_def)
+
+lemma zero_nat_transfer:
+  "(=) 0 (0::nat)"
+  ..
+
+lemma int_trans_from_all_transfer[transfer_rule]:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) ===> state_rel
+    ===> list_all2((=) \<times>R (=) \<times>R (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel))
+  int_trans_from_all int_trans_from_all_impl"
+  supply [transfer_rule] = zero_nat_transfer n_ps_transfer
+  unfolding int_trans_from_all_def int_trans_from_all_impl_def Let_def
+  by transfer_prover
+
+lemma int_trans_from_vec_transfer[transfer_rule]:
+  "(list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R (=)) ===> (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps)
+    ===> state_rel
+    ===> list_all2((=) \<times>R (=) \<times>R (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel))
+  int_trans_from_vec int_trans_from_vec_impl"
+  unfolding int_trans_from_vec_def int_trans_from_vec_impl_def Let_def
+  by transfer_prover
+
+private definition R where "R \<equiv> (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps)"
+
+lemma get_committed_transfer[transfer_rule]:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) ===> list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R (=)))
+    get_committed get_committed"
+proof -
+  have [transfer_rule]:
+    "R automata automata"
+    unfolding R_def by (simp add: n_ps_def list.rel_eq)
+  show ?thesis
+  supply [transfer_rule] = zero_nat_transfer n_ps_transfer
+  unfolding get_committed_def
+  unfolding Let_def
+  apply transfer_prover_start
+  using [[goals_limit=15]]
+  prefer 8
+  apply transfer_step
+                prefer 8
+  unfolding R_def
+       apply transfer_step+
+  apply transfer_prover
+  done
+qed
+
+lemma eq_transfer:
+  "(list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R (=)) ===> list_all2 (=) ===> (=)) (=) (=)"
+  unfolding eq_onp_def
+  apply (intro rel_funI)
+  apply (drule list_all2_mono[where Q = "(=)"])
+   apply (auto simp add: list.rel_eq rel_prod.simps)
+  done
+
+lemma int_trans_from_transfer:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+  ===> list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  int_trans_from int_trans_impl"
+  supply [transfer_rule] = eq_transfer
+  unfolding int_trans_impl_def int_trans_from_def Let_def
+  by transfer_prover
+
+lemma pairs_by_action_transfer[transfer_rule]:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n) ===> state_rel ===>
+    list_all2 ((=) \<times>R eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=) \<times>R (=))
+    ===>
+    list_all2 ((=) \<times>R eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=) \<times>R (=))
+    ===>
+    list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n) \<times>R state_rel))
+   pairs_by_action pairs_by_action_impl"
+  supply [transfer_rule] = list_update_transfer'''
+  unfolding pairs_by_action_def pairs_by_action_impl_def by transfer_prover
+
+lemmas rel_elims =
+  rel_prod.cases
+  rel_funD
+
+lemmas rel_intros =
+  rel_funI
+
+lemma pairs_by_action_transfer':
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n) ===> state_rel ===>
+    list_all2 (B \<times>R eq_onp valid_check \<times>R C \<times>R D \<times>R list_all2 (eq_onp valid_upd) \<times>R E \<times>R F) ===>
+    list_all2 (B \<times>R eq_onp valid_check \<times>R C \<times>R D \<times>R list_all2 (eq_onp valid_upd) \<times>R E \<times>R F) ===>
+    list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n) \<times>R state_rel))
+   pairs_by_action pairs_by_action_impl"
+  if "\<And>x y. B x y \<Longrightarrow> x = y"
+    "\<And>x y. C x y \<Longrightarrow> x = y" "\<And>x y. D x y \<Longrightarrow> x = y"
+    "\<And>x y. E x y \<Longrightarrow> x = y" "\<And>x y. F x y \<Longrightarrow> x = y"
+  for B C D E F
+  apply (intro rel_funI)
+  apply (rule pairs_by_action_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD, THEN rel_funD])
+     apply (assumption | erule that list_all2_mono prod.rel_mono_strong)+
+  done
+
+lemma trans_in_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=)
+    ===> list_all2 (
+    eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>a. a < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))
+   ) trans_in_map trans_in_map"
+  supply [transfer_rule] = trans_map_transfer[folded act.rel_eq_onp]
+  unfolding trans_in_map_def by transfer_prover
+
+lemma trans_in_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=)
+    ===> list_all2 (eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))
+   ) trans_in_map trans_in_map"
+  unfolding trans_in_map_def oops
+
+lemma trans_out_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=)
+    ===> list_all2 (
+    eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>a. a < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)
+   )) trans_out_map trans_out_map"
+  supply [transfer_rule] = trans_map_transfer[folded act.rel_eq_onp]
+  unfolding trans_out_map_def by transfer_prover
+
+lemma trans_out_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> list_all2 (
+    eq_onp valid_check \<times>R (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+  trans_out_map trans_out_map"
+  unfolding trans_out_map_def oops
+
+lemma actions_by_state_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i < n_ps) ===>
+    list_all2 ((=) \<times>R (=) \<times>R eq_onp (\<lambda>i. i < n) \<times>R (=)) ===>
+    (\<lambda> x y. list_all2 (list_all2 (eq_onp (\<lambda>i. i<n_ps) \<times>R (=) \<times>R (=) \<times>R eq_onp (\<lambda>x. x < n) \<times>R (=))) x y \<and> length x = n) ===>
+    (\<lambda> x y. list_all2 (list_all2 (eq_onp (\<lambda>i. i<n_ps) \<times>R (=) \<times>R (=) \<times>R eq_onp (\<lambda>x. x < n) \<times>R (=))) x y \<and> length x = n)
+  )
+  actions_by_state actions_by_state" for n
+  supply [transfer_rule] = list_update_transfer''
+  unfolding actions_by_state_def by transfer_prover
+
+lemma actions_by_state_transfer'[transfer_rule]:
+  "(
+    eq_onp (\<lambda>i. i < n_ps) ===>
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>i. i < n) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)) ===>
+    (\<lambda> x y. list_all2 (list_all2 (
+        eq_onp (\<lambda>i. i<n_ps) \<times>R eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < n)
+        \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y \<and> length x = n) ===>
+    (\<lambda> x y. list_all2 (list_all2 (
+        eq_onp (\<lambda>i. i<n_ps) \<times>R eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < n)
+        \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y \<and> length x = n)
+  )
+  actions_by_state actions_by_state"
+  for n
+  supply [transfer_rule] = list_update_transfer''
+  unfolding actions_by_state_def by transfer_prover
+
+lemma transfer_consts:
+  "(eq_onp (\<lambda>x. x = num_actions)) num_actions num_actions" "(eq_onp (\<lambda>x. x = 0)) (0::nat) 0"
+  "(eq_onp (\<lambda>x. x = n_ps)) n_ps n_ps" 
+  by (auto simp: eq_onp_def)
+
+lemma all_actions_by_state_transfer[transfer_rule]:
+  "
+  (
+    (eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>i. i < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+    ===>
+    (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps)
+    ===>
+    (\<lambda> x y. list_all2 (list_all2 (eq_onp (\<lambda>i. i<n_ps) \<times>R eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>i. i<num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y \<and> length x = num_actions)
+  )
+  all_actions_by_state all_actions_by_state"
+  supply [transfer_rule] = map_transfer_length upt_0_transfer upt_length_transfer transfer_consts
+    n_ps_transfer
+  unfolding all_actions_by_state_def by transfer_prover
+
+lemma all_actions_from_vec_transfer[transfer_rule]:
+  "
+  (
+    (eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>i. i < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+    ===>
+    list_all2 (eq_onp (\<lambda>i. i < n_ps) \<times>R (=))
+    ===>
+    (\<lambda> x y. list_all2 (list_all2 (eq_onp (\<lambda>i. i<n_ps) \<times>R eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>i. i<num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y \<and> length x = num_actions)
+  )
+  all_actions_from_vec all_actions_from_vec"
+  supply [transfer_rule] = map_transfer_length upt_length_transfer transfer_consts
+  unfolding all_actions_from_vec_def all_actions_from_vec_def by transfer_prover
+
+lemma bin_actions_transfer[transfer_rule]:
+  "(list_all2 (eq_onp (\<lambda>x. x < num_actions))) bin_actions bin_actions"
+proof -
+  have *: "list_all2 (eq_onp (\<lambda>x. x < num_actions)) [0..<num_actions] [0..<num_actions]"
+    by (rule list.rel_refl_strong) (auto simp: eq_onp_def)
+  show ?thesis
+    unfolding bin_actions_def
+    by (rule filter_transfer[THEN rel_funD, THEN rel_funD, OF _ *]) (auto simp: eq_onp_def)
+qed
+
+lemma Let_transfer_bin_aux:
+  "((\<lambda>x y. list_all2 (list_all2
+    (eq_onp (\<lambda>i. i < n_ps) \<times>R eq_onp valid_check \<times>R list_all2 (rel_acconstraint (=) (=)) \<times>R
+         eq_onp (\<lambda>i. i < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))) x y
+    \<and> length x = num_actions) ===>
+   ((\<lambda>x y. list_all2 (list_all2
+    ((=) \<times>R eq_onp valid_check \<times>R list_all2 (rel_acconstraint (=) (=)) \<times>R
+         (=) \<times>R list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))) x y
+    \<and> length x = num_actions) ===>
+   list_all2
+    (list_all2 (rel_acconstraint (=) (=)) \<times>R rel_label (=) \<times>R
+     list_all2 (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)) ===>
+  list_all2
+    (list_all2 (rel_acconstraint (=) (=)) \<times>R rel_label (=) \<times>R
+     list_all2 (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel))
+  Let Let"
+  unfolding Let_def
+  by (intro rel_funI, drule rel_funD)
+     (auto simp: eq_onp_def elim!: list_all2_mono prod.rel_mono_strong)
+
+lemma bin_trans_from_transfer:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+  ===> list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  bin_trans_from bin_trans_from_impl"
+  unfolding bin_trans_from_impl_def bin_trans_from_def
+  supply [transfer_rule] =
+    map_transfer_length upt_length_transfer transfer_consts eq_transfer Let_transfer_bin_aux
+  by transfer_prover
+
+(*
+definition
+  "trans_in_broad_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (g, a, m, l').
+      case a of In a \<Rightarrow> if a \<in> set broadcast then Some (g, a, m, l') else None | _ \<Rightarrow> None)
+    (trans_map i j)"
+
+definition
+  "trans_out_broad_map i j \<equiv>
+    List.map_filter
+      (\<lambda> (g, a, m, l').
+      case a of Out a \<Rightarrow> if a \<in> set broadcast then Some (g, a, m, l') else None | _ \<Rightarrow> None)
+    (trans_map i j)"
+*)
+
+(*
+definition
+  "actions_by_state' xs \<equiv>
+    fold (\<lambda> t acc. acc[fst (snd t) := t # (acc ! fst (snd t))]) xs (repeat [] num_actions)"
+*)
+
+lemma trans_map_transfer'':
+  "((=) ===> (=) ===>
+  list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (pred_act (\<lambda>x. x < num_actions)) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+  trans_map trans_map"
+  apply (intro rel_funI, simp add: eq_onp_def, intro list.rel_refl_strong)
+  apply clarsimp
+  apply (auto 4 4 dest!: trans_mapD dest: action_setD var_setD intro: list.rel_refl_strong simp: valid_upd_def)
+  oops
+
+lemma compute_upds_transfer:
+  "(
+  (list_all2 (=) \<times>R (=) \<times>R list_all2 (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n) \<times>R state_rel)
+  ===> list_all2 (list_all2
+        ((=) \<times>R eq_onp valid_check \<times>R list_all2 (=) \<times>R (=) \<times>R list_all2 (eq_onp valid_upd)
+        \<times>R list_all2 (=) \<times>R (=))) ===>
+  list_all2 (
+    list_all2 (=) \<times>R (=) \<times>R list_all2 (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n) \<times>R state_rel
+  )) compute_upds compute_upds_impl"
+  supply [transfer_rule] = list_update_transfer'''
+  unfolding compute_upds_def compute_upds_impl_def by transfer_prover
+
+lemma in_transfer:
+  "(eq_onp (\<lambda>x. x < num_actions) ===> (=) ===> (=)) (\<in>) (\<in>)"
+  by (intro rel_funI, rule member_transfer[of "(=)", THEN rel_funD, THEN rel_funD])
+     (auto simp: eq_onp_def rel_set_eq intro: bi_unique_eq)
+
+lemma trans_in_broad_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+  trans_in_broad_map trans_in_broad_map"
+  supply [transfer_rule] = trans_map_transfer[folded act.rel_eq_onp] in_transfer
+  unfolding trans_in_broad_map_def by transfer_prover
+
+lemma trans_out_broad_map_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=) ===> list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=)))
+  trans_out_broad_map trans_out_broad_map"
+  supply [transfer_rule] = trans_map_transfer[folded act.rel_eq_onp] in_transfer
+  unfolding trans_out_broad_map_def by transfer_prover
+
+text \<open>We are using the ``equality version'' of parametricty for @{term "(!)"} here.\<close>
+lemma actions_by_state'_transfer[transfer_rule]:
+  "(list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))
+  ===> (\<lambda> x y. list_all2 (
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y
+    \<and> length x = num_actions
+  ))
+  actions_by_state' actions_by_state'"
+  supply [transfer_rule] = transfer_consts
+    upt_length_transfer
+    map_transfer_length
+    list_update_transfer''
+  unfolding actions_by_state'_def by transfer_prover
+
+lemma trans_in_broad_grouped_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=)
+  ===> (\<lambda> x y. list_all2 (
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y
+    \<and> length x = num_actions
+  )) trans_in_broad_grouped trans_in_broad_grouped"
+  unfolding trans_in_broad_grouped_def by transfer_prover
+
+lemma trans_out_broad_grouped_transfer[transfer_rule]:
+  "(eq_onp (\<lambda>i. i<n_ps) ===> (=)
+  ===> (\<lambda> x y. list_all2 (
+    list_all2 (eq_onp valid_check \<times>R (=) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R (=))) x y
+    \<and> length x = num_actions
+  )) trans_out_broad_grouped trans_out_broad_grouped"
+  unfolding trans_out_broad_grouped_def by transfer_prover
+
+lemma make_combs_transfer:
+  fixes R
+  assumes "\<And>x y. R x y \<Longrightarrow> x = y"
+  shows
+    "(eq_onp (\<lambda>x. x < n_ps)
+  ===> eq_onp (\<lambda>x. x < num_actions)
+  ===> (\<lambda>x y. list_all2 (\<lambda>x y. list_all2 (list_all2 R) x y \<and> length x = num_actions) x y
+        \<and> length x = n_ps)
+  ===> list_all2 (list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R R)))
+  make_combs make_combs"
+proof -
+  have [transfer_rule]:
+    "(eq_onp (\<lambda>x. x < n_ps) ===> eq_onp (\<lambda>x. x < n_ps) ===> (=)) (=) (=)"
+    "(list_all2 R ===> list_all2 R ===> (=)) (=) (=)"
+    "(list_all2 (list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R R))
+      ===> list_all2 (list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R R)) ===> (=)) (=) (=)"
+    apply (simp_all add: eq_onp_def)
+    subgoal
+      by (smt assms list.rel_eq list_all2_mono rel_funI)
+    subgoal
+      by (smt assms fun.rel_eq list_all2_eq list_all2_mono rel_fun_mono rel_prod.cases)
+    subgoal
+      by (smt assms fun.rel_eq list_all2_eq list_all2_mono rel_fun_mono rel_prod.cases)
+    done
+  show ?thesis
+    supply [transfer_rule] = upt_0_transfer transfer_consts
+    unfolding make_combs_def by transfer_prover
+qed
+
+lemma broad_trans_from_alt_def2:
+  "broad_trans_from = (\<lambda>(L, s).
+    let
+      pairs = get_committed L;
+      In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+      Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps];
+      In = map (map (filter (\<lambda>(b, _). bval (the \<circ> s) b))) In;
+      Out = map (map (filter (\<lambda>(b, _). bval (the \<circ> s) b))) Out
+    in
+    if pairs = [] then
+      concat (
+        map (\<lambda>a.
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+                compute_upds init combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    else
+      concat (
+        map (\<lambda>a.
+          let
+            ins_committed =
+              List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs;
+            always_committed = (length ins_committed > 1)
+          in
+          concat (map (\<lambda>p.
+            let
+              outs = Out ! p ! a
+            in if outs = [] then []
+            else if
+              \<not> always_committed \<and> (ins_committed = [p] \<or> ins_committed = [])
+              \<and> \<not> list_ex (\<lambda> (q, _). q = p) pairs
+            then []
+            else
+              let
+                combs = make_combs p a In;
+                outs = map (\<lambda>t. (p, t)) outs;
+                combs = (
+                  if combs = [] then [[x]. x \<leftarrow> outs]
+                  else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                init = ([], Broad a, [], (L, s))
+              in
+                compute_upds init combs
+          )
+          [0..<n_ps])
+        )
+      [0..<num_actions])
+    )
+    "
+  unfolding broad_trans_from_def compute_upds_def
+  apply (rule HOL.refl)
+  done
+
+lemma concat_length_transfer:
+  "((\<lambda> x y. list_all2 (list_all2 A) x y \<and> length x = n) ===> list_all2 A) concat concat" for A n
+  by (intro rel_funI concat_transfer[THEN rel_funD], elim conjunct1)
+
+lemma broad_trans_from_transfer:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+  ===> list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  broad_trans_from broad_trans_from_impl"
+proof -
+
+  have compute_upds_impl_transfer[transfer_rule]: "
+    (list_all2 (=) \<times>R
+      rel_label (eq_onp (\<lambda>x. x < num_actions)) \<times>R
+      list_all2 (=) \<times>R
+      (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel ===>
+      list_all2
+       (list_all2
+         (eq_onp (\<lambda>x. x < n_ps) \<times>R
+          eq_onp valid_check \<times>R
+          list_all2 (rel_acconstraint (=) (=)) \<times>R
+          eq_onp (\<lambda>x. x < num_actions) \<times>R
+          list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))) ===>
+      list_all2
+       (list_all2 (=) \<times>R
+        (=) \<times>R
+        list_all2 (=) \<times>R (\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel))
+     compute_upds compute_upds_impl"
+    apply (intro rel_funI)
+    apply (rule compute_upds_transfer[THEN rel_funD, THEN rel_funD])
+     apply (elim rel_prod.cases)
+     apply (simp only:)
+     apply (intro rel_prod.intros)
+         apply assumption
+    subgoal
+      by (drule label.rel_mono_strong[of _ _ _ "(=)"]; simp add: eq_onp_def label.rel_eq)
+       apply (simp; fail)+
+    apply (elim list_all2_mono rel_prod.cases)
+    apply (simp only:)
+    apply (intro rel_prod.intros)
+          apply (assumption | simp add: eq_onp_def acconstraint.rel_eq)+
+    done
+
+  have [is_at_least_equality]: "is_equality (rel_acconstraint (=) (=))"
+    by (tactic \<open>Transfer.eq_tac @{context} 1\<close>)
+    (* by (simp add: acconstraint.rel_eq list.rel_eq is_equality_def) *)
+
+  have eq_transfer1:
+    "(list_all2
+      (eq_onp valid_check \<times>R  list_all2 (rel_acconstraint (=) (=)) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R
+         list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=)) ===>
+    list_all2
+      (eq_onp valid_check \<times>R list_all2 (rel_acconstraint (=) (=)) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R
+         list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))
+    ===> (=)) (=) (=)
+    "
+    by (intro is_at_least_equality_cong2 is_at_least_equality)
+
+  let ?R = "(eq_onp (\<lambda>x. x < n_ps) \<times>R eq_onp valid_check \<times>R
+         list_all2 (rel_acconstraint (=) (=)) \<times>R
+         eq_onp (\<lambda>x. x < num_actions) \<times>R list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))"
+
+  have eq_transfer5:
+    "(list_all2 (list_all2 ?R) ===> list_all2 (list_all2 ?R) ===> (=)) (=) (=)"
+    by (intro is_at_least_equality_cong2 is_at_least_equality)
+
+  have eq_transfer2:
+    "(list_all2 (eq_onp (\<lambda>x. x < n_ps)) ===> list_all2 (eq_onp (\<lambda>x. x < n_ps)) ===> (=)) (=) (=)"
+    by (intro is_at_least_equality_cong2 is_at_least_equality)
+
+  have eq_transfer3:
+    "(eq_onp (\<lambda>x. x < n_ps) ===> eq_onp (\<lambda>x. x < n_ps) ===> (=)) (=) (=)"
+    by (intro is_at_least_equality_cong2 is_at_least_equality)
+
+  have eq_transfer4:
+    "(list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R (=)) ===>
+      list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R (=)) ===> (=)) (=) (=)"
+    by (intro is_at_least_equality_cong2 is_at_least_equality)
+
+  have make_combs_transfer: "
+    (eq_onp (\<lambda>x. x < n_ps) ===>
+        eq_onp (\<lambda>x. x < num_actions) ===>
+        (\<lambda>x y. list_all2 (\<lambda>x y. list_all2 (list_all2
+          (eq_onp valid_check \<times>R list_all2 (rel_acconstraint (=) (=)) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R
+           list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=)))
+          x y \<and> length x = num_actions) x y \<and> length x = n_ps) ===>
+        list_all2 (list_all2 (eq_onp (\<lambda>x. x < n_ps) \<times>R
+          eq_onp valid_check \<times>R list_all2 (rel_acconstraint (=) (=)) \<times>R eq_onp (\<lambda>x. x < num_actions) \<times>R
+          list_all2 (eq_onp valid_upd) \<times>R list_all2 (=) \<times>R (=))
+        )
+    ) make_combs make_combs"
+    apply (intro rel_funI)
+    apply (rule make_combs_transfer[THEN rel_funD, THEN rel_funD, THEN rel_funD])
+    subgoal
+      apply (simp add: acconstraint.rel_eq list.rel_eq prod.rel_eq)
+      apply (drule prod.rel_mono_strong[of _ _ _ _ "(=)" "(=)"], erule eq_onp_to_eq)
+       apply (drule prod.rel_mono_strong[of _ _ _ _ "(=)" "(=)"])
+         apply assumption
+        apply (drule prod.rel_mono_strong[of _ _ _ _ "(=)" "(=)"], erule eq_onp_to_eq)
+         apply (drule prod.rel_mono_strong[of _ _ _ _ "(=)" "(=)"])
+           apply (drule list_all2_mono[of _ _ _ "(=)"], erule eq_onp_to_eq, simp only: list.rel_eq)
+          apply assumption
+         apply (drule (2) prod.rel_mono_strong[of _ _ _ _ "(=)" "(=)"];
+                simp add: acconstraint.rel_eq list.rel_eq prod.rel_eq)+
+      done
+      apply (simp add: prod.rel_eq)+
+    done
+  have "
+  ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+    ===> list_all2 (
+          (=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  (
+  \<lambda>(L, s).
+      let
+        pairs = [];
+        In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+        Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps]
+      in
+        concat (
+          map (\<lambda>a.
+            concat (map (\<lambda>p.
+              let
+                outs = Out ! p ! a
+              in if outs = [] then []
+              else
+                let
+                  combs = make_combs p a In;
+                  outs = map (\<lambda>t. (p, t)) outs;
+                  combs = (
+                    if combs = [] then [[x]. x \<leftarrow> outs]
+                    else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                  init = ([], Broad a, [], (L, s))
+                in
+                  compute_upds init combs
+            )
+            [0..<n_ps])
+          )
+        [0..<num_actions])
+  ) (
+  \<lambda>(L, s).
+      let
+        pairs = [];
+        In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+        Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps]
+      in
+        concat (
+          map (\<lambda>a.
+            concat (map (\<lambda>p.
+              let
+                outs = Out ! p ! a
+              in if outs = [] then []
+              else
+                let
+                  combs = make_combs p a In;
+                  outs = map (\<lambda>t. (p, t)) outs;
+                  combs = (
+                    if combs = [] then [[x]. x \<leftarrow> outs]
+                    else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                  init = ([], Broad a, [], (L, s))
+                in
+                  compute_upds_impl init combs
+            )
+            [0..<n_ps])
+          )
+        [0..<num_actions]))
+  "
+    supply [transfer_rule] =
+      transfer_consts
+      upt_0_transfer
+      map_transfer_length
+      upt_length_transfer
+      make_combs_transfer
+      compute_upds_transfer
+      concat_length_transfer
+      eq_transfer1
+      eq_transfer2
+      eq_transfer3
+      eq_transfer5
+    unfolding Let_def by transfer_prover
+  have "
+  ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+    ===> list_all2 (
+          (=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  (
+  \<lambda>(L, s).
+      let
+        pairs = get_committed L;
+        In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+        Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps]
+      in
+    concat (
+          map (\<lambda>a.
+            let
+              ins_committed =
+                List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs;
+              always_committed = (length ins_committed > 1)
+            in
+            concat (map (\<lambda>p.
+              let
+                outs = Out ! p ! a
+              in if outs = [] then []
+              else if
+                \<not> always_committed \<and> (ins_committed = [p] \<or> ins_committed = [])
+                \<and> \<not> list_ex (\<lambda> (q, _). q = p) pairs
+              then []
+              else
+                let
+                  combs = make_combs p a In;
+                  outs = map (\<lambda>t. (p, t)) outs;
+                  combs = (
+                    if combs = [] then [[x]. x \<leftarrow> outs]
+                    else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                  init = ([], Broad a, [], (L, s))
+                in
+                  compute_upds init combs
+            )
+            [0..<n_ps])
+          )
+        [0..<num_actions])
+  )
+  (
+  \<lambda>(L, s).
+      let
+        pairs = get_committed L;
+        In  = map (\<lambda>p. trans_in_broad_grouped p (L ! p)) [0..<n_ps];
+        Out = map (\<lambda>p. trans_out_broad_grouped p (L ! p)) [0..<n_ps]
+      in
+  concat (
+          map (\<lambda>a.
+            let
+              ins_committed =
+                List.map_filter (\<lambda>(p, _). if In ! p ! a \<noteq> [] then Some p else None) pairs;
+              always_committed = (length ins_committed > 1)
+            in
+            concat (map (\<lambda>p.
+              let
+                outs = Out ! p ! a
+              in if outs = [] then []
+              else if
+                \<not> always_committed \<and> (ins_committed = [p] \<or> ins_committed = [])
+                \<and> \<not> list_ex (\<lambda> (q, _). q = p) pairs
+              then []
+              else
+                let
+                  combs = make_combs p a In;
+                  outs = map (\<lambda>t. (p, t)) outs;
+                  combs = (
+                    if combs = [] then [[x]. x \<leftarrow> outs]
+                    else concat (map (\<lambda>x. map (\<lambda>xs. x # xs) combs) outs));
+                  init = ([], Broad a, [], (L, s))
+                in
+                  compute_upds_impl init combs
+            )
+            [0..<n_ps])
+          )
+        [0..<num_actions])
+  )"
+    supply [transfer_rule] =
+      transfer_consts
+      upt_0_transfer
+      map_transfer_length
+      upt_length_transfer
+      make_combs_transfer
+      compute_upds_transfer
+      concat_length_transfer
+      eq_transfer1
+      eq_transfer2
+      eq_transfer3
+      eq_transfer5
+    unfolding Let_def by transfer_prover
+
+  show ?thesis
+    supply [transfer_rule] =
+      transfer_consts
+      upt_0_transfer
+      map_transfer_length
+      upt_length_transfer
+      make_combs_transfer
+      compute_upds_transfer
+      concat_length_transfer
+      eq_transfer1
+      eq_transfer2
+      eq_transfer3
+      eq_transfer4
+      eq_transfer5
+    unfolding broad_trans_from_alt_def2 broad_trans_from_impl_def Let_def by transfer_prover
+qed
+
+lemma trans_from_transfer:
+  "((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel
+  ===> list_all2 ((=) \<times>R (=) \<times>R (=) \<times>R ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel)))
+  trans_from trans_impl"
+  supply [transfer_rule] = int_trans_from_transfer bin_trans_from_transfer broad_trans_from_transfer
+  unfolding trans_from_def trans_impl_def by transfer_prover
+
+lemma list_all2_swap:
+  "list_all2 S ys xs" if "list_all2 R xs ys" "S = (\<lambda>x y. R y x)" for R S
+  using list_all2_swap that by blast
+
+lemma swap_rel_prod: "(\<lambda> x y. (R \<times>R S) y x) = (\<lambda>x y. R y x) \<times>R (\<lambda>x y. S y x)" for R S
+  by (intro ext) auto
+
+lemma swap_eq:
+  "(\<lambda>x y. y = x) = (=)"
+  by auto
+
+lemma trans_from_refine:
+  "(trans_impl, trans_from) \<in> fun_rel_syn loc_rel (list_rel (Id \<times>r Id \<times>r Id \<times>r loc_rel))"
+proof -
+  have [rel2p]:
+    "rel2p loc_rel = (\<lambda> x y. ((\<lambda>x y. list_all2 (=) x y \<and> length x = n_ps) \<times>R state_rel) y x)"
+    unfolding loc_rel_def rel2p_def by (intro ext) (auto simp: list.rel_eq)
+  have "rel2p (fun_rel_syn
+      {((L', s'), L, s) |L s L' s'. L' = L \<and> length L = n_ps \<and> state_rel s s'}
+      (\<langle>Id \<times>r Id \<times>r Id \<times>r loc_rel\<rangle>list_rel))
+    trans_impl trans_from"
+    unfolding rel2p rel2p_def state_rel_def
+    by (intro rel_funI trans_from_transfer[THEN rel_funD, THEN list_all2_swap])
+       (auto simp: list.rel_eq state_rel_def swap_rel_prod swap_eq)
+  then show ?thesis
+    unfolding rel2p_def relAPP_def loc_rel_def .
+qed
+
+lemma trans_from_correct:
+  "(trans_from, trans_prod) \<in> transition_rel states'"
+  using int_trans_from_correct bin_trans_from_correct broad_trans_from_correct
+  unfolding trans_from_def trans_prod_def transition_rel_def by auto
+
+lemma states'_alt_def:
+  "states' = {(L, s). L \<in> states \<and> bounded bounds s}"
+  unfolding states'_def bounded_def dom_bounds_eq by auto
+
+lemma states'_alt_def2:
+  "states' = {(L, s). L \<in> states \<and> dom s = {0..<n_vs} \<and> check_bounded s}"
+proof -
+  have "states' = {(L, s). L \<in> states \<and> dom s = {0..<n_vs} \<and> bounded bounds s}"
+    unfolding states'_def bounded_def dom_bounds_eq by auto
+  then show ?thesis
+    by (auto simp: check_bounded_iff)
+qed
+
+lemma trans_prod_states'_inv:
+  "l' \<in> states'" if "(l, g, a, r, l') \<in> trans_prod" "l \<in> states'"
+  using that unfolding states'_alt_def
+  by (cases l') (auto dest: trans_prod_bounded_inv trans_prod_states_inv)
+
+lemma prod_ta_states'_inv:
+  "l' \<in> states'" if "prod_ta \<turnstile> l \<longrightarrow>g,a,r l'" "l \<in> states'"
+  using that by simp (rule trans_prod_states'_inv)
+
+lemma dom_eq_transfer [transfer_rule]:
+  "(state_rel ===> (=)) (\<lambda>s. dom s = {0..<n_vs}) (\<lambda>s. length s = n_vs)"
+  by (rule rel_funI) (auto simp: state_rel_def)
+
+lemma states'_memi_correct:
+  "(states'_memi, (\<lambda>l. l \<in> states')) \<in> loc_rel \<rightarrow> bool_rel"
+proof -
+  define t where "t s \<equiv> dom s = {0..<n_vs}" for s :: "nat \<rightharpoonup> int"
+  define ti where "ti s \<equiv> length s = n_vs" for s :: "int list"
+  let ?R = "\<lambda>x y. (eq_onp (\<lambda>L. length L = n_ps) \<times>R state_rel) y x"
+  note [transfer_rule] = dom_eq_transfer[folded t_def ti_def]
+  have [p2rel]: "p2rel ((\<lambda>x y. (eq_onp (\<lambda>L. length L = n_ps) \<times>R state_rel) y x)) = loc_rel"
+    by (auto simp: eq_onp_def p2rel_def loc_rel_def)
+  have *: "(\<lambda>(L, s). L \<in> states \<and> dom s = {0..<n_vs} \<and> check_bounded s) = (\<lambda>l. l \<in> states')"
+    by (intro ext) (auto simp: states'_alt_def2)
+  have "(((=) \<times>R state_rel) ===> (=))
+      (\<lambda>(L, s). L \<in> states \<and> t s \<and> check_bounded s) (\<lambda>(L, s). L \<in> states \<and> ti s \<and> check_boundedi s)"
+    by transfer_prover
+  then have
+    "(((=) \<times>R state_rel) ===> (=))
+    (\<lambda>(L, s). L \<in> states \<and> dom s = {0..<n_vs} \<and> check_bounded s)
+    states'_memi"
+    unfolding t_def ti_def states'_memi_def .
+  then have [p2rel]: "(?R ===> (=)) states'_memi (\<lambda>l. l \<in> states')"
+    unfolding * by (intro rel_funI) (auto simp: eq_onp_def elim!: rel_funE)
+  then have "(states'_memi, (\<lambda>l. l \<in> states')) \<in> p2rel (?R ===> (=))"
+    unfolding p2rel_def by simp
+  then show ?thesis
+    unfolding p2rel .
+qed
+
+end (* Lfiting Syntax *)
+
+end (* Simple Network Impl nat *)
+
+end (* Theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Model_Checking.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Model_Checking.thy
@@ -0,0 +1,1507 @@
+theory Simple_Network_Language_Model_Checking
+  imports Munta_Base.Temporal_Logics
+    Simple_Network_Language_Impl_Refine
+    Munta_Model_Checker.UPPAAL_Model_Checking
+begin
+
+section \<open>Product Bisimulation\<close>
+
+no_notation State_Networks.step_sn ("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow>_ \<langle>_, _, _\<rangle>" [61,61,61,61,61] 61)
+
+text \<open>
+  We have proved the necessary theorems already but we need to lift it to
+  the case where delay and action transitions are strictly alternating.
+\<close>
+inductive step_u' ::
+  "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval
+  \<Rightarrow> 's list \<Rightarrow> ('x \<rightharpoonup> 'v) \<Rightarrow> ('c, 't) cval \<Rightarrow> bool"
+("_ \<turnstile> \<langle>_, _, _\<rangle> \<rightarrow> \<langle>_, _, _\<rangle>" [61,61,61,61] 61) where
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L'', s'', u''\<rangle>" if
+  "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  "a \<noteq> Del"
+  "A \<turnstile> \<langle>L', s', u'\<rangle> \<rightarrow>a \<langle>L'', s'', u''\<rangle>"
+
+inductive_cases step'_elims: "A \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+
+inductive_cases step_u'_elims: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+
+theorem Bisimulation_Invariant_strong_intro:
+  fixes A :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+    and P :: "'a \<Rightarrow> bool"
+    and B :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes "\<And>a b. A a b \<Longrightarrow> P a \<Longrightarrow> B a b"
+    and "\<And>a b. B a b \<Longrightarrow> P a \<Longrightarrow> A a b"
+    and "\<And>a b. P a \<Longrightarrow> A a b \<Longrightarrow> P b"
+  shows "Bisimulation_Invariant A B (=) P P"
+  apply standard
+  subgoal A for a b a'
+    by (blast intro: assms)
+  subgoal B for a b a'
+    by (blast intro: assms)
+  subgoal C for a b
+    by (rule assms)
+  subgoal D for a b
+    by (rule C, assumption) (rule assms)
+  done
+
+context Prod_TA_Defs
+begin
+
+definition
+  "all_prop L s = (L \<in> states \<and> bounded bounds s)"
+
+lemma all_prop_boundedD[dest]:
+  "bounded bounds s" if "all_prop L s"
+  using that unfolding all_prop_def ..
+
+lemma all_prop_statesD[dest]:
+  "L \<in> states" if "all_prop L s"
+  using that unfolding all_prop_def ..
+
+end (* Prod TA Defs *)
+
+
+context Prod_TA
+begin
+
+lemma prod_action_step_not_eq_delay:
+  "a \<noteq> Del" if "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>"
+  using that
+  apply cases
+  unfolding trans_of_def
+  unfolding prod_ta_def trans_prod_def
+  by (auto simp: trans_int_def trans_bin_def trans_broad_def)
+
+end
+
+locale Prod_TA_urge =
+  Prod_TA + assumes urgency_removed: "\<forall>N \<in> set (fst (snd A)). urgent N = {}"
+begin
+
+lemma prod_all_prop_inv:
+  "all_prop L' s'" if "all_prop L s" "prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+  using that unfolding all_prop_def
+  by (auto elim: bounded_inv states_inv simp: step_iff[symmetric, OF _ _ urgency_removed])
+
+lemma prod_all_prop_inv':
+  "all_prop L' s'" if "all_prop L s" "prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+  using that by (auto intro: prod_all_prop_inv elim!: step'_elims)
+
+interpretation prod_bisim:
+  Bisimulation_Invariant
+  "(\<lambda> (L, s, u) (L', s', u'). prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+  "(\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)"
+  "(=)"
+  "(\<lambda> (L, s, u). all_prop L s)"
+  "(\<lambda> (L, s, u). all_prop L s)"
+  apply (rule Bisimulation_Invariant_strong_intro; clarsimp)
+  subgoal
+    by (auto intro: step_u'.intros simp: all_prop_def
+             dest: action_sound prod_action_step_not_eq_delay delay_sound[OF _ _ _ urgency_removed]
+             elim!: step'_elims)
+  subgoal
+    by (auto 4 4 dest: prod_all_prop_inv action_complete elim: delay_complete elim!: step_u'_elims)
+  subgoal
+    by (rule prod_all_prop_inv')
+  done
+
+lemmas prod_bisim_intro = prod_bisim.Bisimulation_Invariant_axioms
+
+end (* Prod TA *)
+
+
+section \<open>The Final Semantics\<close>
+
+text \<open>State formulas\<close>
+datatype ('n, 's, 'a, 'b) sexp =
+  true |
+  \<comment> \<open>Boolean connectives\<close>
+  not "('n, 's, 'a, 'b) sexp" | "and" "('n, 's, 'a, 'b) sexp" "('n, 's, 'a, 'b) sexp" |
+  or "('n, 's, 'a, 'b) sexp" "('n, 's, 'a, 'b) sexp" | imply "('n, 's, 'a, 'b) sexp" "('n, 's, 'a, 'b) sexp" |
+  \<comment> \<open>Does var \<open>a\<close> equal \<open>x\<close>?\<close>
+  eq 'a 'b |
+  le 'a 'b |
+  lt 'a 'b |
+  ge 'a 'b |
+  gt 'a 'b |
+  \<comment> \<open>Is procces \<open>i\<close> in location \<open>l\<close>?\<close>
+  loc 'n 's
+
+datatype ('n, 's, 'a, 'b) formula =
+  EX "('n, 's, 'a, 'b) sexp" | EG "('n, 's, 'a, 'b) sexp"
+| AX "('n, 's, 'a, 'b) sexp" | AG "('n, 's, 'a, 'b) sexp"
+| Leadsto "('n, 's, 'a, 'b) sexp" "('n, 's, 'a, 'b) sexp"
+
+fun check_sexp :: "(nat, 's, 'a, 'b :: linorder) sexp \<Rightarrow> 's list \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
+  "check_sexp sexp.true _ _ \<longleftrightarrow> True" |
+  "check_sexp (not e) L s \<longleftrightarrow> \<not> check_sexp e L s" |
+  "check_sexp (and e1 e2) L s \<longleftrightarrow> check_sexp e1 L s \<and> check_sexp e2 L s" |
+  "check_sexp (sexp.or e1 e2) L s \<longleftrightarrow> check_sexp e1 L s \<or> check_sexp e2 L s" |
+  "check_sexp (imply e1 e2) L s \<longleftrightarrow> check_sexp e1 L s \<longrightarrow> check_sexp e2 L s" |
+  "check_sexp (eq i x) L s \<longleftrightarrow> s i = x" |
+  "check_sexp (le i x) L s \<longleftrightarrow> s i \<le> x" |
+  "check_sexp (lt i x) L s \<longleftrightarrow> s i < x" |
+  "check_sexp (ge i x) L s \<longleftrightarrow> s i \<ge> x" |
+  "check_sexp (gt i x) L s \<longleftrightarrow> s i > x" |
+  "check_sexp (loc i x) L s \<longleftrightarrow> L ! i = x"
+
+fun locs_of_sexp :: "('n, 's, 'a, 'b) sexp \<Rightarrow> 'n set" where
+  "locs_of_sexp (not e) = locs_of_sexp e" |
+  "locs_of_sexp (and e1 e2) = locs_of_sexp e1 \<union> locs_of_sexp e2" |
+  "locs_of_sexp (sexp.or e1 e2) = locs_of_sexp e1 \<union> locs_of_sexp e2" |
+  "locs_of_sexp (imply e1 e2) = locs_of_sexp e1 \<union> locs_of_sexp e2" |
+  "locs_of_sexp (loc i x) = {i}" |
+  "locs_of_sexp _ = {}"
+
+fun vars_of_sexp :: "('n, 's, 'a, 'b) sexp \<Rightarrow> 'a set" where
+  "vars_of_sexp (not e) = vars_of_sexp e" |
+  "vars_of_sexp (and e1 e2) = vars_of_sexp e1 \<union> vars_of_sexp e2" |
+  "vars_of_sexp (sexp.or e1 e2) = vars_of_sexp e1 \<union> vars_of_sexp e2" |
+  "vars_of_sexp (imply e1 e2) = vars_of_sexp e1 \<union> vars_of_sexp e2" |
+  "vars_of_sexp (eq i x) = {i}" |
+  "vars_of_sexp (lt i x) = {i}" |
+  "vars_of_sexp (le i x) = {i}" |
+  "vars_of_sexp (ge i x) = {i}" |
+  "vars_of_sexp (gt i x) = {i}" |
+  "vars_of_sexp _ = {}"
+
+fun locs_of_formula :: "('n, 's, 'a, 'b) formula \<Rightarrow> 'n set" where
+  "locs_of_formula (formula.EX \<phi>) = locs_of_sexp \<phi>" |
+  "locs_of_formula (EG \<phi>) = locs_of_sexp \<phi>" |
+  "locs_of_formula (AX \<phi>) = locs_of_sexp \<phi>" |
+  "locs_of_formula (AG \<phi>) = locs_of_sexp \<phi>" |
+  "locs_of_formula (Leadsto \<phi> \<psi>) = locs_of_sexp \<phi> \<union> locs_of_sexp \<psi>"
+
+fun vars_of_formula :: "('n, 's, 'a, 'b) formula \<Rightarrow> 'a set" where
+  "vars_of_formula (formula.EX \<phi>) = vars_of_sexp \<phi>" |
+  "vars_of_formula (EG \<phi>) = vars_of_sexp \<phi>" |
+  "vars_of_formula (AX \<phi>) = vars_of_sexp \<phi>" |
+  "vars_of_formula (AG \<phi>) = vars_of_sexp \<phi>" |
+  "vars_of_formula (Leadsto \<phi> \<psi>) = vars_of_sexp \<phi> \<union> vars_of_sexp \<psi>"
+
+fun hd_of_formula :: "(nat, 's, 'a, 'b) formula \<Rightarrow> 's list \<Rightarrow> ('a \<Rightarrow> 'b :: linorder) \<Rightarrow> bool" where
+  "hd_of_formula (formula.EX \<phi>) L s = check_sexp \<phi> L s" |
+  "hd_of_formula (EG \<phi>) L s = check_sexp \<phi> L s" |
+  "hd_of_formula (AX \<phi>) L s = Not (check_sexp \<phi> L s)" |
+  "hd_of_formula (AG \<phi>) L s = Not (check_sexp \<phi> L s)" |
+  "hd_of_formula (Leadsto \<phi> _) L s = check_sexp \<phi> L s"
+
+definition models ("_,_ \<Turnstile> _" [61,61] 61) where
+  "A,a0 \<Turnstile> \<Phi> \<equiv> (case \<Phi> of
+    formula.EX \<phi> \<Rightarrow>
+      Graph_Defs.Ex_ev
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_sexp \<phi> L (the o s))
+  | formula.EG \<phi> \<Rightarrow>
+      Graph_Defs.Ex_alw
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_sexp \<phi> L (the o s))
+  | formula.AX \<phi> \<Rightarrow>
+      Graph_Defs.Alw_ev
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_sexp \<phi> L (the o s))
+  | formula.AG \<phi> \<Rightarrow>
+      Graph_Defs.Alw_alw
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_sexp \<phi> L (the o s))
+  | formula.Leadsto \<phi> \<psi> \<Rightarrow>
+      Graph_Defs.leadsto
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_sexp \<phi> L (the o s))
+        (\<lambda> (L, s, _). check_sexp \<psi> L (the o s))
+  ) a0"
+
+lemmas models_iff = models_def[unfolded Graph_Defs.Ex_alw_iff Graph_Defs.Alw_alw_iff]
+
+definition prop_of where
+  "prop_of \<phi> = PropC (\<lambda>(L, s, _). check_sexp \<phi> L (the o s))"
+
+fun ctl_of where
+  "ctl_of (formula.EX \<phi>) = ctl_formula.EX (prop_of \<phi>)"
+| "ctl_of (formula.EG \<phi>) = ctl_formula.EG (prop_of \<phi>)"
+| "ctl_of (formula.AX \<phi>) = ctl_formula.AX (prop_of \<phi>)"
+| "ctl_of (formula.AG \<phi>) = ctl_formula.AG (prop_of \<phi>)"
+| "ctl_of (formula.Leadsto \<phi> \<psi>) =
+  ctl_formula.AG (ctl_formula.ImpliesC (prop_of \<phi>) (ctl_formula.AX (prop_of \<psi>)))"
+
+context
+  fixes A :: "('a, 's, 'c, 't :: time, 'x, 'v :: linorder) nta"
+begin
+
+interpretation Graph_Defs "\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>" .
+
+lemma models_ctl_iff:
+  "A,a0 \<Turnstile> \<Phi> \<longleftrightarrow> models_ctl (ctl_of \<Phi>) a0"
+  by (cases \<Phi>) (simp_all add: models_def prop_of_def leadsto_def)
+
+end
+
+fun check_sexpi :: "(nat, 's, nat, int) sexp \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> bool" where
+  "check_sexpi sexp.true _ _ \<longleftrightarrow> True" |
+  "check_sexpi (not e) L s \<longleftrightarrow> \<not> check_sexpi e L s" |
+  "check_sexpi (and e1 e2) L s \<longleftrightarrow> check_sexpi e1 L s \<and> check_sexpi e2 L s" |
+  "check_sexpi (sexp.or e1 e2) L s \<longleftrightarrow> check_sexpi e1 L s \<or> check_sexpi e2 L s" |
+  "check_sexpi (imply e1 e2) L s \<longleftrightarrow> check_sexpi e1 L s \<longrightarrow> check_sexpi e2 L s" |
+  "check_sexpi (eq i x) L s \<longleftrightarrow> s ! i = x" |
+  "check_sexpi (le i x) L s \<longleftrightarrow> s ! i \<le> x" |
+  "check_sexpi (lt i x) L s \<longleftrightarrow> s ! i < x" |
+  "check_sexpi (ge i x) L s \<longleftrightarrow> s ! i \<ge> x" |
+  "check_sexpi (gt i x) L s \<longleftrightarrow> s ! i > x" |
+  "check_sexpi (loc i x) L s \<longleftrightarrow> L ! i = x"
+
+fun hd_of_formulai :: "(nat, 's, nat, int) formula \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> bool" where
+  "hd_of_formulai (formula.EX \<phi>) L s = check_sexpi \<phi> L s" |
+  "hd_of_formulai (EG \<phi>) L s = check_sexpi \<phi> L s" |
+  "hd_of_formulai (AX \<phi>) L s = Not (check_sexpi \<phi> L s)" |
+  "hd_of_formulai (AG \<phi>) L s = Not (check_sexpi \<phi> L s)" |
+  "hd_of_formulai (Leadsto \<phi> _) L s = check_sexpi \<phi> L s"
+
+
+
+(* Stop locale constant unfolding in simplifier *)
+lemma all_prop_weak_cong[cong]:
+  "Prod_TA_Defs.all_prop A = Prod_TA_Defs.all_prop A" for A
+  by(rule HOL.refl)
+lemma equiv'_weak_cong[cong]:
+  "Bisimulation_Invariant.equiv' R P Q = Bisimulation_Invariant.equiv' R P Q" for R P Q
+  by (rule HOL.refl)
+lemma equiv'_weak_cong2[cong]:
+  "Simulation_Invariant.equiv' R P Q = Simulation_Invariant.equiv' R P Q" for R P Q
+  by (rule HOL.refl)
+lemma models_ctl_weak_cong[cong]:
+  "Graph_Defs.models_ctl E = Graph_Defs.models_ctl E" for E
+  by (rule HOL.refl)
+lemma models_path_weak_cong[cong]:
+  "Graph_Defs.models_path E = Graph_Defs.models_path E" for E
+  by (rule HOL.refl)
+lemma models_state_weak_cong[cong]:
+  "Graph_Defs.models_state E = Graph_Defs.models_state E" for E
+  by (rule HOL.refl)
+lemma models_ltl_weak_cong[cong]:
+  "Graph_Defs.models_ltlc E = Graph_Defs.models_ltlc E" for E
+  by (rule HOL.refl)
+lemma models_weak_cong[cong]:
+  "models E = models E"
+  for E by (rule HOL.refl)
+
+
+
+
+section \<open>Instantiating the Model Checking Locale\<close>
+
+locale Simple_Network_Impl_nat_urge =
+  Simple_Network_Impl_nat + assumes no_urgency: "\<forall> (_, U, _, _) \<in> set automata. U = []"
+
+text \<open>
+  This locale certifies that a given local clock ceiling is correct.
+  Moreover, we certify that the vector of initial locations has outgoing transitions for
+  each automaton, and that all variables of the initial state are in bounds.
+\<close>
+locale Simple_Network_Impl_nat_ceiling_start_state =
+  Simple_Network_Impl_nat_urge +
+  fixes k :: "nat list list list"
+    and L0 :: "nat list"
+    and s0 :: "(nat \<times> int) list"
+    and formula :: "(nat, nat, nat, int) formula"
+    and show_clock :: "nat \<Rightarrow> string"
+    and show_state :: "nat list \<times> int list \<Rightarrow> string"
+  assumes k_ceiling:
+    "\<forall>i < n_ps. \<forall>(l, g) \<in> set ((snd o snd o snd) (automata ! i)).
+      \<forall>(x, m) \<in> collect_clock_pairs g. m \<le> int (k ! i ! l ! x)"
+    "\<forall>i < n_ps. \<forall>(l, _, g, _) \<in> set ((fst o snd o snd) (automata ! i)).
+      (\<forall>(x, m) \<in> collect_clock_pairs g. m \<le> int (k ! i ! l ! x))"
+  and k_resets:
+    "\<forall>i < n_ps. \<forall> (l, b, g, a, upd, r, l') \<in> set ((fst o snd o snd) (automata ! i)).
+       \<forall>c \<in> {0..<m+1} - set r. k ! i ! l' ! c \<le> k ! i ! l ! c"
+  and k_length:
+    "length k = n_ps" "\<forall> i < n_ps. length (k ! i) = num_states i"
+    "\<forall> xs \<in> set k. \<forall> xxs \<in> set xs. length xxs = m + 1"
+  and k_0:
+    "\<forall>i < n_ps. \<forall>l < num_states i. k ! i ! l ! 0 = 0"
+  and inv_unambiguous:
+    "\<forall>(_, _, _, inv) \<in> set automata. distinct (map fst inv)"
+  and s0_bounded: "bounded bounds (map_of s0)"
+  and L0_len: "length L0 = n_ps"
+  and L0_has_trans: "\<forall>i < n_ps. L0 ! i \<in> fst ` set ((fst o snd o snd) (automata ! i))"
+  and vars_of_formula: "vars_of_formula formula \<subseteq> {0..<n_vs}"
+  (* and num_states_length: "\<forall>i<n_ps. num_states i = length (fst (snd (automata ! i)))" *)
+begin
+
+text \<open>
+The ceiling \<open>k\<close> is correct for each individual automaton in the network.
+We now construct a ceiling for the product automaton:
+\<close>
+definition
+  "k_fun l c \<equiv>
+    if (c > 0 \<and> c \<le> m) \<and> fst l \<in> states then Max {k ! i ! (fst l ! i) ! c | i . i < n_ps} else 0"
+
+lemma (in -) default_map_of_distinct:
+  "(k, default_map_of x xs k) \<in> set xs \<union> {(k, x)}" if "distinct (map fst xs)"
+  unfolding default_map_of_alt_def by clarsimp (simp add: map_of_eq_Some_iff[OF that])
+
+lemma N_inv:
+  "(L ! i, inv (N i) (L ! i)) \<in> set ((snd o snd o snd) (automata ! i)) \<union> {(L ! i, [])}"
+  if "i < n_ps"
+  unfolding N_def comp_def fst_conv snd_conv inv_def
+  using that
+  apply (subst nth_map)
+   apply (simp add: n_ps_def; fail)
+  apply (clarsimp split: prod.split simp: automaton_of_def)
+  subgoal for _ _ _ xs
+    using default_map_of_distinct[of xs "L ! i" "[]"] inv_unambiguous that
+    by (auto dest!: nth_mem simp: n_ps_def)
+  done
+
+lemma (in -) subset_nat_0_atLeastLessThan_conv:
+  "S \<subseteq> {0..<n::nat} \<longleftrightarrow> (\<forall> x \<in> S. x < n)"
+  by auto
+
+lemma k_ceiling_rule:
+  "m \<le> int (k ! i ! l ! x)"
+  if "i < n_ps" "(l, b, g, xx) \<in> set ((fst o snd o snd) (automata ! i))"
+     "(x, m) \<in> collect_clock_pairs g"
+  for i l x g xx
+  using that k_ceiling(2) by fastforce
+
+lemma k_ceiling_1:
+  "\<forall>s. \<forall>L \<in> states. \<forall>(x,m) \<in> clkp_set prod_ta (L, s). m \<le> k_fun (L, s) x"
+proof safe
+  fix L s c x
+  assume \<open>L \<in> states\<close> \<open>(c, x) \<in> Closure.clkp_set prod_ta (L, s)\<close>
+  have "0 < c" "c \<le> m"
+  proof -
+    from \<open>(c, x) \<in> _\<close> have "(c, x) \<in> Timed_Automata.clkp_set prod_ta"
+      unfolding TA_clkp_set_unfold by auto
+    with clock_range show "0 < c" "c \<le> m"
+      by auto
+  qed
+  with \<open>L \<in> _\<close> have "k_fun (L, s) c = Max {k ! i ! (L ! i) ! c | i . i < n_ps}"
+    unfolding k_fun_def by auto
+  have Max_aux: "x \<le> int (Max {k ! i ! (L ! i) ! c |i. i < n_ps})"
+    if "x \<le> int (k ! p ! (L ! p) ! c)" "p < n_ps" for p
+  proof -
+    from \<open>p < n_ps \<close> have "k ! p ! (L ! p) ! c \<le> Max {k ! i ! (L ! i) ! c |i. i < n_ps}"
+      by (intro Max_ge) auto
+    with \<open>x \<le> _\<close> show ?thesis
+      by simp
+  qed
+  from \<open>(c, x) \<in> _\<close> show \<open>x \<le> k_fun (L, s) c\<close>
+    unfolding clkp_set_def
+  proof safe
+    assume \<open>(c, x) \<in> Closure.collect_clki (inv_of prod_ta) (L, s)\<close>
+    then show \<open>x \<le> k_fun (L, s) c\<close>
+      using k_ceiling(1) unfolding collect_clki_def \<open>k_fun (L, s) c = _\<close>
+      by (auto 0 8 dest: N_inv
+          intro!: Max_aux simp: prod_inv_def collect_clock_pairs_def k_fun_def)
+  next
+    assume \<open>(c, x) \<in> Closure.collect_clkt (trans_of prod_ta) (L, s)\<close>
+    then show \<open>x \<le> k_fun (L, s) c\<close>
+      unfolding collect_clkt_def \<open>k_fun (L, s) c = _\<close>
+      apply (clarsimp simp: trans_prod_def collect_clock_pairs_def k_fun_def)
+      apply safe
+      subgoal
+        using k_ceiling(2) unfolding trans_int_def
+        apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+        apply (fastforce intro!: Max_aux simp: collect_clock_pairs_def)
+        done
+      subgoal
+        using k_ceiling(2) unfolding trans_bin_def
+        apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+        apply (erule disjE)
+         apply (force intro!: Max_aux simp: collect_clock_pairs_def)+
+        done
+      subgoal
+        using k_ceiling(2) unfolding trans_broad_def
+        apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+        apply (erule disjE)
+         apply (fastforce intro!: Max_aux simp: collect_clock_pairs_def)
+        apply (erule bexE)
+        apply (force intro!: Max_aux simp: collect_clock_pairs_def) \<comment> \<open>slow: 60s\<close>
+        done
+      done
+  qed
+qed
+
+lemma k_fun_mono':
+  "k_fun (L, s) c \<le> k_fun (L', s') c" if
+  "\<forall>i < n_ps. k ! i ! (L ! i) ! c \<le> k ! i ! (L' ! i) ! c" "L \<in> states" "L' \<in> states"
+  using that unfolding k_fun_def
+  apply clarsimp
+  apply (cases "n_ps = 0")
+   apply (simp; fail)
+  apply (rule Max.boundedI)
+    apply (simp; fail)
+   apply blast
+  apply safe
+  subgoal for _ i
+    by - (rule order.trans[where b = "k ! i ! (L' ! i) ! c"], auto intro: Max_ge)
+  done
+
+lemma k_fun_mono:
+  \<open>Max {k ! i ! (L ! i) ! c | i . i < n_ps} \<le> Max {k ! i ! (L' ! i) ! c | i . i < n_ps}\<close>
+  if \<open>\<forall>i < n_ps. k ! i ! (L ! i) ! c \<le> k ! i ! (L' ! i) ! c\<close>
+  apply (cases "n_ps = 0")
+   apply (simp; fail)
+  apply (rule Max.boundedI)
+    apply (simp; fail)
+   apply blast
+  apply safe
+  subgoal for _ i
+    using that by - (rule order.trans[where b = "k ! i ! (L' ! i) ! c"], auto intro: Max_ge)
+  done
+
+lemma (in -) fold_upds_aux1:
+  "fold (\<lambda>p L. L[p := g p]) ps xs ! i = xs ! i" if \<open>i \<notin> set ps\<close>
+  using that by (induction ps arbitrary: xs) auto
+
+lemma (in -) fold_upds_aux2:
+  "fold (\<lambda>p L. L[p := g p]) ps xs ! i = g i" if \<open>distinct ps\<close> \<open>i \<in> set ps\<close> \<open>i < length xs\<close>
+  using that by (induction ps arbitrary: xs) (auto simp: fold_upds_aux1)
+
+lemma (in -) fold_upds_aux_length:
+  "length (fold (\<lambda>p L. L[p := g p]) ps xs) = length xs"
+  by (induction ps arbitrary: xs) auto
+
+lemma prod_ta_step_statesD:
+  assumes "prod_ta \<turnstile> (L, s) \<longrightarrow>g,a,r (L', s')"
+  shows "L \<in> states" "L' \<in> states"
+  using assms state_set_states by (fastforce dest: state_setI1 state_setI2)+
+
+lemma k_ceiling_2:
+  "\<forall>l g a r l'. \<forall> c \<le> m. prod_ta \<turnstile> l \<longrightarrow>g,a,r l' \<and> c \<notin> set r \<longrightarrow> k_fun l' c \<le> k_fun l c"
+proof safe
+  fix L s g a r L' s' c
+  assume A: \<open>c \<le> m\<close> \<open>prod_ta \<turnstile> (L, s) \<longrightarrow>g,a,r (L', s')\<close> \<open>c \<notin> set r\<close>
+  then have "L \<in> states" "L' \<in> states"
+    by - (rule prod_ta_step_statesD, assumption)+
+  from A have \<open>Max {k ! i ! (L' ! i) ! c | i . i < n_ps} \<le> Max {k ! i ! (L ! i) ! c | i . i < n_ps}\<close>
+    apply simp
+    unfolding trans_prod_def
+    apply safe
+    subgoal
+      using k_resets
+      unfolding trans_int_def
+      apply clarsimp
+      apply (rule k_fun_mono)
+      apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+      subgoal for b f p aa l' i
+        by (cases "p = i"; force simp add: L_len)
+      done
+    subgoal
+      using k_resets
+      unfolding trans_bin_def
+      apply clarsimp
+      apply (rule k_fun_mono)
+      apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+      subgoal for _ _ p q b1 g1 f1 r1 l1' b2 g2 f2 r2 l2' i
+        by (cases "p = i"; cases "q = i"; force simp add: L_len)
+      done
+    subgoal
+      using k_resets
+      unfolding trans_broad_def
+      apply clarsimp
+      apply (rule k_fun_mono)
+      apply (clarsimp simp: mem_trans_N_iff L_len subset_nat_0_atLeastLessThan_conv)
+      subgoal premises prems for s'a aa p b ga f ra l' bs gs fs rs ls' ps i
+      proof (cases "p = i")
+        case True
+        with \<open>p \<notin> _\<close> \<open>i < _\<close> \<open>L \<in> states\<close> have "(fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l'] ! i = l'"
+          by (simp add: L_len fold_upds_aux_length)
+        with prems \<open>p = i\<close> show ?thesis
+          by (fastforce simp add: L_len)
+      next
+        case False
+        then have *: "(fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l'] ! i
+          = fold (\<lambda>p L. L[p := ls' p]) ps L ! i"
+          by simp
+        show ?thesis
+        proof (cases "i \<in> set ps")
+          case True
+          then have **: "fold (\<lambda>p L. L[p := ls' p]) ps L ! i = ls' i"
+            using \<open>distinct ps\<close> \<open>i < n_ps\<close> \<open>L \<in> states\<close> by (auto simp: fold_upds_aux2)
+          moreover have
+            "(L ! i, bs i, gs i, In aa, fs i, rs i, ls' i) \<in> set (fst (snd (snd (automata ! i))))"
+            using \<open>p \<noteq> i\<close> True prems by fast
+          moreover have "c\<in>{0..<Suc m} - set (rs i)"
+            using \<open>p \<noteq> i\<close> True prems by force
+          ultimately show ?thesis
+            using prems(2) \<open>i < n_ps\<close> by (auto 4 3 simp add: *)
+        next
+          case False
+          with \<open>p \<noteq> i\<close> show ?thesis
+            by (simp add: fold_upds_aux1)
+        qed
+      qed
+      done
+    done
+  with \<open>L \<in> states\<close> \<open>L' \<in> states\<close> \<open>c \<le> m\<close> show "k_fun (L', s') c \<le> k_fun (L, s) c"
+    by (auto simp: k_fun_def)
+qed
+
+
+
+abbreviation F where "F \<equiv> \<lambda>(L, s). hd_of_formula formula L (the o s)"
+abbreviation "Fi \<equiv> \<lambda>(L, s). hd_of_formulai formula L s"
+
+lemma (in Simple_Network_Impl_nat) check_sexp_check_sexpi:
+  "check_sexp e L (the o s) \<longleftrightarrow> check_sexpi e L s'"
+  if "state_rel s s'" "vars_of_sexp e \<subseteq> {0..<n_vs}"
+  using that unfolding state_rel_def by (induction e) auto
+
+lemma (in Simple_Network_Impl_nat) hd_of_formula_hd_of_formulai:
+  "hd_of_formula \<phi> L (the o s) \<longleftrightarrow> hd_of_formulai \<phi> L s'"
+  if "state_rel s s'" "vars_of_formula \<phi> \<subseteq> {0..<n_vs}"
+  using that by (induction \<phi>) (auto simp: check_sexp_check_sexpi)
+
+lemma F_Fi:
+  "F l \<longleftrightarrow> Fi l'" if "(l', l) \<in> loc_rel"
+  using vars_of_formula that unfolding loc_rel_def by clarsimp (erule hd_of_formula_hd_of_formulai)
+
+
+abbreviation "l0 \<equiv> (L0, map_of s0)"
+abbreviation "s0i \<equiv> map (the o map_of s0) [0..<n_vs]"
+abbreviation "l0i \<equiv> (L0, s0i)"
+
+lemma state_rel_start:
+  "state_rel (map_of s0) s0i"
+  using s0_bounded unfolding state_rel_def bounded_def dom_bounds_eq by auto
+
+lemma statesI:
+  "L \<in> states" if "length L = n_ps" "\<forall>i<n_ps. L ! i \<in> fst ` set (fst (snd (snd (automata ! i))))"
+  using that unfolding states_def by (auto 4 3 simp: mem_trans_N_iff[symmetric])
+
+lemma L0_states[simp, intro]:
+  "L0 \<in> states"
+  using L0_has_trans L0_len by (auto intro: statesI)
+
+lemma l0_states'[simp, intro]:
+  "l0 \<in> states'"
+  using state_rel_start s0_bounded unfolding states'_def state_rel_def by auto
+
+sublocale reach: Reachability_Problem_Defs
+  prod_ta
+  l0
+  m
+  k_fun
+  F
+  .
+
+lemma (in -) collect_clkt_state_setI:
+  assumes "(x, d) \<in> Closure.collect_clkt (trans_of A) l"
+  shows "l \<in> state_set (trans_of A)" "l \<in> Simulation_Graphs_TA.state_set A"
+  using assms unfolding collect_clkt_def by (auto simp: state_set_def)
+
+lemma clkp_set_statesD:
+  fixes x d
+  assumes "(x, d)\<in>Closure.clkp_set prod_ta (L, s)"
+  shows "L \<in> states"
+  using assms
+  unfolding clkp_set_def collect_clki_def
+  apply safe
+  subgoal
+    apply simp
+    unfolding prod_inv_def
+    unfolding collect_clock_pairs_def
+    apply auto
+    done
+  apply (drule collect_clkt_state_setI)
+  using state_set_states by auto
+
+sublocale reach1: Reachability_Problem
+  prod_ta
+  l0
+  m
+  k_fun
+  F
+  apply standard
+  subgoal
+    apply safe
+    using k_ceiling_1
+    subgoal for L s x m
+      apply (cases "L \<in> states")
+       apply blast
+      apply (auto dest: clkp_set_statesD simp: k_fun_def)
+      done
+    done
+  subgoal
+    apply safe
+    subgoal for L s g a r L' s' c
+      apply (cases "c \<le> m")
+      using k_ceiling_2
+       apply force
+      apply (auto simp: k_fun_def)
+      done
+    done
+  subgoal
+    by (simp add: k_fun_def)
+  subgoal
+    by (simp add: k_fun_def)
+  done  \<comment> \<open>slow: 60s\<close>
+
+lemma states'_superset:
+  "{l0} \<union> Normalized_Zone_Semantics_Impl_Refine.state_set trans_prod \<subseteq> states'"
+  (is "{l0} \<union> ?S \<subseteq> states'")
+proof -
+  have "?S \<subseteq> states'"
+  proof safe
+    fix L s
+    assume "(L, s) \<in> ?S"
+    then have "L \<in> states"
+      using state_set_states[unfolded trans_of_prod state_set_eq] by blast
+    moreover have "bounded bounds s"
+      using \<open>(L, s) \<in> _\<close>
+      unfolding state_set_def
+      unfolding trans_prod_def
+      unfolding trans_int_def trans_bin_def trans_broad_def
+      by auto
+    ultimately show "(L, s) \<in> states'"
+      by (auto simp: states'_alt_def)
+  qed
+  then show ?thesis
+    by simp
+qed
+
+
+
+definition "k_i \<equiv> IArray (map (IArray o (map (IArray o map int))) k)"
+
+definition
+  "k_impl \<equiv> \<lambda>(l, _). IArray (map (\<lambda> c. Max {k_i !! i !! (l ! i) !! c | i. i < n_ps}) [0..<m+1])"
+
+(* Duplication with UPPAAL_State_Networks_Impl_Refine *)
+lemma Max_int_commute:
+  "int (Max S) = Max (int ` S)" if "finite S" "S \<noteq> {}"
+  apply (rule mono_Max_commute)
+    apply rule
+  using that by auto
+
+lemma (in Simple_Network_Impl_nat) n_ps_gt_0: "n_ps > 0"
+  using length_automata_eq_n_ps non_empty by auto
+
+lemma statesD:
+  "L ! i \<in> fst ` set (fst (snd (snd (automata ! i))))
+ \<or> L ! i \<in> (snd o snd o snd o snd o snd o snd) ` set (fst (snd (snd (automata ! i))))"
+  if "L \<in> states" "length L = n_ps" "i < n_ps"
+  using that unfolding states_def
+  apply safe
+  apply -
+  apply (elim allE impE, assumption)
+  apply safe
+   apply (force simp: mem_trans_N_iff)+
+  done
+
+lemma k_impl_k_fun:
+  "k_impl (L, s) = IArray (map (k_fun (L, s)) [0..<m+1])" if "L \<in> states"
+proof -
+  define k_i2 where "k_i2 i c = k_i !! i !! (L ! i) !! c" for i c
+  have k_i2_k: "k_i2 i c = k ! i ! (L ! i) ! c" if "i < n_ps" "c \<le> m" for i c
+  proof -
+    have "i < length k"
+      by (simp add: k_length(1) that(1))
+    moreover have "L ! i < length (k ! i)"
+      using L_i_len[OF _ \<open>L \<in> states\<close>] k_length(2) \<open>i < n_ps\<close> by auto
+    moreover have "c < length (k ! i ! (L ! i))"
+      using k_length(3) \<open>c \<le> m\<close> \<open>i < length k\<close> \<open>L ! i < length (k ! i)\<close> by (auto dest: nth_mem)
+    ultimately show ?thesis
+      unfolding k_i2_def k_i_def by simp
+  qed
+  have k_impl_alt_def: "k_impl (L, s) = IArray (map (\<lambda> c. Max {k_i2 i c | i. i < n_ps}) [0..<m+1])"
+    unfolding k_impl_def k_i2_def by auto
+  have Max_cong: "Max {f i | i. i < n_ps} = Max {g i | i. i < n_ps}"
+    if "\<And> i. i < n_ps \<Longrightarrow> f i = g i" for f g :: "nat \<Rightarrow> int"
+    by (rule arg_cong[where f = Max]) (force simp: that)
+  from that n_ps_gt_0 show ?thesis
+    unfolding k_impl_alt_def
+    unfolding k_i2_def[symmetric]
+    apply (clarsimp simp: k_fun_def k_i2_k cong: Max_cong)
+    apply safe
+    subgoal
+      by (subst Max_int_commute; force simp: setcompr_eq_image image_comp comp_def)
+    subgoal
+      using k_0 L_i_len[OF _ \<open>L \<in> states\<close>] by (intro linorder_class.Max_eqI) auto
+    subgoal
+      by (subst Max_int_commute; force simp: setcompr_eq_image image_comp comp_def)
+    done
+qed
+
+
+sublocale impl: Reachability_Problem_Impl
+  where trans_fun = trans_from
+  and inv_fun = inv_fun
+  and F_fun = Fi
+  and ceiling = k_impl
+  and A = prod_ta
+  and l0 = l0
+  and l0i = l0i
+  and F = "PR_CONST F"
+  and n = m
+  and k = k_fun
+  and trans_impl = trans_impl
+  and states' = states'
+  and loc_rel = loc_rel
+  apply standard
+
+(* trans_from *)
+  subgoal
+    unfolding trans_of_prod by (rule trans_from_correct)
+
+(* trans_impl *)
+  subgoal
+    apply (rule trans_from_refine)
+    done
+
+(* inv_fun *)
+  subgoal
+    unfolding trans_of_prod
+    by (rule set_mp[OF _ inv_fun_inv_of'[where R = loc_rel and S = "{(s, s'). state_rel s' s}"]])
+       (auto simp: loc_rel_def)
+
+(* state set *)
+  subgoal
+    using states'_superset by simp
+
+(* loc_rel l0 l0i*)
+  subgoal
+    using state_rel_start unfolding loc_rel_def by auto
+
+(* loc_rel left unique *)
+  subgoal for l li li'
+    unfolding trans_of_prod by (rule state_rel_left_unique)
+
+(* loc_rel right unique *)
+  subgoal for l l' li
+    unfolding trans_of_prod by (rule state_rel_right_unique)
+
+(* ceiling *)
+  subgoal
+    unfolding inv_rel_def using L0_states
+    by (auto simp: loc_rel_def state_rel_def reach.k'_def k_fun_def k_impl_k_fun)
+
+(* F_fun *)
+  subgoal
+    unfolding inv_rel_def by (clarsimp dest!: F_Fi)
+
+  done  \<comment> \<open>slow: 50s\<close>
+
+end (* Simple_Network_Impl_nat_ceiling_start_state *)
+
+no_notation UPPAAL_Model_Checking.models ("_,_ \<Turnstile>_ _" [61,61] 61)
+
+
+context Reachability_Problem_Impl
+begin
+
+lemma F_reachable_correct:
+  "op.F_reachable
+  \<longleftrightarrow> (\<exists>l'. \<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> (\<exists> u'. conv_A A \<turnstile>' \<langle>l0, u0\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l'))"
+  using E_op''.E_from_op_reachability_check[symmetric] reachability_check_new
+  unfolding E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+  unfolding F_rel_def by auto
+
+lemma E_op''_F_reachable_correct:
+  "E_op''.F_reachable
+  \<longleftrightarrow> (\<exists>l'. \<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> (\<exists> u'. conv_A A \<turnstile>' \<langle>l0, u0\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l'))"
+  using E_op''.E_from_op_reachability_check[symmetric] reachability_check_new
+  unfolding E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+  unfolding F_rel_def by auto
+
+lemma Ex_ev_impl_hnr:
+  assumes "\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0)"
+  shows
+    "
+  (uncurry0
+    (pw_impl (return \<circ> fst) state_copy_impl tracei subsumes_impl a0_impl F_impl succs_impl
+      emptiness_check_impl),
+   uncurry0 (SPEC (\<lambda>r. (r \<longleftrightarrow> (\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> Ex_ev (\<lambda>(l, _). F l) (l0, u0))))))
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+proof -
+  interpret Bisimulation_Invariant
+    "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). l' = l \<and> (\<forall> c. c \<in> clk_set (conv_A A) \<longrightarrow> u c = u' c))"
+    "(\<lambda>_. True)" "(\<lambda>_. True)"
+    apply (rule ta_bisimulation)
+    done
+  define spec where "spec = (\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> Ex_ev (\<lambda>(l, _). F l) (l0, u0))"
+  have *: "E_op''.F_reachable \<longleftrightarrow> spec"
+    unfolding spec_def
+    unfolding E_op''_F_reachable_correct
+  proof safe
+    fix l' :: \<open>'s\<close> and u0 :: \<open>nat \<Rightarrow> real\<close>
+    assume
+      \<open>\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> (\<exists>u'. conv_A A \<turnstile>' \<langle>l0, u0\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l')\<close> and
+      \<open>\<forall>c\<in>{1..n}. u0 c = 0\<close>
+    then show \<open>Ex_ev (\<lambda>(l, _). F l) (l0, u0)\<close>
+      using assms by (subst Ex_ev; unfold reaches_steps'[symmetric]) blast+
+  next
+    assume \<open>\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> Ex_ev (\<lambda>(l, _). F l) (l0, u0)\<close>
+    then have "Ex_ev (\<lambda>(l, _). F l) (l0, \<lambda>_. 0)"
+      by auto
+    then obtain l' u' where "conv_A A \<turnstile>' \<langle>l0, (\<lambda>_. 0)\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l'"
+      apply (subst (asm) Ex_ev)
+      using assms
+      unfolding reaches_steps'[symmetric]
+      by auto
+    then show \<open>\<exists>l'. \<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> (\<exists>u'. conv_A A \<turnstile>' \<langle>l0, u0\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l')\<close>
+    proof (inst_existentials l', safe, unfold reaches_steps'[symmetric])
+      fix u0 :: \<open>nat \<Rightarrow> real\<close>
+      assume \<open>reaches (l0, \<lambda>_. 0) (l', u')\<close> and \<open>F l'\<close> and \<open>\<forall>c\<in>{1..n}. u0 c = 0\<close>
+      then have "equiv' (l0, \<lambda>_. 0) (l0, u0)"
+        unfolding equiv'_def using clocks_I[of "\<lambda>_. 0" u0] by auto
+      with \<open>reaches _ _\<close> \<open>F l'\<close> show \<open>\<exists>u'. reaches (l0, u0) (l', u') \<and> F l'\<close>
+        by - (drule (1) bisim.A_B.simulation_reaches, unfold equiv'_def, auto)
+    qed
+  qed
+  show ?thesis
+    unfolding spec_def[symmetric] using pw_impl_hnr_F_reachable[to_hnr, unfolded hn_refine_def]
+    by sepref_to_hoare (sep_auto simp: *)
+qed
+
+end (* Reachability Problem Impl *)
+
+
+context Simple_Network_Impl_nat_urge
+begin
+
+sublocale conv: Prod_TA_urge
+  "(set broadcast, map (Simple_Network_Language.conv_A o automaton_of) automata, map_of bounds')"
+  by standard
+     (use no_urgency in \<open>auto simp: Simple_Network_Language.conv_A_def automaton_of_def urgent_def\<close>)
+
+abbreviation "A \<equiv> (set broadcast, map automaton_of automata, map_of bounds')"
+
+interpretation conv_eq_bisim:
+  Bisimulation_Invariant
+  "(\<lambda>(l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  "(\<lambda>(L, s, u) (L', s', u'). conv.prod_ta   \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+  "(\<lambda>((L, s), u) (L', s', u'). L = L' \<and> u = u' \<and> s = s')"
+  "(\<lambda>((L, s), u). conv.all_prop L s)"
+  "(\<lambda>(L, s, u). conv.all_prop L s)"
+proof goal_cases
+  case 1
+  interpret Bisimulation_Invariant
+  "(\<lambda>(L, s, u) (L', s', u'). conv_A prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+  "(\<lambda>(L, s, u) (L', s', u'). conv.prod_ta   \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+  "(=)"
+  "(\<lambda>(L, s, u). conv.all_prop L s)"
+  "(\<lambda>(L, s, u). conv.all_prop L s)"
+  by (rule Bisimulation_Invariant_strong_intro)
+     (auto simp: conv_prod_ta elim: conv.prod_all_prop_inv')
+  show ?case
+    by standard (auto simp: conv_prod_ta elim: conv.prod_all_prop_inv')
+qed
+
+interpretation Bisimulation_Invariant
+  "(\<lambda>(l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  "(\<lambda>(L, s, u) (L', s', u'). Simple_Network_Language.conv A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)"
+  "(\<lambda>((L, s), u) (L', s', u'). L = L' \<and> u = u' \<and> s = s')"
+  "(\<lambda>((L, s), u). conv.all_prop L s)"
+  "(\<lambda>(L, s, u). conv.all_prop L s)"
+  unfolding conv_alt_def
+  apply (rule Bisimulation_Invariant_sim_replace, rule Bisimulation_Invariant_composition)
+    apply (rule conv_eq_bisim.Bisimulation_Invariant_axioms conv.prod_bisim_intro)+
+  apply auto
+  done
+
+lemmas prod_bisim = Bisimulation_Invariant_axioms
+
+lemmas deadlock_iff = deadlock_iff
+
+lemma conv_all_prop:
+  "conv.all_prop = all_prop"
+  unfolding conv.all_prop_def all_prop_def by simp
+
+definition prop_of2 where
+  "prop_of2 \<phi> = PropC (\<lambda>((L, s), _). check_sexp \<phi> L (the o s))"
+
+fun ctl_of2 where
+  "ctl_of2 (formula.EX \<phi>) = ctl_formula.EX (prop_of2 \<phi>)"
+| "ctl_of2 (formula.EG \<phi>) = ctl_formula.EG (prop_of2 \<phi>)"
+| "ctl_of2 (formula.AX \<phi>) = ctl_formula.AX (prop_of2 \<phi>)"
+| "ctl_of2 (formula.AG \<phi>) = ctl_formula.AG (prop_of2 \<phi>)"
+| "ctl_of2 (formula.Leadsto \<phi> \<psi>) =
+  ctl_formula.AG (ctl_formula.ImpliesC (prop_of2 \<phi>) (ctl_formula.AX (prop_of2 \<psi>)))"
+
+lemma models_correct:
+  "Simple_Network_Language.conv A,(L0, s0, u0) \<Turnstile> \<Phi> = (case \<Phi> of
+    formula.EX \<phi> \<Rightarrow>
+      Graph_Defs.Ex_ev
+        (\<lambda> (l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). check_sexp \<phi> L (the o s))
+  | formula.EG \<phi> \<Rightarrow>
+      Not o Graph_Defs.Alw_ev
+        (\<lambda> (l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). \<not> check_sexp \<phi> L (the o s))
+  | formula.AX \<phi> \<Rightarrow>
+      Graph_Defs.Alw_ev
+        (\<lambda> (l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). check_sexp \<phi> L (the o s))
+  | formula.AG \<phi> \<Rightarrow>
+      Not o Graph_Defs.Ex_ev
+        (\<lambda> (l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). \<not> check_sexp \<phi> L (the o s))
+  | formula.Leadsto \<phi> \<psi> \<Rightarrow>
+      Graph_Defs.leadsto
+        (\<lambda> (l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). check_sexp \<phi> L (the o s))
+        (\<lambda> ((L, s), _). check_sexp \<psi> L (the o s))
+  ) ((L0, s0), u0)
+  " if "L0 \<in> states" "Simple_Network_Language.bounded bounds s0"
+proof -
+  have *: "((Not \<circ>\<circ> case_prod) (\<lambda>(L, s) _. check_sexp \<phi> L (the o s)))
+    = (\<lambda>((L, s), _). \<not> check_sexp \<phi> L (the o s))"
+    for \<phi> by (auto simp: comp_def)
+  have "rel_ctl_formula compatible (ctl_of2 \<Phi>) (ctl_of \<Phi>)"
+    by (cases \<Phi>; simp add: prop_of_def prop_of2_def rel_fun_def A_B.equiv'_def)
+  moreover have "A_B.equiv' ((L0, s0), u0) (L0, s0, u0)"
+    unfolding A_B.equiv'_def unfolding conv_all_prop all_prop_def using that by simp
+  ultimately show ?thesis
+    apply (subst models_ctl_iff)
+    apply (subst CTL_compatible[THEN rel_funD, symmetric])
+      apply assumption+
+    apply (cases \<Phi>; simp add:
+        Graph_Defs.models_ctl.simps prop_of2_def
+        Graph_Defs.Ex_alw_iff Graph_Defs.Alw_alw_iff Graph_Defs.leadsto_def *)
+    done
+qed
+
+end (* Simple_Network_Impl_nat *)
+
+context Simple_Network_Impl_nat_ceiling_start_state \<comment> \<open>slow: 70s\<close>
+begin
+
+definition Alw_ev_checker where
+  "Alw_ev_checker = dfs_map_impl'
+     (impl.succs_P_impl' Fi) impl.a0_impl impl.subsumes_impl (return \<circ> fst)
+     impl.state_copy_impl"
+
+definition leadsto_checker where
+  "leadsto_checker \<psi> = do {
+      r \<leftarrow> leadsto_impl
+      impl.state_copy_impl (impl.succs_P_impl' (\<lambda> (L, s). \<not> check_sexpi \<psi> L s))
+      impl.a0_impl impl.subsumes_impl (return \<circ> fst)
+      impl.succs_impl' impl.emptiness_check_impl impl.F_impl
+      (impl.Q_impl (\<lambda> (L, s). \<not> check_sexpi \<psi> L s))
+      impl.tracei;
+      return (\<not> r)
+    }"
+
+definition
+  "reachability_checker \<equiv>
+     pw_impl
+      (return o fst) impl.state_copy_impl impl.tracei impl.subsumes_impl impl.a0_impl impl.F_impl
+      impl.succs_impl impl.emptiness_check_impl"
+
+definition model_checker where
+  "model_checker = (
+  case formula of
+    formula.EX _ \<Rightarrow> reachability_checker |
+    formula.AG _ \<Rightarrow> do {
+      r \<leftarrow> reachability_checker;
+      return (\<not> r)
+    } |
+    formula.AX _ \<Rightarrow> do {
+      r \<leftarrow> if PR_CONST F l0
+      then Alw_ev_checker
+      else return False;
+      return (\<not> r)
+    } |
+    formula.EG _ \<Rightarrow>
+      if PR_CONST F l0
+      then Alw_ev_checker
+      else return False |
+    formula.Leadsto _ \<psi> \<Rightarrow> leadsto_checker \<psi>
+  )
+  "
+
+abbreviation "u0 \<equiv> (\<lambda> _. 0 :: real)"
+
+lemma all_prop_start:
+  "all_prop L0 (map_of s0)"
+  using L0_states s0_bounded unfolding all_prop_def ..
+
+lemma deadlock_start_iff:
+  "Graph_Defs.deadlock
+   (\<lambda>(L, s, u) (L', s', u'). Simple_Network_Language.conv A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>) (L0, (map_of s0), u0)
+  \<longleftrightarrow> reach.deadlock (l0, u0)"
+  using all_prop_start by (subst deadlock_iff[symmetric]) (auto simp: conv_all_prop)
+
+lemma F_Fi':
+  "check_sexp \<psi> L (the o s) \<longleftrightarrow> check_sexpi \<psi> L' s'"
+  if "((L', s'), (L, s)) \<in> loc_rel" "formula = Leadsto \<phi> \<psi>"
+  using vars_of_formula that unfolding loc_rel_def by (auto elim!: check_sexp_check_sexpi)
+
+theorem model_check':
+  "(uncurry0 model_checker,
+    uncurry0 (
+      SPEC (\<lambda>r.
+  \<not> Graph_Defs.deadlock
+    (\<lambda>(L, s, u) (L', s', u').
+      Simple_Network_Language.conv A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>) (L0, (map_of s0), u0)
+      \<longrightarrow> r = (Simple_Network_Language.conv A,(L0, (map_of s0), u0) \<Turnstile> formula)
+      )
+    )
+   )
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+proof -
+  define protect where
+    "protect = ((\<lambda>(l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>))"
+
+  have start: "l0 \<in> Normalized_Zone_Semantics_Impl_Refine.state_set trans_prod"
+    if "\<not> Graph_Defs.deadlock protect (l0, \<lambda>_. 0)"
+    using that unfolding protect_def by (rule impl.init_state_in_state_set[simplified])
+
+  interpret ta_bisim: Bisimulation_Invariant
+    "(\<lambda>(l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). conv_A prod_ta \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). l' = l \<and> (\<forall> c. c \<in> clk_set (conv_A prod_ta) \<longrightarrow> u c = u' c))"
+    "(\<lambda>_. True)" "(\<lambda>_. True)"
+    by (rule ta_bisimulation[of "conv_A prod_ta"])
+
+  let ?\<phi>1 = "\<lambda>\<phi>. \<lambda>(p, _). case p of (L, s) \<Rightarrow> \<not> check_sexp \<phi> L (the \<circ> s)"
+  let ?\<phi>2 = "\<lambda>\<phi>. \<lambda>(p, _). case p of (L, s) \<Rightarrow> check_sexp \<phi> L (the \<circ> s)"
+
+  have start_sim:
+    "ta_bisim.A_B.equiv' (l0, u) (l0, \<lambda>_. 0)" "ta_bisim.A_B.equiv' (l0, \<lambda>_. 0) (l0, u)"
+    if "\<forall>c\<in>{Suc 0..m}. u c = 0" for u
+    using impl.clocks_I[of u "\<lambda>_. 0"] that unfolding ta_bisim.A_B.equiv'_def by auto
+
+  have compatibleI: "\<phi> a = \<phi> b"
+    if "ta_bisim.A_B.equiv' a b" "\<And> l u u'. \<phi> (l, u) = \<phi> (l, u')" for a b \<phi>
+    using that unfolding ta_bisim.A_B.equiv'_def by auto
+
+  have bisims:
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> reach.Ex_ev (?\<phi>1 \<phi>) (l0, u0))
+    \<longleftrightarrow> reach.Ex_ev (?\<phi>1 \<phi>) (l0, u0)"
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> reach.Ex_ev (?\<phi>2 \<phi>) (l0, u0))
+    \<longleftrightarrow> reach.Ex_ev (?\<phi>2 \<phi>) (l0, u0)"
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> reach.Alw_ev (?\<phi>1 \<phi>) (l0, u0))
+    \<longleftrightarrow> reach.Alw_ev (?\<phi>1 \<phi>) (l0, u0)"
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> reach.Alw_ev (?\<phi>2 \<phi>) (l0, u0))
+    \<longleftrightarrow> reach.Alw_ev (?\<phi>2 \<phi>) (l0, u0)"
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> reach.leadsto (?\<phi>2 \<phi>) (?\<phi>2 \<psi>) (l0, u0))
+    \<longleftrightarrow> reach.leadsto (?\<phi>2 \<phi>) (?\<phi>2 \<psi>) (l0, u0)"
+    for \<phi> \<psi>
+    apply safe
+    apply (all \<open>(solves auto)?\<close>)
+    subgoal for u
+      using start_sim
+      by (subst (asm) ta_bisim.Ex_ev_iff[of "?\<phi>1 \<phi>" "?\<phi>1 \<phi>"]) (erule compatibleI, auto)
+    subgoal for u
+      using start_sim
+      by (subst (asm) ta_bisim.Ex_ev_iff[of "?\<phi>2 \<phi>" "?\<phi>2 \<phi>"]) (erule compatibleI, auto)
+    subgoal for u
+      using start_sim
+      by (subst (asm) ta_bisim.Alw_ev_iff[of "?\<phi>1 \<phi>" "?\<phi>1 \<phi>"]) (erule compatibleI, auto)
+    subgoal for u
+      using start_sim
+      by (subst (asm) ta_bisim.Alw_ev_iff[of "?\<phi>2 \<phi>" "?\<phi>2 \<phi>"]) (erule compatibleI, auto)
+    subgoal for u
+      using start_sim
+      by (subst (asm) ta_bisim.Leadsto_iff[of "?\<phi>2 \<phi>" "?\<phi>2 \<phi>" "?\<phi>2 \<psi>" "?\<phi>2 \<psi>"])
+        ((erule compatibleI)?; auto)+
+    done
+
+  have deadlock_bisim:
+    "(\<forall>u0. (\<forall>c\<in>{1..m}. u0 c = 0) \<longrightarrow> \<not> reach.deadlock (l0, u0))
+    \<longleftrightarrow> \<not> Graph_Defs.deadlock protect (l0, \<lambda>_. 0)"
+    unfolding protect_def
+    apply (safe; clarsimp)
+    apply (drule start_sim(2))
+    apply (subst (asm) ta_bisim.deadlock_iff)
+    unfolding ta_bisim.A_B.equiv'_def
+    by auto
+
+  have *****:
+    "return True = (return False \<bind> return o Not)"
+    by auto
+
+  show ?thesis
+    unfolding deadlock_start_iff
+    using models_correct[OF L0_states s0_bounded]
+    unfolding protect_def[symmetric]
+    apply (simp split del: if_split)
+    apply sepref_to_hoare
+    apply sep_auto
+    unfolding model_checker_def Alw_ev_checker_def leadsto_checker_def reachability_checker_def
+    apply (cases formula; simp)
+
+    subgoal
+      apply (cases "Graph_Defs.deadlock protect (l0, \<lambda>_. 0)")
+      subgoal
+        using impl.pw_impl_hnr_F_reachable[to_hnr, unfolded hn_refine_def]
+        by (sep_auto elim!: cons_post_rule)
+      subgoal
+        using impl.Ex_ev_impl_hnr[unfolded deadlock_bisim, to_hnr, unfolded hn_refine_def]
+        by (sep_auto simp: pure_def protect_def bisims)
+      done
+
+    subgoal premises prems for \<phi>
+    proof -
+      have *: "(\<lambda>(l, u). \<not> F l) = ((\<lambda>(p, _). case p of (L, s) \<Rightarrow> \<not> check_sexp \<phi> L (the \<circ> s)))"
+        using prems by auto
+      show ?thesis
+        using impl.Alw_ev_impl_hnr[to_hnr, unfolded hn_refine_def] prems start
+        unfolding PR_CONST_def * protect_def
+        by (sep_auto elim!: cons_post_rule simp: pure_def bisims)
+    qed
+
+    subgoal premises prems for \<phi>
+    proof -
+      have *: "(\<lambda>(l, u). \<not> F l) = ((\<lambda>(p, _). case p of (L, s) \<Rightarrow> check_sexp \<phi> L (the \<circ> s)))"
+        using prems by auto
+      show ?thesis
+        apply intros
+        subgoal
+          using impl.Alw_ev_impl_hnr[to_hnr, unfolded hn_refine_def] prems start
+          unfolding PR_CONST_def * protect_def
+          apply simp
+          unfolding *****
+          apply (erule bind_rule)
+          apply (sep_auto simp: pure_def bisims(2-))
+          done
+        subgoal
+          using impl.Alw_ev_impl_hnr[to_hnr, unfolded hn_refine_def] prems start
+          unfolding PR_CONST_def * protect_def
+          apply (sep_auto elim!: cons_post_rule simp: pure_def bisims(2-))
+          done
+        done
+    qed
+
+    subgoal
+      apply (cases "Graph_Defs.deadlock protect (l0, \<lambda>_. 0)")
+      subgoal
+        using impl.pw_impl_hnr_F_reachable[to_hnr, unfolded hn_refine_def]
+        by (sep_auto elim!: cons_post_rule)
+      subgoal
+        using impl.Ex_ev_impl_hnr[unfolded deadlock_bisim, to_hnr, unfolded hn_refine_def]
+        apply (sep_auto simp: pure_def protect_def bisims)
+        done
+      done
+
+    subgoal premises prems for \<phi> \<psi>
+    proof -
+      have *: "(\<lambda>(l, u). F l) = ((\<lambda>(p, _). case p of (L, s) \<Rightarrow> check_sexp \<phi> L (the \<circ> s)))"
+        using prems by auto
+      have **:"(\<lambda>(L, s). \<not> check_sexpi \<psi> L s, \<lambda>(L, s). \<not> check_sexp \<psi> L (the \<circ> s))
+               \<in> inv_rel loc_rel states'"
+        unfolding trans_of_prod using prems by (auto simp: F_Fi' inv_rel_def)
+      have ****: "(\<lambda>(l, u). \<not> (case l of (L, s) \<Rightarrow> \<not> check_sexp \<psi> L (the \<circ> s)))
+      = (\<lambda>(l, u). (case l of (L, s) \<Rightarrow> check_sexp \<psi> L (the \<circ> s)))"
+        by auto
+      show ?thesis
+        apply (cases "reach.deadlock (l0, \<lambda>_. 0)")
+        subgoal
+          using impl.leadsto_impl_hnr'[OF **, to_hnr, unfolded hn_refine_def]
+          unfolding protect_def \<open>formula = _\<close> by (sep_auto simp: pure_def)
+        using impl.leadsto_impl_hnr[unfolded deadlock_bisim, to_hnr, unfolded hn_refine_def, OF **]
+          prems start
+        unfolding PR_CONST_def * protect_def by (sep_auto simp: pure_def bisims ****)
+    qed
+    done
+qed
+
+
+subsection \<open>Extracting an Efficient Implementation\<close>
+
+lemma reachability_checker_alt_def':
+  "reachability_checker \<equiv>
+    do {
+      x \<leftarrow> do {
+        let key = return \<circ> fst;
+        let sub = impl.subsumes_impl;
+        let copy = impl.state_copy_impl;
+        let start = impl.a0_impl;
+        let final = impl.F_impl;
+        let succs =  impl.succs_impl;
+        let empty = impl.emptiness_check_impl;
+        let trace = impl.tracei;
+        pw_impl key copy trace sub start final succs empty
+      };
+      _ \<leftarrow> return ();
+      return x
+    }"
+  unfolding reachability_checker_def by simp
+
+lemma Alw_ev_checker_alt_def':
+  "Alw_ev_checker \<equiv>
+    do {
+      x \<leftarrow> let
+        key = return \<circ> fst;
+        sub = impl.subsumes_impl;
+        copy = impl.state_copy_impl;
+        start = impl.a0_impl;
+        succs =  impl.succs_P_impl' Fi
+      in dfs_map_impl' succs start sub key copy;
+      _ \<leftarrow> return ();
+      return x
+    }"
+  unfolding Alw_ev_checker_def by simp
+
+lemma leadsto_checker_alt_def':
+  "leadsto_checker \<psi> \<equiv>
+    do {
+      r \<leftarrow> let
+        key = return \<circ> fst;
+        sub = impl.subsumes_impl;
+        copy = impl.state_copy_impl;
+        start = impl.a0_impl;
+        final = impl.F_impl;
+        final' = (impl.Q_impl (\<lambda>(L, s). \<not> check_sexpi \<psi> L s));
+        succs =  impl.succs_P_impl' (\<lambda>(L, s). \<not> check_sexpi \<psi> L s);
+        succs' =  impl.succs_impl';
+        empty = impl.emptiness_check_impl;
+        trace = impl.tracei
+      in
+        leadsto_impl copy succs start sub key succs' empty final final' trace;
+      return (\<not> r)
+    }"
+  unfolding leadsto_checker_def by simp
+
+theorem k_impl_alt_def:
+  "k_impl = (\<lambda>(l, s). IArray (map (\<lambda>c. MAX i \<in> {0..<n_ps}. k_i !! i !! (l ! i) !! c) [0..<m + 1]))"
+proof -
+  have "{i. i < p} = {0..<p}" for p :: nat
+    by auto
+  then show ?thesis
+    unfolding k_impl_def setcompr_eq_image by simp
+qed
+
+definition
+  "trans_map_inner \<equiv> map (\<lambda>i. union_map_of (fst (snd (snd (automata ! i))))) [0..<n_ps]"
+
+lemma trans_map_alt_def:
+  "trans_map i j = (case (IArray trans_map_inner !! i) j of None \<Rightarrow> [] | Some xs \<Rightarrow> xs)"
+  if "i < n_ps"
+  using that unfolding trans_map_inner_def trans_map_def by (auto simp: n_ps_def)
+
+schematic_goal succs_impl_alt_def:
+  "impl.succs_impl \<equiv> ?impl"
+  unfolding impl.succs_impl_def
+  apply (abstract_let impl.E_op''_impl E_op''_impl)
+  unfolding impl.E_op''_impl_def fw_impl'_int
+  apply (abstract_let trans_impl trans_impl)
+  unfolding trans_impl_def
+  apply (abstract_let int_trans_impl int_trans_impl)
+  apply (abstract_let bin_trans_from_impl bin_trans_impl)
+  apply (abstract_let broad_trans_from_impl broad_trans_impl)
+  unfolding int_trans_impl_def bin_trans_from_impl_def broad_trans_from_impl_def
+  apply (abstract_let trans_in_broad_grouped trans_in_broad_grouped)
+  apply (abstract_let trans_out_broad_grouped trans_out_broad_grouped)
+  apply (abstract_let trans_in_map trans_in_map)
+  apply (abstract_let trans_out_map trans_out_map)
+  apply (abstract_let int_trans_from_all_impl int_trans_from_all_impl)
+  unfolding int_trans_from_all_impl_def
+  apply (abstract_let int_trans_from_vec_impl int_trans_from_vec_impl)
+  unfolding int_trans_from_vec_impl_def
+  apply (abstract_let int_trans_from_loc_impl int_trans_from_loc_impl)
+  unfolding int_trans_from_loc_impl_def
+  apply (abstract_let trans_i_map trans_i_map)
+  unfolding trans_out_broad_grouped_def trans_out_broad_map_def
+  unfolding trans_in_broad_grouped_def trans_in_broad_map_def
+  unfolding trans_in_map_def trans_out_map_def
+  unfolding trans_i_map_def
+  apply (abstract_let trans_map trans_map)
+  apply (abstract_let "inv_fun :: nat list \<times> int list \<Rightarrow> _" inv_fun)
+  unfolding inv_fun_alt_def
+  apply (abstract_let invs2 invs)
+  unfolding invs2_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i) (* Could be killed *)
+  apply (abstract_let n_ps n_ps)
+  by (rule Pure.reflexive)
+
+schematic_goal succs_P_impl_alt_def:
+  "impl.succs_P_impl Pi \<equiv> ?impl"
+  if "(Pi, P) \<in> inv_rel loc_rel states'"
+  for P Pi
+  unfolding impl.succs_P_impl_def[OF that]
+  apply (abstract_let impl.E_op''_impl E_op''_impl)
+  unfolding impl.E_op''_impl_def fw_impl'_int
+  apply (abstract_let "trans_impl" trans_impl)
+  unfolding trans_impl_def
+  apply (abstract_let "int_trans_impl" int_trans_impl)
+  apply (abstract_let "bin_trans_from_impl" bin_trans_impl)
+  apply (abstract_let "broad_trans_from_impl" broad_trans_impl)
+  unfolding int_trans_impl_def bin_trans_from_impl_def broad_trans_from_impl_def
+  apply (abstract_let trans_in_broad_grouped trans_in_broad_grouped)
+  apply (abstract_let trans_out_broad_grouped trans_out_broad_grouped)
+  apply (abstract_let trans_in_map trans_in_map)
+  apply (abstract_let trans_out_map trans_out_map)
+  apply (abstract_let int_trans_from_all_impl int_trans_from_all_impl)
+  unfolding int_trans_from_all_impl_def
+  apply (abstract_let int_trans_from_vec_impl int_trans_from_vec_impl)
+  unfolding int_trans_from_vec_impl_def
+  apply (abstract_let int_trans_from_loc_impl int_trans_from_loc_impl)
+  unfolding int_trans_from_loc_impl_def
+  apply (abstract_let trans_i_map trans_i_map)
+  unfolding trans_out_broad_grouped_def trans_out_broad_map_def
+  unfolding trans_in_broad_grouped_def trans_in_broad_map_def
+  unfolding trans_in_map_def trans_out_map_def
+  unfolding trans_i_map_def
+  apply (abstract_let trans_map trans_map)
+  apply (abstract_let "inv_fun :: nat list \<times> int list \<Rightarrow> _" inv_fun)
+  unfolding inv_fun_alt_def
+  apply (abstract_let invs2 invs)
+  unfolding invs2_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i) (* Could be killed *)
+  apply (abstract_let n_ps n_ps)
+  by (rule Pure.reflexive)
+
+
+(* These implementations contain unnecessary list reversals *)
+lemmas succs_P'_impl_alt_def =
+  impl.succs_P_impl'_def[OF impl.F_fun, unfolded succs_P_impl_alt_def[OF impl.F_fun]]
+
+schematic_goal Alw_ev_checker_alt_def:
+  "Alw_ev_checker \<equiv> ?impl"
+  unfolding Alw_ev_checker_alt_def'
+  unfolding succs_P'_impl_alt_def
+  unfolding k_impl_alt_def k_i_def
+  (* The following are just to unfold things that should have been defined in a defs locale *)
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+lemma \<psi>_compatibleI:
+  "(\<lambda>(L, s). \<not> check_sexpi \<psi> L s, \<lambda>(L, s). \<not> check_sexp \<psi> L (the \<circ> s)) \<in> inv_rel loc_rel states'"
+  if "formula = Leadsto \<phi> \<psi>"
+  using F_Fi'[OF _ that] unfolding inv_rel_def by auto
+
+lemma Q_impl_alt_def:
+  "impl.Q_impl (\<lambda>(L, s). \<not> check_sexpi \<psi> L s) \<equiv>
+  \<lambda>xi. return (case xi of (a1, a2) \<Rightarrow> (\<lambda>(L, s). \<not> check_sexpi \<psi> L s) a1)"
+  if "formula = Leadsto \<phi> \<psi>"
+  by (intro impl.Q_impl_def[where Q = "\<lambda>(L, s). \<not> check_sexp \<psi> L (the o s)"] \<psi>_compatibleI[OF that])
+
+schematic_goal leadsto_checker_alt_def:
+  "leadsto_checker \<psi> \<equiv> ?impl" if "formula = Leadsto \<phi> \<psi>"
+  unfolding leadsto_checker_alt_def'
+  unfolding Q_impl_alt_def[OF that] impl.F_impl_def
+  unfolding impl.succs_P_impl'_def[OF \<psi>_compatibleI[OF that]]
+  unfolding succs_P_impl_alt_def[OF \<psi>_compatibleI[OF that]]
+  unfolding impl.succs_impl'_def succs_impl_alt_def
+  unfolding k_impl_alt_def k_i_def
+  (* The following are just to unfold things that should have been defined in a defs locale *)
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+schematic_goal reachability_checker_alt_def:
+  "reachability_checker \<equiv> ?impl"
+  unfolding reachability_checker_alt_def'
+  unfolding succs_impl_alt_def
+  unfolding k_impl_alt_def k_i_def
+  (* The following are just to unfold things that should have been defined in a defs locale *)
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+end (* Simple_Network_Impl_nat_ceiling_start_state *)
+
+concrete_definition reachability_checker
+  uses Simple_Network_Impl_nat_ceiling_start_state.reachability_checker_alt_def
+
+concrete_definition Alw_ev_checker
+  uses Simple_Network_Impl_nat_ceiling_start_state.Alw_ev_checker_alt_def
+
+concrete_definition leadsto_checker
+  uses Simple_Network_Impl_nat_ceiling_start_state.leadsto_checker_alt_def
+
+context Simple_Network_Impl_nat_ceiling_start_state \<comment> \<open>slow: 100s\<close>
+begin
+
+lemma model_checker_unfold_leadsto:
+  "model_checker = (
+  case formula of Simple_Network_Language_Model_Checking.formula.EX xa \<Rightarrow> reachability_checker
+    | Simple_Network_Language_Model_Checking.formula.EG xa \<Rightarrow>
+      if PR_CONST F l0 then Alw_ev_checker else return False
+    | Simple_Network_Language_Model_Checking.formula.AX xa \<Rightarrow>
+      (if PR_CONST F l0 then Alw_ev_checker else return False) \<bind> (\<lambda>r. return (\<not> r))
+    | Simple_Network_Language_Model_Checking.formula.AG xa \<Rightarrow>
+      reachability_checker \<bind> (\<lambda>r. return (\<not> r))
+    | Simple_Network_Language_Model_Checking.formula.Leadsto \<phi> \<psi> \<Rightarrow>
+      Simple_Network_Language_Model_Checking.leadsto_checker
+        broadcast bounds' automata m num_states num_actions k L0 s0 formula \<psi> show_clock show_state)
+"
+  unfolding model_checker_def
+  using leadsto_checker.refine[OF Simple_Network_Impl_nat_ceiling_start_state_axioms]
+  by (simp split: formula.split)
+
+lemmas model_checker_def_refined = model_checker_unfold_leadsto[unfolded
+    reachability_checker.refine[OF Simple_Network_Impl_nat_ceiling_start_state_axioms]
+    Alw_ev_checker.refine[OF Simple_Network_Impl_nat_ceiling_start_state_axioms]
+  ]
+
+end (* Simple_Network_Impl_nat_ceiling_start_state *)
+
+concrete_definition model_checker uses
+  Simple_Network_Impl_nat_ceiling_start_state.model_checker_def_refined
+
+definition precond_mc where
+  "precond_mc
+    show_clock show_state broadcast bounds' automata m num_states num_actions k L0 s0 formula \<equiv>
+    if Simple_Network_Impl_nat_ceiling_start_state
+      broadcast bounds' automata m num_states num_actions k L0 s0 formula
+    then
+      model_checker
+        broadcast bounds' automata m num_states num_actions k L0 s0 formula show_clock show_state
+      \<bind> (\<lambda> x. return (Some x))
+    else return None"
+
+
+abbreviation N where
+  "N broadcast automata bounds \<equiv>
+  Simple_Network_Language.conv
+    (set broadcast, map automaton_of automata, map_of bounds)"
+
+definition has_deadlock where
+  "has_deadlock A a0 \<equiv>
+    Graph_Defs.deadlock (\<lambda> (L, s, u) (L', s', u'). A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>) a0"
+
+theorem model_check:
+  "<emp> precond_mc
+    show_clock show_state broadcast bounds automata m num_states num_actions k L0 s0 formula
+    <\<lambda> Some r \<Rightarrow> \<up>(
+        Simple_Network_Impl_nat_ceiling_start_state
+          broadcast bounds automata m num_states num_actions k L0 s0 formula \<and>
+        (\<not> has_deadlock (N broadcast automata bounds) (L0, map_of s0, \<lambda>_ . 0) \<longrightarrow>
+          r = N broadcast automata bounds,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> formula
+        ))
+     | None \<Rightarrow> \<up>(\<not> Simple_Network_Impl_nat_ceiling_start_state
+        broadcast bounds automata m num_states num_actions k L0 s0 formula)
+    >t"
+proof -
+  define A where "A \<equiv> N broadcast automata bounds"
+  define check where "check \<equiv> A,(L0, map_of s0, \<lambda>_ . 0) \<Turnstile> formula"
+  note [sep_heap_rules] =
+    Simple_Network_Impl_nat_ceiling_start_state.model_check'[
+      to_hnr, unfolded hn_refine_def,
+      of broadcast bounds automata m num_states num_actions k L0 s0 formula,
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      folded A_def check_def has_deadlock_def,
+      simplified
+      ]
+  show ?thesis
+    unfolding A_def[symmetric] check_def[symmetric]
+    unfolding precond_mc_def
+    by (sep_auto simp: model_checker.refine[symmetric] pure_def)
+qed
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Renaming.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/Simple_Network_Language_Renaming.thy
@@ -0,0 +1,3545 @@
+section \<open>Renaming Identifiers in Simple Networks\<close>
+
+theory Simple_Network_Language_Renaming
+  imports Simple_Network_Language_Model_Checking
+begin
+
+text \<open>The part justifies a ``renaming'' step to normalize identifiers in the input.\<close>
+
+unbundle no_library_syntax
+notation (input) TAG ("_ \<bbar> _" [40, 40] 41)
+
+text \<open>Helpful methods and theorems to work with tags:\<close>
+
+lemmas TAG_cong = arg_cong[where f = "TAG t" for t]
+
+lemma TAG_mp:
+  assumes "TAG t x" "x \<Longrightarrow> y"
+  shows "TAG t y"
+  using assms unfolding TAG_def by blast
+
+lemma TAG_mp':
+  assumes "TAG t x" "x \<Longrightarrow> y"
+  shows y
+  using assms unfolding TAG_def by blast
+
+
+lemmas all_mono_rule = all_mono[THEN mp, OF impI, rotated]
+
+lemma imp_mono_rule:
+  assumes "P1 \<longrightarrow> P2"
+    and "Q1 \<Longrightarrow> P1"
+    and "Q1 \<Longrightarrow> P2 \<Longrightarrow> Q2"
+  shows "Q1 \<longrightarrow> Q2"
+  using assms by blast
+
+lemma Ball_mono:
+  assumes "\<forall>x \<in> S. P x" "\<And>x. x \<in> S \<Longrightarrow> P x \<Longrightarrow> Q x"
+  shows "\<forall>x \<in> S. Q x"
+  using assms by blast
+
+method prop_monos =
+  erule all_mono_rule
+  | erule (1) imp_mono_rule
+  | erule disj_mono[rule_format, rotated 2]
+
+locale Prod_TA_Defs_urge =
+  fixes A :: "('a, 's, 'c, 't :: {zero}, 'x, 'v :: linorder) nta" and urge :: 'c
+begin
+
+definition
+  "add_reset \<equiv> \<lambda>(C, U, T, I).
+    (C, {}, (\<lambda>(l, b, g, a, f, r, l'). (l, b, g, a, f, urge # r, l')) ` T, I)"
+
+definition
+  "add_inv \<equiv> \<lambda>(C, U, T, I). (C, U, T,
+    (\<lambda>l. if l \<in> U then acconstraint.LE urge 0 # I l else I l))"
+
+definition "A_urge \<equiv>
+  (fst A, map add_reset (map add_inv (fst (snd A))), snd (snd A))"
+
+definition
+  "N_urge i \<equiv> map add_reset (map add_inv (fst (snd A))) ! i"
+
+end
+
+locale Prod_TA_sem_urge =
+  Prod_TA_Defs_urge A urge + Prod_TA_sem A
+  for A :: "('a, 's, 'c, 't :: {zero, time}, 'x, 'v :: linorder) nta" and urge :: 'c +
+  assumes urge_not_in_invariants:
+    "\<forall>(_, _, _, I) \<in> set (fst (snd A)). \<forall>l. urge \<notin> constraint_clk ` set (I l)"
+  assumes urge_not_in_guards:
+    "\<forall>(_, _, T, _) \<in> set (fst (snd A)). \<forall>(l, b, g, a, f, r, l') \<in> T. urge \<notin> constraint_clk ` set g"
+  assumes urge_not_in_resets:
+    "\<forall>(_, _, T, _) \<in> set (fst (snd A)). \<forall>(l, b, g, a, f, r, l') \<in> T. urge \<notin> set r"
+begin
+
+lemma N_urge_simp:
+  "N_urge i = add_reset (add_inv (N i))" if "i < n_ps"
+  using that unfolding N_urge_def N_def by (simp add: n_ps_def)
+
+lemma inv_add_reset:
+  "inv (add_reset A') = inv A'"
+  unfolding inv_def add_reset_def by (simp split: prod.splits)
+
+lemma inv_add_inv:
+  "inv (add_inv (C, U, T, I)) = (\<lambda>l. if l \<in> U then acconstraint.LE urge 0 # I l else I l)"
+  unfolding add_inv_def inv_def by simp
+
+definition
+  "is_urgent L \<equiv> \<exists>p<n_ps. L ! p \<in> urgent (N p)"
+
+lemma map_nth':
+  "map ((!) xs) [0..<n] = xs" if "n = length xs"
+  using that List.map_nth by simp
+
+lemma A_urge_split:
+  "A_urge = (broadcast, map N_urge [0..<n_ps], bounds)"
+  unfolding broadcast_def N_urge_def bounds_def n_ps_def A_urge_def by (cases A)(simp add: map_nth')
+
+lemma inv_add_inv':
+  "inv (add_inv A') =
+  (\<lambda>l. if l \<in> urgent A' then acconstraint.LE urge 0 # inv A' l else inv A' l)"
+  apply (cases A')
+  apply (simp add: inv_add_inv urgent_def)
+  unfolding inv_def
+  apply auto
+  done
+
+lemma A_split':
+  "A = (fst A, fst (snd A), snd (snd A))"
+  by auto
+
+lemma is_urgent_simp:
+  "(\<exists>p<n_ps. L ! p \<in> urgent (map N [0..<n_ps] ! p)) \<longleftrightarrow> is_urgent L"
+  unfolding is_urgent_def by auto
+
+lemma urgent_N_urge:
+  "urgent (N_urge p) = {}" if "p < n_ps"
+  using that unfolding N_def N_urge_def n_ps_def urgent_def add_reset_def add_inv_def
+  by (auto split: prod.split_asm)
+
+lemma committed_N_urge:
+  "committed (N_urge p) = committed (N p)" if "p < n_ps"
+  using that unfolding N_def N_urge_def n_ps_def committed_def add_reset_def add_inv_def
+  by (auto split: prod.split)
+
+lemma is_urgent_simp2:
+  "(\<exists>p<n_ps. L ! p \<in> urgent (map N_urge [0..<n_ps] ! p)) \<longleftrightarrow> False"
+  unfolding is_urgent_def by (auto simp: urgent_N_urge)
+
+lemma inv_add_invI1:
+  "u \<turnstile> inv (add_inv (N p)) (L ! p)" if "u \<turnstile> inv (N p) (L ! p)" "\<not> is_urgent L" "p < n_ps"
+  using that unfolding inv_add_inv' is_urgent_def by simp
+
+lemma inv_add_invI2:
+  "u \<turnstile> inv (add_inv (N p)) (L ! p)" if "u \<turnstile> inv (N p) (L ! p)" "u urge \<le> 0" "p < n_ps"
+  using that unfolding inv_add_inv' is_urgent_def by (auto elim: guard_consI[rotated])
+
+lemma inv_add_invD:
+  "u \<turnstile> inv A' l" if "u \<turnstile> inv (add_inv A') l"
+  using that unfolding inv_add_inv' by (auto split: if_split_asm simp: clock_val_def)
+
+lemma inv_N_urge:
+  "inv (N_urge p) = inv (add_inv (N p))" if "p < n_ps"
+  using that by (simp add: N_urge_simp inv_add_reset)
+
+lemma N_urge_trans_simp:
+  "trans (N_urge i) = (\<lambda>(l, b, g, a, f, r, l'). (l, b, g, a, f, urge # r, l')) ` (trans (N i))"
+  if "i < n_ps"
+  unfolding N_urge_simp[OF that] add_inv_def add_reset_def trans_def by (simp split: prod.splits)
+
+lemma trans_N_urgeD:
+  "(l, b, g, a, f, urge # r, l') \<in> trans (N_urge p)"
+  if "(l, b, g, a, f, r, l') \<in> trans (N p)" "p < n_ps"
+  using that by (force simp add: N_urge_trans_simp)
+
+lemma trans_N_urgeE:
+  assumes "(l, b, g, a, f, r, l') \<in> trans (N_urge p)" "p < n_ps"
+  obtains "(l, b, g, a, f, tl r, l') \<in> trans (N p)" "r = urge # tl r"
+  using assms by (force simp add: N_urge_trans_simp)
+
+lemma urge_not_in_inv:
+  "urge \<notin> constraint_clk ` set (inv (N p) l)" if "p < n_ps"
+  using urge_not_in_invariants that unfolding N_def inv_def n_ps_def by (fastforce dest! :nth_mem)
+
+lemma urge_not_in_guards':
+  assumes "(l, b, g, a, f, r, l') \<in> trans (N p)" "p < n_ps"
+  shows "urge \<notin> constraint_clk ` set g"
+  using urge_not_in_guards assms unfolding N_def trans_def n_ps_def by (fastforce dest! :nth_mem)
+
+lemma urge_not_in_resets':
+  assumes "(l, b, g, a, f, r, l') \<in> trans (N p)" "p < n_ps"
+  shows "urge \<notin> set r"
+  using urge_not_in_resets assms unfolding N_def trans_def n_ps_def by (fastforce dest! :nth_mem)
+
+lemma clk_upd_clock_val_a_simp:
+  "u(c := d) \<turnstile>a ac \<longleftrightarrow> u \<turnstile>a ac" if "c \<noteq> constraint_clk ac"
+  using that by (cases ac) auto
+
+lemma clk_upd_clock_val_simp:
+  "u(c := d) \<turnstile> cc \<longleftrightarrow> u \<turnstile> cc" if "c \<notin> constraint_clk ` set cc"
+  using that unfolding clock_val_def list_all_def
+  by (force simp: image_def clk_upd_clock_val_a_simp)
+
+lemma inv_N_urgeI:
+  assumes "u \<turnstile> inv (N p) l" "p < n_ps"
+  shows "u(urge := 0) \<turnstile> inv (N_urge p) l"
+  using assms urge_not_in_inv[of p]
+  by (auto simp: inv_N_urge inv_add_inv' clk_upd_clock_val_simp intro!: guard_consI)
+
+lemma inv_N_urge_updI:
+  assumes "u \<turnstile> inv (N p) l" "p < n_ps"
+  shows "u(urge := d) \<turnstile> inv (N p) l"
+  using assms urge_not_in_inv[of p] by (auto simp: clk_upd_clock_val_simp)
+
+lemma inv_N_urge_upd_simp:
+  assumes "p < n_ps"
+  shows "u(urge := d) \<turnstile> inv (N p) l \<longleftrightarrow> u \<turnstile> inv (N p) l"
+  using assms urge_not_in_inv[of p] by (auto simp: clk_upd_clock_val_simp)
+
+lemma inv_N_urge_updD:
+  assumes "u(urge := d) \<turnstile> inv (N p) l" "p < n_ps"
+  shows "u \<turnstile> inv (N p) l"
+  using assms urge_not_in_inv[of p] by (auto simp: clk_upd_clock_val_simp)
+
+lemma inv_N_urgeD:
+  assumes "u \<turnstile> inv (N_urge p) l" "p < n_ps"
+  shows "u(urge := d) \<turnstile> inv (N p) l"
+  using assms urge_not_in_inv[of p]
+  by (auto simp: inv_N_urge inv_add_inv' clk_upd_clock_val_simp split: if_split_asm elim: guard_consE)
+
+lemma inv_N_urge_urges:
+  assumes "\<forall>p<n_ps. u(urge := 0) \<oplus> d \<turnstile> inv (N_urge p) (L ! p)" "is_urgent L"
+  shows "d \<le> 0"
+proof -
+  from \<open>is_urgent L\<close> obtain p where "p < n_ps" "L ! p \<in> urgent (N p)"
+    unfolding is_urgent_def by auto
+  then have "acconstraint.LE urge 0 \<in> set (inv (N_urge p) (L ! p))"
+    by (simp add: inv_N_urge inv_add_inv')
+  with assms \<open>p < n_ps \<close> have "u(urge := 0) \<oplus> d \<turnstile>a acconstraint.LE urge 0"
+    unfolding clock_val_def list_all_iff by auto
+  then show ?thesis
+    by (auto simp: cval_add_def)
+qed
+
+lemma trans_urge_upd_iff:
+  assumes "(l, b, g, a, f, r, l') \<in> trans (N p)" "p < n_ps"
+  shows "u(urge := d) \<turnstile> g \<longleftrightarrow> u \<turnstile> g"
+  using assms urge_not_in_guards' by (auto simp: clk_upd_clock_val_simp)
+
+lemma cval_add_0[simp]:
+  "u \<oplus> (0 :: 'tt :: time) = u"
+  unfolding cval_add_def by simp
+
+lemma clock_val_reset_delay:
+  "u(c := 0) \<oplus> (d :: 'tt :: time) = (u \<oplus> d)(c := d)"
+  unfolding cval_add_def by auto
+
+lemma clock_set_upd_simp:
+  "[r \<rightarrow> d]u(c := d') = [r \<rightarrow> d]u" if "c \<in> set r"
+  using that
+  apply (induction r arbitrary: u)
+   apply auto
+  oops
+
+lemma fun_upd_twist2:
+  "f(a := x, b := x) = f(b := x, a := x)"
+  by auto
+
+lemma clock_set_upd_simp:
+  "[c # r \<rightarrow> d]u(c := d') = [c # r \<rightarrow> d]u"
+  apply (induction r arbitrary: u)
+   apply simp
+  apply simp
+  apply (subst fun_upd_twist2)
+  apply auto
+  done
+
+lemma clock_set_commute_single:
+  "[r \<rightarrow> d]u(c := d') = ([r \<rightarrow> d]u)(c := d')" if "c \<notin> set r"
+  using that by (induction r) auto
+
+lemma clock_set_upd_simp2:
+  "([xs @ c # ys \<rightarrow> d]u)(c := d') = ([xs @ ys \<rightarrow> d]u)(c := d')"
+  apply (induction xs)
+   apply (solves auto)
+  apply (intro ext)
+  subgoal for a xs x
+    by (drule fun_cong[where x = x]) auto
+  done
+
+lemma clock_set_upd_simp3:
+  "([xs \<rightarrow> d]u)(c := d') = ([filter (\<lambda>x. x \<noteq> c) xs \<rightarrow> d]u)(c := d')"
+  apply (induction xs)
+   apply (solves auto)
+  apply (rule ext)
+  apply (clarsimp, safe)
+  apply clarsimp
+  subgoal for xs x
+    by (drule fun_cong[where x = x]) auto
+  subgoal for a xs x
+    by (drule fun_cong[where x = x]) auto
+  done
+
+interpretation urge: Prod_TA_sem A_urge .
+
+lemma urge_n_ps_eq:
+  "urge.n_ps = n_ps"
+  unfolding urge.n_ps_def n_ps_def unfolding A_urge_def by simp
+
+lemma urge_N_simp[simp]:
+  "urge.N = N_urge"
+  unfolding urge.N_def N_urge_def unfolding A_urge_def by simp
+
+lemma urge_states_eq:
+  "urge.states = states"
+  unfolding states_def urge.states_def urge_n_ps_eq
+  by (auto simp add: N_urge_trans_simp split: prod.splits)
+
+interpretation Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' A L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' A_urge L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = L \<and> s' = s \<and> u' = u(urge := 0)"
+  "\<lambda>(L, s, u). L \<in> states" "\<lambda>(L, s, u). True"
+proof -
+  cancel\<open>case (1 L s u L' s' u')
+  then show ?case\<close>
+  have "A_urge \<turnstile> \<langle>L, s, u (urge := 0)\<rangle> \<rightarrow> \<langle>L', s', u'(urge := 0)\<rangle>"
+    if steps: "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>" "L \<in> states"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> 'v option"
+      and u :: "'c \<Rightarrow> 't"
+      and L' :: "'s list"
+      and s' :: "'x \<Rightarrow> 'v option"
+      and u' :: "'c \<Rightarrow> 't"
+    using that
+  proof cases
+    case steps: (1 L'' s'' u'' a)
+    have *: "(u \<oplus> d)(urge := (u \<oplus> d) urge - u urge) = u(urge := 0) \<oplus> d" for d
+      by (auto simp: cval_add_def)
+    from steps(1) \<open>L \<in> states\<close> have "L'' \<in> states"
+      by (rule states_inv)
+    then have [simp]: "length L'' = n_ps"
+      by (rule states_lengthD)
+    from steps(1) have
+      "A_urge \<turnstile> \<langle>L, s, u(urge := 0)\<rangle> \<rightarrow>Del \<langle>L'', s'', u''(urge := u'' urge - u urge)\<rangle>"
+      apply cases
+      apply (subst A_urge_split)
+      apply (subst (asm) A_split)
+      apply (drule sym)
+      apply simp
+      unfolding TAG_def[of "''urgent''"]
+      apply (simp add: is_urgent_simp *)
+      apply (rule step_u.intros)
+      apply tag
+      apply simp
+      subgoal
+        by (cases "is_urgent L")
+          (auto 4 3
+            intro: inv_N_urge_updI inv_add_invI1 inv_add_invI2
+            simp: inv_N_urge clock_val_reset_delay)
+      by (tag, simp add: is_urgent_simp2)+
+    moreover from steps(3,2) have
+      "A_urge \<turnstile> \<langle>L'', s'', u''(urge := u'' urge - u urge)\<rangle> \<rightarrow>a \<langle>L', s', u'(urge := 0)\<rangle>"
+      apply (cases; subst A_urge_split; subst (asm) A_split)
+         apply (all \<open>drule sym\<close>)
+      subgoal delay
+        by simp
+      subgoal internal premises prems[tagged] for l b g a' f r l' N p B broadcast
+        using prems apply tag usingT "''range''"
+        apply simp
+        apply (rule step_u.intros)
+        preferT \<open>TRANS ''silent''\<close>
+          apply (solves \<open>tag, simp, drule (1) trans_N_urgeD, subst nth_map, simp, simp\<close>)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''bexp''\<close>
+          apply (tag, assumption)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "''guard''" "TRANS _"; auto simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''new valuation''\<close>
+          apply (tag, unfold clock_set.simps(2)[symmetric] clock_set_upd_simp, simp; fail)
+        by (tag; auto intro: inv_N_urgeI)+
+      subgoal binary premises prems[tagged]
+        for a' broadcast' l1 b1 g1 f1 r1 l1' N' p l2 b2 g2 f2 r2 l2' q s'' B
+        using prems apply tag
+        usingT "RECV ''range''" "SEND ''range''"
+        apply simp
+        apply (rule step_u.intros)
+        preferT \<open>''not broadcast''\<close>
+          apply (tag; simp; fail)
+        preferT \<open>TRANS ''in''\<close>
+          apply (solves \<open>tag, simp, drule (1) trans_N_urgeD, subst nth_map, simp, simp\<close>)
+        preferT \<open>TRANS ''out''\<close>
+          apply (solves \<open>tag, simp, drule (1) trans_N_urgeD, subst nth_map, simp, simp\<close>)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''bexp''\<close>
+          apply (tag, assumption)
+        preferT \<open>''bexp''\<close>
+          apply (tag, assumption)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "''guard''" "TRANS _"; auto simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "''guard''" "TRANS _"; auto simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''new valuation''\<close>
+          apply (solves \<open>tag, simp, unfold clock_set.simps(2)[symmetric] clock_set_upd_simp,
+            simp add: clock_set_upd_simp2\<close>)
+        by (solves \<open>tag; auto intro: inv_N_urgeI; simp\<close>)+
+      subgoal broadcast premises prems[tagged]
+        for a' broadcast' l b g f r l' N' p ps bs gs fs rs ls' s'' B
+        using prems apply tag
+        usingT "SEL ''range''" "SEND ''range''"
+        apply (simp add: subset_nat_0_atLeastLessThan_conv)
+        apply (rule step_u.intros)
+        preferT \<open>''broadcast''\<close>
+          apply (tag; simp; fail)
+        preferT \<open>TRANS ''out''\<close>
+          apply (solves \<open>tag, simp, drule (1) trans_N_urgeD, subst nth_map, simp, simp\<close>)
+        preferT \<open>TRANS ''in''\<close>
+          apply (solves \<open>tag, auto 4 3 dest: trans_N_urgeD\<close>)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''bexp''\<close>
+          apply (tag, assumption)
+        preferT \<open>''bexp''\<close>
+          apply (tag, assumption)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "''guard''" "TRANS _"; auto simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "''guard''" "TRANS _"; auto simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''maximal''\<close>
+          apply (all \<open>tag; auto intro: inv_N_urgeI; fail | succeed\<close>)
+        preferT \<open>''maximal''\<close>
+          apply (solves \<open>tag; auto; erule (1) trans_N_urgeE, force simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''new valuation''\<close> subgoal
+          apply tag
+          apply clarsimp
+          unfolding clock_set.simps(2)[symmetric] clock_set_upd_simp
+          apply simp
+          apply (subst (2) clock_set_upd_simp3)
+          apply (subst clock_set_upd_simp3)
+          apply (simp add: filter_concat comp_def)
+          done
+        done
+      done
+    ultimately show "A_urge \<turnstile> \<langle>L, s, u(urge := 0)\<rangle> \<rightarrow> \<langle>L', s', u'(urge := 0)\<rangle>"
+      using \<open>a \<noteq> _\<close> by (intro step_u'.intros)
+  qed
+cancel\<open>next
+  case (2 L s u L' s' u')\<close>
+  moreover have "\<exists>ba. A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', ba\<rangle> \<and> u' = ba(urge := 0)"
+    if "A_urge \<turnstile> \<langle>L, s, u (urge := 0)\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+      and "L \<in> states"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> 'v option"
+      and u :: "'c \<Rightarrow> 't"
+      and L' :: "'s list"
+      and s' :: "'x \<Rightarrow> 'v option"
+      and u' :: "'c \<Rightarrow> 't"
+    using that
+  proof cases
+    case steps: (1 L'' s'' u'' a)
+    have *: "(u(urge := 0) \<oplus> d)(urge := (u(urge := 0) \<oplus> d) urge + u urge) = u \<oplus> d" for d
+      unfolding cval_add_def by auto
+    from steps(1) \<open>L \<in> states\<close> have "L'' \<in> states"
+      by (rule urge.states_inv[unfolded urge_states_eq])
+    then have [simp]: "length L'' = n_ps"
+      by (rule states_lengthD)
+    have 1: "([r\<rightarrow>0]u'') urge = 0" if "r = urge # xs" for r xs
+      by (subst that) simp
+    have 2: "filter (\<lambda>x. x \<noteq> urge) r = filter (\<lambda>x. x \<noteq> urge) (tl r)" if "r = urge # tl r" for r
+      by (subst that) simp
+    from steps(3,2) have "u' urge = 0"
+      apply (cases; subst (asm) A_urge_split)
+         apply (all \<open>drule sym\<close>)
+         apply (simp; fail)
+      subgoal delay
+        by (tag- "''new valuation''" "TRANS _" "''range''")
+           (solves \<open>auto elim!: trans_N_urgeE simp: 1\<close>)
+      subgoal binary for a' broadcast' l1 b1 g1 f1 r1 l1' N p l2 b2 g2 f2 r2 l2' q s'' B
+        by (tag- "''new valuation''" "TRANS _" "RECV ''range''")
+          (auto elim!: trans_N_urgeE simp: 1[of _ "tl r1 @ r2"])
+      subgoal broadcast for a' broadcast' l b g f r l' N p ps bs gs fs rs ls' s'a B
+        by (tag- "''new valuation''" "TRANS _"  "SEND ''range''")
+          (auto elim!: trans_N_urgeE simp: 1[of _ "tl r @ concat (map rs ps)"])
+      done
+    have **: "
+      ([r\<rightarrow>0]u'')(urge := u'' urge + u urge) = ([tl r\<rightarrow>0]u'')(urge := u'' urge + u urge)"
+      if "r = urge # tl r" for r
+      by (subst that) simp
+    from steps(1) have "A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L'', s'', u''(urge := u'' urge + u urge)\<rangle>"
+      apply cases
+      apply (subst (asm) A_urge_split)
+      apply (subst A_split)
+      apply (drule sym)
+      apply simp
+      unfolding TAG_def[of "''urgent''"]
+      apply (simp add: is_urgent_simp *)
+      apply (rule step_u.intros)
+      subgoal
+        by tag (auto dest!: inv_add_invD inv_N_urge_updD simp: inv_N_urge clock_val_reset_delay)
+        apply (tag, assumption)
+       apply (tag- "''target invariant''" "''nonnegative''")
+       apply (auto simp add: is_urgent_simp2 is_urgent_simp dest: inv_N_urge_urges)
+      done
+    moreover from steps(3,2) have
+      "A \<turnstile> \<langle>L'', s'', u''(urge := u'' urge + u urge)\<rangle> \<rightarrow>a \<langle>L', s', u'(urge := u'' urge + u urge)\<rangle>"
+      apply (cases; subst (asm) A_urge_split; subst A_split)
+         apply (all \<open>drule sym\<close>)
+      subgoal delay
+        by simp
+      subgoal internal premises prems[tagged] for l b g a' f r l' N p B broadcast
+        using prems apply del_tags
+        usingT "''range''"
+        apply simp
+        apply (rule step_u.intros)
+        preferT \<open>TRANS ''silent''\<close>
+          apply (solves \<open>tag, simp, erule (1) trans_N_urgeE, subst nth_map, simp, simp\<close>)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''bexp''\<close>
+          apply (tag; assumption)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "TRANS _"; auto intro: trans_N_urgeE simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''target invariant''\<close>
+          apply (all \<open>tag; auto intro: inv_N_urgeD; fail | succeed\<close>)
+        preferT \<open>''new valuation''\<close>
+          apply (solves \<open>tag "TRANS _", simp, erule (1) trans_N_urgeE,
+            subst clock_set_commute_single, rule urge_not_in_resets', (simp add: **)+\<close>)
+        done
+      subgoal binary premises prems[tagged]
+        for a' broadcast' l1 b1 g1 f1 r1 l1' Na p l2 b2 g2 f2 r2 l2' q s'' B
+        using prems apply del_tags
+        usingT "RECV ''range''" "SEND ''range''"
+        apply simp
+        apply (rule step_u.intros)
+        preferT \<open>''not broadcast''\<close> apply (tag; simp; fail)
+        preferT \<open>TRANS ''in''\<close>
+          apply (solves \<open>tag, simp, erule (1) trans_N_urgeE, subst nth_map, simp, simp\<close>)
+        preferT \<open>TRANS ''out''\<close>
+          apply (solves \<open>tag, simp, erule (1) trans_N_urgeE, subst nth_map, simp, simp\<close>)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''bexp''\<close>
+          apply (tag; assumption)
+        preferT \<open>''bexp''\<close>
+          apply (tag; assumption)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "TRANS _"; auto intro: trans_N_urgeE simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "TRANS _"; auto intro: trans_N_urgeE simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''new valuation''\<close> subgoal premises prems[tagged]
+          using prems apply tag
+          apply (tag "TRANS _", simp, erule (1) trans_N_urgeE, erule (1) trans_N_urgeE)
+          apply (subst clock_set_upd_simp3)
+          apply (subst clock_set_commute_single)
+           apply (simp; rule conjI; erule (1) urge_not_in_resets'; fail)
+          apply (rule arg_cong)
+          subgoal premises prems
+            using urge_not_in_resets'[OF prems(7) \<open>p < n_ps\<close>]
+                  urge_not_in_resets'[OF prems(9) \<open>q < n_ps\<close>]
+            by (subst \<open>r1 = _\<close>, subst \<open>r2 = _\<close>, simp) (subst filter_True, auto)+
+          done
+        by (solves \<open>tag; auto intro: inv_N_urgeD\<close>)+
+      subgoal broadcast premises prems[tagged]
+        for a' broadcast' l b g f r l' N' p ps bs gs fs rs ls' s2 B
+        using prems apply del_tags
+        usingT "SEL ''range''" "SEND ''range''"
+        apply (simp add: subset_nat_0_atLeastLessThan_conv)
+        apply (rule step_u.intros)
+        preferT \<open>TRANS ''out''\<close>
+          \<comment> \<open>instantiates schematics\<close>
+          apply (solves \<open>tag, simp, erule (1) trans_N_urgeE, subst nth_map, simp, simp\<close>)
+        preferT \<open>TRANS ''in''\<close>
+          apply (tag, auto 4 3 elim: trans_N_urgeE; fail) \<comment> \<open>instantiates schematics\<close>
+        preferT \<open>''broadcast''\<close>
+          apply (tag; simp; fail)
+        preferT \<open>''committed''\<close>
+          apply (solves \<open>tag, subst nth_map, simp, simp add: committed_N_urge\<close>)
+        preferT \<open>''guard''\<close>
+          apply (solves \<open>tag "TRANS _"; auto elim: trans_N_urgeE simp: trans_urge_upd_iff\<close>)
+        preferT \<open>''guard''\<close> subgoal guards
+          by (tag "TRANS _"; auto elim!: trans_N_urgeE simp: trans_urge_upd_iff)
+        preferT \<open>''maximal''\<close> subgoal maximal
+          by (tag, force simp: trans_urge_upd_iff dest: trans_N_urgeD)
+        preferT \<open>''new valuation''\<close> subgoal new_valuation premises prems[tagged]
+          apply (tag "TRANS _")
+          using prems apply tag
+          apply (subst [[get_tagged \<open>''new valuation''\<close>]])
+          apply (subst clock_set_upd_simp3)
+          apply (simp add: filter_concat comp_def)
+          apply (subst clock_set_commute_single[symmetric], (simp; fail))
+          apply (rule arg_cong)
+          apply (fo_rule arg_cong2)
+          subgoal
+            by (auto simp: 2 urge_not_in_resets' filter_id_conv elim!: trans_N_urgeE)
+          subgoal
+            apply (fo_rule arg_cong)
+            apply (fastforce elim: trans_N_urgeE simp: 2 urge_not_in_resets' filter_id_conv)+
+            done
+          done
+        by (solves \<open>tag; auto intro: inv_N_urgeD\<close>)+
+      done
+    ultimately show ?thesis
+      using \<open>a \<noteq> _\<close> \<open>u' urge = 0\<close> by  (intros add: step_u'.intros) auto
+  qed
+  moreover have "ab \<in> states"
+    if "a \<in> states"
+      and "A \<turnstile> \<langle>a, aa, b\<rangle> \<rightarrow> \<langle>ab, ac, ba\<rangle>"
+    for a :: "'s list"
+      and aa :: "'x \<Rightarrow> 'v option"
+      and b :: "'c \<Rightarrow> 't"
+      and ab :: "'s list"
+      and ac :: "'x \<Rightarrow> 'v option"
+      and ba :: "'c \<Rightarrow> 't"
+    using that by (elim step_u'_elims states_inv)
+  ultimately show
+    "Bisimulation_Invariant
+     (\<lambda>(L, s, u) (L', s', u').
+         A \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u').
+         A_urge \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u').
+         L' = L \<and> s' = s \<and> u' = u(urge := 0))
+     (\<lambda>(L, s, u). L \<in> states) (\<lambda>(L, s, u). True)"
+    by - (standard; clarsimp)
+qed
+
+lemmas urge_bisim = Bisimulation_Invariant_axioms
+
+end (* Prod TA sem urge *)
+
+
+context Simple_Network_Impl_Defs
+begin
+
+lemma dom_bounds: "dom bounds = fst ` set bounds'"
+  unfolding bounds_def by (simp add: dom_map_of_conv_image_fst)
+
+lemma mem_trans_N_iff:
+  "t \<in> Simple_Network_Language.trans (N i) \<longleftrightarrow> t \<in> set (fst (snd (snd (automata ! i))))"
+  if "i < n_ps"
+  unfolding N_eq[OF that] by (auto split: prod.splits simp: automaton_of_def trans_def)
+
+lemma length_automata_eq_n_ps:
+  "length automata = n_ps"
+  unfolding n_ps_def by simp
+
+lemma N_p_trans_eq:
+  "Simple_Network_Language.trans (N p) = set (fst (snd (snd (automata ! p))))" if "p < n_ps"
+  using mem_trans_N_iff[OF that] by auto
+
+lemma loc_set_compute:
+  "loc_set = \<Union> ((\<lambda>(_, _, t, _). \<Union> ((\<lambda>(l, _, _, _, _, _, l'). {l, l'}) ` set t)) ` set automata)"
+  unfolding loc_set_def setcompr_eq_image
+  apply (safe; clarsimp simp: N_p_trans_eq n_ps_def)
+     apply (drule nth_mem, erule bexI[rotated], force)+
+   apply (drule mem_nth, force)+
+  done
+
+lemma var_set_compute:
+  "var_set =
+  (\<Union>S \<in> (\<lambda>t. (fst \<circ> snd) ` set t) ` ((\<lambda>(_, _, t, _). t) `  set automata). \<Union>b \<in> S. vars_of_bexp b) \<union>
+  (\<Union>S \<in> (\<lambda>t. (fst \<circ> snd o snd \<circ> snd \<circ> snd) ` set t) ` ((\<lambda>(_, _, t, _). t) `  set automata).
+    \<Union>f \<in> S. \<Union> (x, e) \<in> set f. {x} \<union> vars_of_exp e)"
+  unfolding var_set_def setcompr_eq_image
+  by (rule arg_cong2[where f = "(\<union>)"]; auto simp: N_p_trans_eq n_ps_def,
+     (drule nth_mem, erule bexI[rotated],
+      metis (no_types, lifting) insert_iff mem_case_prodI prod.collapse)+,
+      (drule mem_nth, force)+)
+
+lemma states_loc_setD:
+  "set L \<subseteq> loc_set" if "L \<in> states"
+  using states_loc_set that by auto
+
+end (* Simple Network Impl Defs *)
+
+
+context Simple_Network_Impl
+begin
+
+lemma sem_bounds_eq: "sem.bounds = bounds"
+  unfolding sem.bounds_def bounds_def unfolding sem_def by simp
+
+lemma n_ps_eq[simp]:
+  "sem.n_ps = n_ps"
+  unfolding n_ps_def sem.n_ps_def unfolding sem_def by auto
+
+lemma sem_loc_set_eq:
+  "sem.loc_set = loc_set"
+  unfolding sem.loc_set_def loc_set_def n_ps_eq setcompr_eq_image
+  apply (simp add: sem_N_eq N_eq)
+  apply (rule arg_cong2[where f = "(\<union>)"])
+  apply (safe; clarsimp;
+      force split: prod.splits simp:  conv_automaton_def trans_def automaton_of_def n_ps_def)+
+  done
+
+lemma sem_states_eq:
+  "sem.states = states"
+  unfolding sem.states_def states_def n_ps_eq setcompr_eq_image
+  by (clarsimp simp: sem_N_eq N_eq;
+      force simp:  conv_automaton_def trans_def automaton_of_def n_ps_def split: prod.splits)+
+
+end (* Simple Network Impl *)
+
+
+locale Simple_Network_Rename_Defs =
+  Simple_Network_Impl_Defs automata broadcast bounds' for automata ::
+    "('s list \<times> 's list \<times> (('a :: countable) act, 's, 'c, 't, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, 't) cconstraint) list) list"
+    and broadcast bounds' +
+  fixes renum_acts   :: "'a \<Rightarrow> nat"
+    and renum_vars   :: "'x \<Rightarrow> nat"
+    and renum_clocks :: "'c \<Rightarrow> nat"
+    and renum_states :: "nat \<Rightarrow> 's \<Rightarrow> nat"
+begin
+
+definition
+  "map_cconstraint f g xs = map (map_acconstraint f g) xs"
+
+definition
+  "renum_cconstraint = map_cconstraint renum_clocks id"
+
+definition
+  "renum_act = map_act renum_acts"
+
+definition
+  "renum_bexp = map_bexp renum_vars"
+
+definition
+  "renum_exp = map_exp renum_vars"
+
+definition
+  "renum_upd = (\<lambda>(x, upd). (renum_vars x, renum_exp upd))"
+
+abbreviation
+  "renum_upds \<equiv> map renum_upd"
+
+definition
+  "renum_reset = map renum_clocks"
+
+definition renum_automaton where
+  "renum_automaton i \<equiv> \<lambda>(committed, urgent, trans, inv). let
+    committed' = map (renum_states i) committed;
+    urgent' = map (renum_states i) urgent;
+    trans' = map (\<lambda>(l, b, g, a, upd, r, l').
+      (renum_states i l, renum_bexp b, renum_cconstraint g, renum_act a, renum_upds upd,
+       renum_reset r,  renum_states i l')
+    ) trans;
+    inv' = map (\<lambda>(l, g). (renum_states i l, renum_cconstraint g)) inv
+  in (committed', urgent', trans', inv')
+"
+
+definition
+  "vars_inv \<equiv> the_inv renum_vars"
+
+definition
+  "map_st \<equiv> \<lambda>(L, s). (map_index renum_states L, s o vars_inv)"
+
+definition
+  "clocks_inv \<equiv> the_inv renum_clocks"
+
+definition
+  "map_u u = u o clocks_inv"
+
+lemma map_u_add[simp]:
+  "map_u (u \<oplus> d) = map_u u \<oplus> d"
+  by (auto simp: map_u_def cval_add_def)
+
+definition renum_label where
+  "renum_label = map_label renum_acts"
+
+sublocale renum: Simple_Network_Impl_Defs
+  "map_index renum_automaton automata"
+  "map renum_acts broadcast"
+  "map (\<lambda>(a,p). (renum_vars a, p)) bounds'"
+  .
+
+lemma renum_n_ps_simp[simp]:
+  "renum.n_ps = n_ps"
+  unfolding n_ps_def renum.n_ps_def by simp
+
+end (* Simple Network Rename Defs *)
+
+
+locale Simple_Network_Rename_Defs_int =
+  Simple_Network_Rename_Defs automata +
+  Simple_Network_Impl automata
+  for automata ::
+    "('s list \<times> 's list \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list"
+begin
+
+sublocale renum: Simple_Network_Impl
+  "map_index renum_automaton automata"
+  "map renum_acts broadcast"
+  "map (\<lambda>(a,p). (renum_vars a, p)) bounds'"
+  .
+
+end (* Simple Network Rename Defs int *)
+
+lemma
+  fixes f :: "'b :: countable \<Rightarrow> nat"
+  assumes "inj_on f S" "finite S" "infinite (UNIV :: 'b set)"
+  shows extend_bij_inj: "inj (extend_bij f S)" and extend_bij_surj: "surj (extend_bij f S)"
+    and extend_bij_extends: "\<forall>x \<in> S. extend_bij f S x = f x"
+proof -
+  have "infinite (\<nat> :: nat set)"
+    unfolding Nats_def by simp
+  from assms this have "bij (extend_bij f S)"
+    by (intro extend_bij_bij) auto
+  then show "inj (extend_bij f S)" and "surj (extend_bij f S)"
+    unfolding bij_def by fast+
+  from assms \<open>infinite \<nat>\<close> show "\<forall>x \<in> S. extend_bij f S x = f x"
+    by (intro extend_bij_extends) auto
+qed
+
+lemma default_map_of_map:
+  "default_map_of y (map (\<lambda>(a, b). (f a, g b)) xs) (f a) = g (default_map_of x xs a)"
+  if "inj f" "y = g x"
+  using that unfolding default_map_of_alt_def
+  by (auto 4 4 dest: injD[OF that(1)] map_of_SomeD simp: map_of_eq_None_iff map_of_mapk_SomeI)
+
+lemma default_map_of_map_2:
+  "default_map_of y (map (\<lambda>(a, b). (a, g b)) xs) a = g (default_map_of x xs a)" if "y = g x"
+  unfolding default_map_of_alt_def using that by (auto simp: map_of_map)
+
+lemma map_of_map':
+  "map_of (map (\<lambda>(k, v). (k, f k v)) xs)
+  = (\<lambda>k. case map_of xs k of Some v \<Rightarrow> Some (f k v) | _ \<Rightarrow> None)"
+  by (induct xs) (auto simp: fun_eq_iff)
+
+lemma default_map_of_map_3:
+  "default_map_of y (map (\<lambda>(a, b). (a, g a b)) xs) a = g a (default_map_of x xs a)"
+  if "\<And>k. y = g k x"
+  unfolding default_map_of_alt_def using that by (auto simp: map_of_map')
+
+lemma dom_map_of_map:
+  "dom (map_of (map (\<lambda> (a, b). (f a, g b)) xs)) = f ` fst ` set xs"
+  unfolding dom_map_of_conv_image_fst by (auto 4 3)
+
+lemma inj_image_eqI:
+  "S = T" if "inj f" "f ` S = f ` T"
+  using that by (meson inj_image_eq_iff)
+
+lemmas [finite_intros] = finite_set
+
+lemma map_of_NoneI:
+  "map_of xs x = None" if "x \<notin> fst ` set xs"
+  by (simp add: map_of_eq_None_iff that)
+
+lemma bij_f_the_inv_f:
+  "f (the_inv f x) = x" if "bij f"
+  using that f_the_inv_f unfolding bij_def by (simp add: f_the_inv_f)
+
+
+fun map_sexp ::
+  "(nat \<Rightarrow> 's \<Rightarrow> 's1) \<Rightarrow> ('a \<Rightarrow> 'a1) \<Rightarrow> ('b \<Rightarrow> 'b1) \<Rightarrow> (nat, 's, 'a, 'b) sexp
+    \<Rightarrow> (nat, 's1, 'a1, 'b1) sexp"
+  where
+  "map_sexp _ _ _ sexp.true = sexp.true" |
+  "map_sexp f g h (not e) = not (map_sexp f g h e)" |
+  "map_sexp f g h (and e1 e2) = and (map_sexp f g h e1) (map_sexp f g h e2)" |
+  "map_sexp f g h (sexp.or e1 e2) = sexp.or (map_sexp f g h e1) (map_sexp f g h e2)" |
+  "map_sexp f g h (imply e1 e2) = imply (map_sexp f g h e1) (map_sexp f g h e2)" |
+  "map_sexp f g h (eq i x) = eq (g i) (h x)" |
+  "map_sexp f g h (lt i x) = lt (g i) (h x)" |
+  "map_sexp f g h (le i x) = le (g i) (h x)" |
+  "map_sexp f g h (ge i x) = ge (g i) (h x)" |
+  "map_sexp f g h (gt i x) = gt (g i) (h x)" |
+  "map_sexp f g h (loc i x) = loc i (f i x)"
+
+fun map_formula ::
+  "(nat \<Rightarrow> 's \<Rightarrow> 's1) \<Rightarrow> ('a \<Rightarrow> 'a1) \<Rightarrow> ('b \<Rightarrow> 'b1)
+  \<Rightarrow> (nat, 's, 'a, 'b) formula \<Rightarrow> (nat, 's1, 'a1, 'b1) formula"
+where
+  "map_formula f g h (formula.EX \<phi>) = formula.EX (map_sexp f g h \<phi>)" |
+  "map_formula f g h (EG \<phi>) = EG (map_sexp f g h \<phi>)" |
+  "map_formula f g h (AX \<phi>) = AX (map_sexp f g h \<phi>)" |
+  "map_formula f g h (AG \<phi>) = AG (map_sexp f g h \<phi>)" |
+  "map_formula f g h (Leadsto \<phi> \<psi>) = Leadsto (map_sexp f g h \<phi>) (map_sexp f g h \<psi>)"
+
+
+locale Simple_Network_Impl_real =
+  fixes automata ::
+    "('s list \<times> 's list \<times> ('a act, 's, 'c, real, 'x, int) transition list
+      \<times> ('s \<times> ('c, real) cconstraint) list) list"
+    and broadcast :: "'a list"
+    and bounds' :: "('x \<times> (int \<times> int)) list"
+begin
+
+sublocale Simple_Network_Impl_Defs automata broadcast bounds' .
+
+abbreviation "sem \<equiv> (set broadcast, map automaton_of automata, map_of bounds')"
+
+sublocale Prod_TA_sem "(set broadcast, map automaton_of automata, map_of bounds')" .
+
+end
+
+context Simple_Network_Impl
+begin
+
+lemma sem_state_guard_eq:
+  "(fst \<circ> snd) ` trans (sem.N p) = (fst \<circ> snd) ` trans (N p)" if "p < n_ps"
+  unfolding sem_N_eq[OF \<open>p < n_ps\<close>] N_eq[OF \<open>p < n_ps\<close>]
+  unfolding automaton_of_def conv_automaton_def trans_def
+  by (force split: prod.splits)
+
+lemma sem_state_update_eq:
+  "(fst \<circ> snd \<circ> snd \<circ> snd \<circ> snd) ` trans (sem.N p) = (fst \<circ> snd \<circ> snd \<circ> snd \<circ> snd) ` trans (N p)"
+  if "p < n_ps"
+  unfolding sem_N_eq[OF \<open>p < n_ps\<close>] N_eq[OF \<open>p < n_ps\<close>]
+  unfolding automaton_of_def conv_automaton_def trans_def
+  by (force split: prod.splits)
+
+lemma sem_var_set_eq:
+  "sem.var_set = var_set"
+  unfolding sem.var_set_def var_set_def n_ps_eq using sem_state_guard_eq sem_state_update_eq
+  by (simp cong: image_cong_simp add: setcompr_eq_image)
+
+end
+
+locale Simple_Network_Rename_Defs_real =
+  Simple_Network_Rename_Defs automata +
+  Simple_Network_Impl_real automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, real, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, real) cconstraint) list) list"
+begin
+
+sublocale renum: Simple_Network_Impl_real
+  "map_index renum_automaton automata"
+  "map renum_acts broadcast"
+  "map (\<lambda>(a,p). (renum_vars a, p)) bounds'"
+  .
+
+end (* Simple Network Rename Defs *)
+
+locale Simple_Network_Rename' =
+  Simple_Network_Rename_Defs where automata = automata for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, 't, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, 't) cconstraint) list) list" +
+  assumes bij_renum_clocks: "bij renum_clocks"
+      and renum_states_inj: "\<forall>p<n_ps. inj (renum_states p)"
+      and bij_renum_vars: "bij renum_vars"
+      and bounds'_var_set: "fst ` set bounds' = var_set"
+      and inj_renum_acts: "inj renum_acts"
+
+locale Simple_Network_Rename_real =
+  Simple_Network_Rename_Defs_real where automata = automata +
+  Simple_Network_Rename' where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, real, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, real) cconstraint) list) list" +
+  assumes urgency_removed: "\<forall>i < n_ps. urgent (N i) = {}"
+begin
+
+lemma aux_1:
+  "(\<lambda>x. case case x of
+          (l, g) \<Rightarrow> (renum_states p l, renum_cconstraint g) of
+     (s, cc) \<Rightarrow> (s, map conv_ac cc))
+  = (\<lambda> (l, g). (renum_states p l, map conv_ac (renum_cconstraint g)))
+"
+  by auto
+
+lemma clocks_inv_inv:
+  "clocks_inv (renum_clocks a) = a"
+  unfolding clocks_inv_def by (subst the_inv_f_f; rule bij_renum_clocks[THEN bij_is_inj] HOL.refl)
+
+lemma map_u_renum_cconstraint_clock_valI:
+  "map_u u \<turnstile> renum_cconstraint cc" if "u \<turnstile> cc"
+  using that
+  unfolding clock_val_def list_all_length
+  unfolding renum_cconstraint_def map_cconstraint_def
+  unfolding map_u_def
+  apply clarsimp
+  subgoal for n
+    by (cases "cc ! n") (auto 4 4 simp: clocks_inv_inv split: acconstraint.split)
+  done
+
+lemma map_u_renum_cconstraint_clock_valD:
+  "u \<turnstile> cc" if "map_u u \<turnstile> renum_cconstraint cc"
+  using that
+  unfolding clock_val_def list_all_length
+  unfolding renum_cconstraint_def map_cconstraint_def
+  unfolding map_u_def
+  apply clarsimp
+  subgoal for n
+    by (cases "cc ! n") (auto 4 4 simp: clocks_inv_inv split: acconstraint.split)
+  done
+
+lemma inj_renum_states: "inj (renum_states p)" if "p < n_ps"
+  using renum_states_inj \<open>p < n_ps\<close> by blast
+
+lemma inv_renum_sem_I:
+  assumes
+    "u \<turnstile> Simple_Network_Language.inv (N p) l" "p < n_ps" "l \<in> loc_set"
+  shows
+    "map_u u \<turnstile> Simple_Network_Language.inv (renum.N p) (renum_states p l)"
+  using assms
+  unfolding inv_def
+  apply -
+  apply (subst (asm) N_eq, assumption)
+  apply (subst renum.N_eq, subst renum_n_ps_simp, assumption)
+  apply (subst nth_map_index, simp add: n_ps_def)
+  unfolding conv_automaton_def automaton_of_def
+  apply (clarsimp split: prod.split_asm simp: renum_automaton_def comp_def)
+  unfolding aux_1
+  apply (subst default_map_of_map[where x = "[]"])
+  subgoal
+    by (intro inj_renum_states \<open>p < n_ps\<close>)
+   apply (simp add: renum_cconstraint_def map_cconstraint_def; fail)
+  apply (erule map_u_renum_cconstraint_clock_valI)
+  done
+
+lemma inv_renum_sem_D:
+  assumes
+    "map_u u \<turnstile> Simple_Network_Language.inv (renum.N p) (renum_states p l)" "p < n_ps" "l \<in> loc_set"
+  shows
+    "u \<turnstile> Simple_Network_Language.inv (N p) l"
+  using assms
+  unfolding inv_def
+  apply -
+  apply (subst N_eq, assumption)
+  apply (subst (asm) renum.N_eq, subst renum_n_ps_simp, assumption)
+  apply (subst (asm) nth_map_index, simp add: n_ps_def)
+  unfolding conv_automaton_def automaton_of_def
+  apply (clarsimp split: prod.split simp: renum_automaton_def comp_def)
+  unfolding aux_1
+  apply (subst (asm) default_map_of_map[where x = "[]"])
+  subgoal
+    by (intro inj_renum_states \<open>p < n_ps\<close>)
+   apply (simp add: renum_cconstraint_def map_cconstraint_def; fail)
+  apply (erule map_u_renum_cconstraint_clock_valD)
+  done
+
+lemma dom_the_inv_comp:
+  "dom (m o the_inv f) = f ` dom m" if "inj f" "range f = UNIV"
+  unfolding dom_def
+  apply (clarsimp, safe)
+  subgoal for x y
+    apply (subgoal_tac "f (the_inv f x) = x")
+     apply force
+    using that by (auto intro: f_the_inv_f)
+  subgoal for x y
+    using that by (auto simp: the_inv_f_f)
+  done
+
+lemma inj_renum_vars:
+  "inj renum_vars"
+  using bij_renum_vars unfolding bij_def ..
+
+lemma surj_renum_vars:
+  "surj renum_vars"
+  using bij_renum_vars unfolding bij_def ..
+
+lemma map_of_inv_map:
+  "map_of xs (the_inv f x) = map_of (map (\<lambda> (a, b). (f a, b)) xs) x"
+  if "inj f" "surj f" "the_inv f x \<in> dom (map_of xs)"
+  apply (subgoal_tac "x = f (the_inv f x)")
+  subgoal premises prems
+    using domD[OF that(3)] \<open>inj f\<close>
+    apply (subst (2) prems)
+    apply safe
+    apply (subst map_of_mapk_SomeI; assumption)
+    done
+  subgoal
+    apply (rule sym, rule f_the_inv_f)
+    using that by auto
+  done
+
+lemma dom_vars_invD:
+  assumes "x \<in> dom (s \<circ> vars_inv)"
+  shows "x \<in> renum_vars ` dom s" (is ?A) and "the_inv renum_vars x \<in> dom s" (is ?B)
+proof -
+  show ?A
+    using assms unfolding vars_inv_def dom_the_inv_comp[OF inj_renum_vars surj_renum_vars] .
+  then show ?B
+    apply (elim imageE)
+    apply simp
+    apply (subst the_inv_f_f, rule inj_renum_vars, assumption)
+    done
+qed
+
+lemma bounded_renumI:
+  assumes "bounded (map_of bounds') s"
+  shows "bounded (map_of (map (\<lambda>(a, y). (renum_vars a, y)) bounds')) (s o vars_inv)"
+  using assms unfolding bounded_def
+  apply elims
+  apply intros
+  subgoal
+    unfolding dom_map_of_conv_image_fst vars_inv_def
+    by (auto intro!: image_cong simp: dom_the_inv_comp[OF inj_renum_vars surj_renum_vars] image_comp)
+  subgoal for x
+    apply (frule dom_vars_invD)
+    apply (frule dom_vars_invD(2))
+    apply (drule bspec, assumption)
+    apply (auto simp: map_of_inv_map[OF inj_renum_vars surj_renum_vars] vars_inv_def)
+    done
+  subgoal for x
+    apply (frule dom_vars_invD)
+    apply (frule dom_vars_invD(2))
+    apply (drule bspec, assumption)
+    apply (auto simp: map_of_inv_map[OF inj_renum_vars surj_renum_vars] vars_inv_def)
+    done
+  done
+
+lemma map_of_renum_vars_simp:
+  assumes
+    "dom (s o vars_inv)
+     = dom (map_of (map (\<lambda>(a, y). (renum_vars a, y)) bounds'))"
+    "x \<in> dom s" "dom s \<subseteq> var_set"
+  shows "map_of (map (\<lambda>(a, y). (renum_vars a, y)) bounds') (renum_vars x) = map_of bounds' x"
+proof -
+  have *:
+    "map (\<lambda>(a, y). (renum_vars a, y)) bounds' = map (\<lambda>(a, y). (renum_vars a, y)) bounds'"
+    by auto
+  have "x \<in> dom (map_of bounds')"
+    unfolding dom_map_of_conv_image_fst
+    using assms
+    unfolding vars_inv_def
+    apply -
+    apply (subst (asm) dom_the_inv_comp, rule inj_renum_vars, rule surj_renum_vars)
+    apply (subst (asm) dom_map_of_map)
+    apply clarsimp
+    apply (subst (asm) inj_on_image_eq_iff[OF inj_renum_vars])
+    by auto
+  then obtain y where "map_of bounds' x = Some y"
+    by auto
+  then show ?thesis
+    by (subst map_of_mapk_SomeI) (auto intro: inj_renum_vars)
+qed
+
+lemma bounded_renumD:
+  assumes
+    "Simple_Network_Language.bounded
+     (map_of (map (\<lambda>(a, y). (renum_vars a, y)) bounds')) (s o vars_inv)"
+    and "dom s \<subseteq> var_set"
+  shows "Simple_Network_Language.bounded (map_of bounds') s"
+proof -
+  show ?thesis
+    using assms(1)
+    unfolding bounded_def
+    apply elims
+    apply intros
+    subgoal
+      unfolding vars_inv_def
+      apply (subst (asm) dom_the_inv_comp[OF inj_renum_vars surj_renum_vars])
+      apply (subst (asm) dom_map_of_map)
+      apply (rule inj_image_eqI[OF inj_renum_vars], simp add: dom_map_of_conv_image_fst)
+      done
+    subgoal for x
+      using assms(2) unfolding vars_inv_def
+      apply (subst (asm) (2) dom_the_inv_comp[OF inj_renum_vars surj_renum_vars])
+      apply (drule bspec, erule imageI)
+      apply (simp add: the_inv_f_f[OF inj_renum_vars] map_of_renum_vars_simp[unfolded vars_inv_def])
+      done
+    subgoal for x
+      using assms(2) unfolding vars_inv_def
+      apply (subst (asm) (2) dom_the_inv_comp[OF inj_renum_vars surj_renum_vars])
+      apply (drule bspec, erule imageI)
+      apply (simp add: the_inv_f_f[OF inj_renum_vars] map_of_renum_vars_simp[unfolded vars_inv_def])
+      done
+    done
+qed
+
+lemma dom_bounds_var_set: "dom bounds = var_set"
+  unfolding dom_bounds using bounds'_var_set .
+
+lemma sem_states_loc_setD: "L ! p \<in> loc_set" if "p < length automata" "L \<in> states" for L p
+  using that states_loc_set by (force simp: n_ps_def)
+
+lemma trans_N_renumD:
+  assumes "(l, b, g, a, f, r, l') \<in> Simple_Network_Language.trans (N p)" "p < n_ps"
+  shows "(renum_states p l, renum_bexp b, renum_cconstraint g, renum_act a, renum_upds f, renum_reset r, renum_states p l')
+  \<in> Simple_Network_Language.trans (renum.N p)"
+  using assms
+  unfolding mem_trans_N_iff[OF assms(2)] renum.mem_trans_N_iff[unfolded renum_n_ps_simp, OF assms(2)]
+  by (force split: prod.split simp: renum_automaton_def n_ps_def)
+
+lemma match_assumption2:
+  assumes "P a1 b1" "a1 = a" "b1 = b" shows "P a b"
+  using assms by auto
+
+lemma inj_pair:
+  assumes "inj f" "inj g"
+  shows "inj (\<lambda>(a, b). (f a, g b))"
+  using assms unfolding inj_on_def by auto
+
+lemma injective_functions:
+  "inj renum_reset" "inj renum_upd" "inj renum_act" "inj renum_cconstraint" "inj renum_bexp"
+  "\<And>p. p < length automata \<Longrightarrow> inj (renum_states p)"
+  subgoal
+    unfolding renum_reset_def using bij_renum_clocks[THEN bij_is_inj] by simp
+  subgoal
+    unfolding renum_upd_def renum_exp_def using bij_renum_vars[THEN bij_is_inj]
+    by (intro inj_pair exp.inj_map inj_mapI)
+  subgoal
+    unfolding renum_act_def using inj_renum_acts by (rule act.inj_map)
+  subgoal
+    unfolding renum_cconstraint_def map_cconstraint_def
+    by (intro inj_mapI acconstraint.inj_map bij_renum_clocks bij_is_inj bij_id)
+  subgoal
+    unfolding renum_bexp_def by (intro bexp.inj_map inj_renum_vars)
+  subgoal
+    by (rule inj_renum_states, simp add: n_ps_def)
+  done
+
+lemma trans_N_renumI:
+  assumes "(
+    renum_states p l, renum_bexp b, renum_cconstraint g, renum_act a, renum_upds f, renum_reset r,
+    renum_states p l')
+  \<in> trans (renum.N p)"  "p < n_ps"
+  shows "(l, b, g, a, f, r, l') \<in> trans (N p)"
+  using assms
+  unfolding mem_trans_N_iff[OF assms(2)] renum.mem_trans_N_iff[unfolded renum_n_ps_simp, OF assms(2)]
+  apply (clarsimp split: prod.split_asm simp: renum_automaton_def n_ps_def)
+  apply (fo_rule match_assumption2, assumption)
+   apply (auto elim!: injD[rotated, THEN sym] intro: injective_functions)
+  done
+
+lemma renum_acconstraint_eq_convD:
+  assumes "map_acconstraint renum_clocks id g = conv_ac g'"
+  obtains g1 where "g = conv_ac g1" "g' = map_acconstraint renum_clocks id g1"
+  subgoal premises
+    using assms
+    apply (cases g'; cases g; clarsimp)
+    subgoal for m c
+      by (rule that[of "acconstraint.LT c m"]; simp)
+    subgoal for m c
+      by (rule that[of "acconstraint.LE c m"]; simp)
+    subgoal for m c
+      by (rule that[of "acconstraint.EQ c m"]; simp)
+    subgoal for m c
+      by (rule that[of "acconstraint.GT c m"]; simp)
+    subgoal for m c
+      by (rule that[of "acconstraint.GE c m"]; simp)
+    done
+  done
+
+lemma renum_cconstraint_eq_convD:
+  assumes "renum_cconstraint g = conv_cc g'"
+  obtains g1 where "g = conv_cc g1" "g' = renum_cconstraint g1"
+proof -
+  let ?f = "\<lambda>(ac, ac'). SOME ac1. ac = conv_ac ac1 \<and> ac' = map_acconstraint renum_clocks id ac1"
+  let ?g = "map ?f (zip g g')"
+  from assms have "length g = length g'"
+    unfolding renum_cconstraint_def map_cconstraint_def by -(drule arg_cong[where f = length], simp)
+  then have "g = conv_cc ?g \<and> g' = renum_cconstraint ?g"
+    using assms
+    by (simp add: comp_def renum_cconstraint_def map_cconstraint_def)
+       (induction rule: list_induct2; simp; elim conjE renum_acconstraint_eq_convD; smt someI)
+  then show ?thesis
+    by (blast intro: that)
+qed
+
+lemma trans_sem_N_renumI:
+  assumes "(
+  renum_states p l, renum_bexp b, renum_cconstraint g, renum_act a, renum_upds f, renum_reset r,
+  renum_states p l')
+  \<in> trans (renum.N p)" "p < n_ps"
+  shows "(l, b, g, a, f, r, l') \<in> trans (N p)"
+  using assms(1) \<open>p < n_ps\<close> by (simp add: trans_N_renumI)
+
+lemma trans_sem_N_renumI':
+  assumes "(renum_states p l, b, g, a, f, r, l') \<in> trans (renum.N p)" "p < n_ps"
+  shows "\<exists> b' g' a' f' r' l1.
+    TRANS ''orig'' \<bbar> (l, b', g', a', f', r', l1) \<in> Simple_Network_Language.trans (N p)
+    \<and> ''renum bexp''   \<bbar> b = renum_bexp b'  \<and> ''renum guard'' \<bbar> g = renum_cconstraint g'
+    \<and> ''renum action'' \<bbar> a = renum_act a'   \<and> ''renum upds''  \<bbar> f = renum_upds f'
+    \<and> ''renum reset''  \<bbar> r = renum_reset r' \<and> ''renum loc''   \<bbar> l' = renum_states p l1"
+proof -
+  obtain b' g' a' f' r' l1 where "b = renum_bexp b'" "g = renum_cconstraint g'" "f = renum_upds f'"
+    "a = renum_act a'" "r = renum_reset r'" "l' = renum_states p l1"
+    using assms
+    unfolding N_eq[OF assms(2)] renum.N_eq[unfolded renum_n_ps_simp, OF assms(2)]
+    apply (subst (asm) nth_map_index)
+    subgoal
+      by (simp add: n_ps_def)
+    unfolding renum_automaton_def automaton_of_def trans_def by (auto split: prod.split_asm)
+  with assms show ?thesis
+    by untag (fastforce dest!: trans_sem_N_renumI)
+qed
+
+lemma committed_renum_eq:
+  "committed (renum.N p) = renum_states p ` committed (N p)" if "p < n_ps"
+  unfolding
+    committed_def N_eq[OF \<open>p < n_ps\<close>] renum.N_eq[unfolded renum_n_ps_simp, OF \<open>p < n_ps\<close>]
+  apply (subst nth_map_index)
+  subgoal
+    using \<open>p < n_ps\<close> by (simp add: n_ps_def)
+  unfolding automaton_of_def conv_automaton_def renum_automaton_def by (simp split: prod.split)
+
+lemma urgent_renum_eq:
+  "urgent (renum.N p) = renum_states p ` urgent (N p)" if "p < n_ps"
+  unfolding
+    urgent_def N_eq[OF \<open>p < n_ps\<close>] renum.N_eq[unfolded renum_n_ps_simp, OF \<open>p < n_ps\<close>]
+  apply (subst nth_map_index)
+  subgoal
+    using \<open>p < n_ps\<close> by (simp add: n_ps_def)
+  unfolding automaton_of_def conv_automaton_def renum_automaton_def by (simp split: prod.split)
+
+lemma renum_urgency_removed:
+  "\<forall>i < n_ps. urgent (renum.N i) = {}"
+  using urgency_removed
+  apply intros
+  apply (simp only: urgent_renum_eq)
+  apply simp
+  done
+
+lemma check_bexp_renumD:
+  "check_bexp s b bv \<Longrightarrow> check_bexp (s o vars_inv) (renum_bexp b) bv"
+and is_val_renumD:
+  "is_val s e v \<Longrightarrow> is_val (s o vars_inv) (renum_exp e) v"
+   apply (induction rule: check_bexp_is_val.inducts)
+   apply (solves \<open>
+    auto
+      intro: check_bexp_is_val.intros
+      simp: renum_bexp_def renum_exp_def vars_inv_def the_inv_f_f[OF inj_renum_vars]
+  \<close>)+
+  subgoal
+    apply (clarsimp simp: renum_bexp_def renum_exp_def vars_inv_def, safe)
+    using is_val.simps apply fastforce+
+    done
+  apply (auto intro: check_bexp_is_val.intros simp: renum_exp_def)
+  done
+
+method solve_case =
+  auto 
+    intro: check_bexp_is_val.intros
+    simp: renum_bexp_def renum_exp_def vars_inv_def the_inv_f_f[OF inj_renum_vars];
+  fail
+| use is_val.simps in fastforce
+
+lemma check_bexp_renumI:
+  "check_bexp (s o vars_inv) (renum_bexp b) bv \<Longrightarrow> check_bexp s b bv"
+  and is_val_renumI:
+  "is_val (s o vars_inv) (renum_exp e) v \<Longrightarrow> is_val s e v"
+  apply (induction
+      "s o vars_inv" "renum_bexp b" bv and "s o vars_inv" "renum_exp e" _
+      arbitrary: b and e rule: check_bexp_is_val.inducts)
+  subgoal for b
+    by (cases b; solve_case)
+  subgoal for e bv b
+    by (cases b; solve_case)
+  subgoal for e1 a e2 bv b
+    by (cases b; solve_case)
+  subgoal for e1 a e2 bv b
+    by (cases b; solve_case)
+  subgoal for e1 a e2 bv b
+    by (cases b; solve_case)
+  subgoal for a v bv x b
+    by (cases b; solve_case)
+  subgoal for a v bv x b
+    by (cases b; solve_case)
+  subgoal for a v bv x b
+    by (cases b; solve_case)
+  subgoal for a v bv x b
+    by (cases b; solve_case)
+  subgoal for a v bv x b
+    by (cases b; solve_case)
+  subgoal for c e
+    by (cases e; solve_case)
+  subgoal for x v e
+    by (cases e; solve_case)
+  subgoal for e1 v1 e2 v2 b bv e
+    by (clarsimp simp: renum_bexp_def renum_exp_def; safe; cases e; solve_case)
+  subgoal for e1 v1 e2 v2 f e
+    by (clarsimp simp: renum_bexp_def renum_exp_def; cases e; solve_case)
+  subgoal for e1 v f e
+    by (cases e; solve_case)
+  done
+
+lemma renum_reset_map_u:
+  "[renum_reset r\<rightarrow>0]map_u u = map_u ([r\<rightarrow>0]u)"
+  unfolding map_u_def
+  apply (rule ext)
+  subgoal for x
+    apply (cases "clocks_inv x \<in> set r"; simp add: clocks_inv_def)
+    subgoal
+      using bij_renum_clocks[THEN bij_is_surj]
+      by (auto
+            simp: renum_reset_def f_the_inv_f[OF bij_is_inj, OF bij_renum_clocks]
+            dest: imageI[where f = renum_clocks]
+         )
+    subgoal
+      by (subst clock_set_id)
+         (auto simp: renum_reset_def the_inv_f_f[OF bij_is_inj, OF bij_renum_clocks])
+    done
+  done
+
+lemma bounded_renumI':
+  assumes "bounded bounds s'"
+  shows "bounded renum.bounds (s' o vars_inv)"
+  using assms unfolding renum.bounds_def bounds_def by (simp add: bounded_renumI)
+
+lemma bounded_renumD':
+  assumes "bounded renum.bounds (s' o vars_inv)" "dom s' \<subseteq> var_set"
+  shows "bounded bounds s'"
+  using assms unfolding renum.bounds_def bounds_def by (simp add: bounded_renumD)
+
+lemma is_upd_renumD:
+  assumes "is_upd s f s'"
+  shows "is_upd (s o vars_inv) (renum_upd f) (s' o vars_inv)"
+  using assms
+  unfolding is_upd_def renum_upd_def
+  by (force dest!: is_val_renumD
+            simp: bij_f_the_inv_f[OF bij_renum_vars] the_inv_f_f[OF inj_renum_vars] vars_inv_def)
+
+lemma is_upds_renumD:
+  assumes "is_upds s1 ps s'"
+  shows "is_upds (s1 o vars_inv) (renum_upds ps) (s' o vars_inv)"
+  using assms by induction (auto intro: is_upds.intros simp: comp_def dest!: is_upd_renumD)
+
+lemma is_upds_concat_renumD:
+  assumes "is_upds s1 (concat ps) s'"
+  shows "is_upds (s1 o vars_inv) (concat_map renum_upds ps) (s' o vars_inv)"
+  using assms by (induction ps arbitrary: s1) (auto simp: comp_def map_concat dest!: is_upds_renumD)
+
+lemma Ball_mono:
+  assumes "\<forall>x \<in> S. P x" "\<And>x. x \<in> S \<Longrightarrow> P x \<Longrightarrow> Q x"
+  shows "\<forall>x \<in> S. Q x"
+  using assms by blast
+
+lemma atLeastLessThan_upperD:
+  assumes "S \<subseteq> {l..<u}" "x \<in> S"
+  shows "x < u"
+  using assms by auto
+
+lemma imp_mono_rule:
+  assumes "P1 \<longrightarrow> P2"
+    and "Q1 \<Longrightarrow> P1"
+    and "Q1 \<Longrightarrow> P2 \<Longrightarrow> Q2"
+  shows "Q1 \<longrightarrow> Q2"
+  using assms by blast
+
+lemma inj_id:
+  "inj id"
+  by auto
+
+lemma step_single_renumD:
+  assumes "step_u sem L s u a L' s' u'" "L \<in> states" "dom s \<subseteq> var_set"
+  shows "step_u renum.sem
+    (map_index renum_states L) (s o vars_inv) (map_u u)
+    (renum_label a)
+    (map_index renum_states L') (s' o vars_inv) (map_u u')"
+  using assms(1-3)
+  apply (cases a)
+
+     supply [simp del] = map_map_index set_map
+     \<comment> \<open>To keep the simplifier from breaking through locale abstractions\<close>
+
+  \<comment> \<open>Delay\<close>
+  subgoal
+    supply [simp] = length_automata_eq_n_ps L_len
+    apply (subst (asm) A_split)
+    apply (subst renum.A_split)
+    apply (simp only: renum_label_def label.map renum_n_ps_simp)
+    apply (erule step_u_elims')
+    apply simp
+    apply (rule step_u.intros)
+    preferT \<open>''target invariant''\<close> subgoal
+      by tag (auto 4 3 simp: dest: inv_renum_sem_I[OF _ _ sem_states_loc_setD])
+    preferT \<open>''nonnegative''\<close> subgoal
+      by assumption
+    preferT \<open>''urgent''\<close> subgoal
+      using renum_urgency_removed by - (tag, auto)
+    preferT \<open>''bounded''\<close> subgoal
+      by tag (rule bounded_renumI')
+    done
+
+  \<comment> \<open>Internal\<close>
+  subgoal for a'
+    supply [simp] = length_automata_eq_n_ps L_len
+    apply (subst (asm) A_split)
+    apply (subst renum.A_split)
+    apply (simp only: renum_label_def label.map renum_n_ps_simp)
+    apply (erule step_u_elims')
+    unfolding TAG_def[of "''range''"]
+    apply simp
+    apply (rule step_u.intros)
+    preferT \<open>TRANS ''silent''\<close>
+      apply (solves \<open>tag, drule (1) trans_N_renumD, subst nth_map, (simp add: renum_act_def)+\<close>)
+    preferT \<open>''committed''\<close> subgoal
+      by tag (auto simp: committed_renum_eq dest!: inj_renum_states[THEN injD, rotated])
+    preferT \<open>''bexp''\<close> subgoal
+      by tag (erule check_bexp_renumD)
+    preferT \<open>''guard''\<close> subgoal
+      by tag (rule map_u_renum_cconstraint_clock_valI)
+    preferT \<open>''target invariant''\<close> subgoal
+      apply (tag- "''new loc''" "TRANS ''silent''")
+      apply clarsimp
+      apply (rule inv_renum_sem_I[OF _ _ sem_states_loc_setD[OF _ state_preservation_updI]])
+          apply fastforce+
+      done
+    preferT \<open>''loc''\<close>
+      apply (tag, simp; fail)
+    preferT \<open>''range''\<close>
+      apply (tag, simp; fail)
+    preferT \<open>''new loc''\<close>
+      apply (tag, simp only: map_index_update; fail)
+    preferT \<open>''new valuation''\<close>
+      apply (tag, simp only: renum_reset_map_u; fail)
+    preferT \<open>''is_upd''\<close> subgoal
+      by (tag, erule is_upds_renumD)
+    preferT \<open>''bounded''\<close>
+      apply (tag, erule bounded_renumI'; fail)
+    done
+
+  \<comment> \<open>Binary\<close>
+  subgoal for a'
+    supply [simp] = length_automata_eq_n_ps L_len
+    apply (subst (asm) A_split)
+    apply (subst renum.A_split)
+    apply (simp only: renum_label_def label.map renum_n_ps_simp)
+    apply (erule step_u_elims')
+    unfolding TAG_def[of "RECV ''range''"] TAG_def[of "SEND ''range''"]
+    apply simp
+    apply (rule step_u.intros)
+    preferT \<open>''not broadcast''\<close> subgoal
+      apply tag
+      unfolding renum.broadcast_def
+      apply (clarsimp simp: set_map)
+      apply (drule injD[OF inj_renum_acts])
+      apply (subst (asm) broadcast_def, simp)
+      done
+    preferT \<open>TRANS ''in''\<close>
+      apply (solves \<open>tag, drule (1) trans_N_renumD, subst nth_map, (simp add: renum_act_def)+\<close>)
+    preferT \<open>TRANS ''out''\<close>
+      apply (solves \<open>tag, drule (1) trans_N_renumD, subst nth_map, (simp add: renum_act_def)+\<close>)
+    preferT \<open>''committed''\<close> subgoal
+      by tag (auto simp: committed_renum_eq dest!: inj_renum_states[THEN injD, rotated])
+
+    preferT \<open>''bexp''\<close> subgoal
+      by tag (erule check_bexp_renumD)
+
+    preferT \<open>''bexp''\<close> subgoal
+      by tag (erule check_bexp_renumD)
+
+    preferT \<open>''guard''\<close> subgoal
+      by tag (rule map_u_renum_cconstraint_clock_valI)
+
+    preferT \<open>''guard''\<close> subgoal
+      by tag (rule map_u_renum_cconstraint_clock_valI)
+
+    preferT \<open>''target invariant''\<close> subgoal
+      apply (tag- "''new loc''" "TRANS _")
+      apply clarsimp
+      apply (rule inv_renum_sem_I[OF _ _ sem_states_loc_setD[OF _ state_preservation_updI]])
+          apply (fastforce intro!: state_preservation_updI)+
+      done
+    preferT \<open>''new loc''\<close>
+      apply (tag, simp add: map_index_update; fail)
+    preferT \<open>''new valuation''\<close>
+      apply (tag, simp add: renum_reset_map_u[symmetric] renum_reset_def; fail)
+    preferT \<open>''upd''\<close>
+      apply (tag, erule is_upds_renumD; fail)
+    preferT \<open>''upd''\<close>
+      apply (tag, erule is_upds_renumD; fail)
+    preferT \<open>''bounded''\<close>
+      apply (tag, erule bounded_renumI'; fail)
+    apply (tag, simp; fail)+
+    done
+
+  \<comment> \<open>Broadcast\<close>
+  subgoal for a'
+    supply [simp] = length_automata_eq_n_ps L_len
+    apply (subst (asm) A_split)
+    apply (subst renum.A_split)
+    apply (simp only: renum_label_def label.map renum_n_ps_simp)
+    apply (erule step_u_elims')
+    apply simp
+    unfolding TAG_def[of "SEND ''range''"]
+
+    apply (rule step_u.intros)
+
+    preferT \<open>''broadcast''\<close> subgoal
+      unfolding renum.broadcast_def by (tag, subst (asm) broadcast_def, simp add: set_map)
+
+    preferT \<open>TRANS ''out''\<close>
+    apply (solves\<open>tag, simp, drule (1) trans_N_renumD, subst nth_map, (simp add: renum_act_def)+\<close>)
+
+    preferT \<open>TRANS ''in''\<close>
+      apply (tag- "SEL _")
+      apply (solves \<open>auto dest!: trans_N_renumD simp add: renum_act_def atLeastLessThan_upperD\<close>)
+
+    preferT \<open>''committed''\<close> subgoal
+      apply (tag- "SEL _")
+      apply (simp add: committed_renum_eq)
+      apply (erule disj_mono[rule_format, rotated 2], (simp; fail))
+      apply (erule disj_mono[rule_format, rotated 2])
+      subgoal
+        apply clarify
+        apply (rule rev_bexI, assumption)
+        apply (auto simp: committed_renum_eq)
+        done
+      apply (auto simp: committed_renum_eq dest!: inj_renum_states[THEN injD, rotated]; fail)
+      done
+
+    preferT \<open>''maximal''\<close> subgoal
+      apply tag
+      apply simp
+      apply (erule all_mono[THEN mp, OF impI, rotated])
+      apply (erule (1) imp_mono_rule)
+      apply (erule (1) imp_mono_rule)
+      apply clarify
+      apply (drule trans_sem_N_renumI', assumption)
+      apply (untag, clarsimp simp: renum_act_def)
+      apply (drule check_bexp_renumI)
+      apply (drule InD2[OF inj_renum_acts])
+      apply (fastforce dest!: map_u_renum_cconstraint_clock_valD)
+      done
+
+    preferT \<open>''target invariant''\<close> subgoal for l b g f r l' p ps bs gs fs rs ls' s'
+      apply (tag- "SEL _" "''new loc''" "TRANS _")
+      apply (subgoal_tac "fold (\<lambda>p L. L[p := ls' p]) ps L \<in> states")
+      subgoal
+        apply simp
+        apply (erule all_mono[THEN mp, OF impI, rotated], erule (1) imp_mono_rule,
+               drule (1) inv_renum_sem_I)
+        subgoal
+          apply (rule sem_states_loc_setD, simp)
+          apply (rule state_preservation_updI)
+          subgoal
+            by blast
+          .
+        subgoal
+          apply (fo_rule match_assumption2[where P = clock_val], assumption, rule HOL.refl)
+          apply (drule states_lengthD, simp)
+          done
+        done
+      subgoal
+        apply (rule state_preservation_fold_updI)
+         apply (erule Ball_mono)
+         apply (simp add: atLeastLessThan_upperD; blast)
+        by simp
+      done
+
+    preferT \<open>''bexp''\<close>
+      apply (tag, erule check_bexp_renumD; fail)
+    preferT \<open>''bexp''\<close>
+      apply (solves \<open>tag, auto elim: check_bexp_renumD\<close>)
+    preferT \<open>''guard''\<close>
+      apply (tag, erule map_u_renum_cconstraint_clock_valI; fail)
+    preferT \<open>''guard''\<close>
+      apply (solves \<open>tag, auto elim: map_u_renum_cconstraint_clock_valI\<close>)
+    preferT \<open>''new loc''\<close>
+      apply (tag, simp add: map_trans_broad_aux1[symmetric] map_index_update; fail)
+    preferT \<open>''new valuation''\<close>
+      apply (tag, simp add: renum_reset_map_u[symmetric] renum_reset_def map_concat comp_def; fail)
+    preferT \<open>''upd''\<close>
+      apply (tag, erule is_upds_renumD; fail)
+    preferT \<open>''upds''\<close>
+      apply (tag, drule is_upds_concat_renumD, simp add: comp_def; fail)
+    preferT \<open>''bounded''\<close>
+      apply (tag, erule bounded_renumI'; fail)
+    apply (tag, simp; fail)+
+    done
+  done
+
+lemma inj_the_inv:
+  "inj (the_inv f)" if "bij f"
+  by (auto simp: bij_f_the_inv_f[OF that] dest: arg_cong[where f = f] intro!: injI)
+
+lemma inj_vars_inv:
+  "inj vars_inv"
+  using bij_renum_vars unfolding vars_inv_def by (rule inj_the_inv)
+
+lemma comp_vars_inv_upd_commute:
+  "(s o vars_inv)(x \<mapsto> y) = s(vars_inv x \<mapsto> y) o vars_inv"
+  by (intro ext) (auto dest: injD[OF inj_vars_inv])
+
+lemma is_upd_renumI:
+  assumes "is_upd (s o vars_inv) (renum_upd f) s'"
+  shows "is_upd s f (s' o renum_vars)"
+  using assms
+  unfolding is_upd_def renum_upd_def
+  by (clarsimp dest!: is_val_renumI split: prod.split_asm simp: comp_vars_inv_upd_commute)
+     (force simp: the_inv_f_f[OF inj_renum_vars] vars_inv_def)
+
+lemma is_upd_renumI':
+  assumes "is_upd (s o vars_inv) (renum_upd f) s'"
+  obtains s1 where "is_upd s f s1" "s1 = s' o renum_vars"
+  by (simp add: assms is_upd_renumI)
+
+lemma is_upd_renumI'':
+  assumes "is_upd s (renum_upd f) s'"
+  shows "is_upd (s o renum_vars) f (s' o renum_vars)"
+proof -
+  have "s = (s o renum_vars) o vars_inv"
+    by (intro ext) (auto simp: bij_f_the_inv_f[OF bij_renum_vars] vars_inv_def)
+  with assms  show ?thesis
+    by (intro is_upd_renumI) simp
+qed
+
+lemma is_upds_renumI:
+  assumes "is_upds (s o vars_inv) (renum_upds upds) s'"
+  shows "\<exists>s1. is_upds s upds s1 \<and> s1 = s' o renum_vars"
+  using assms apply (induction "s o vars_inv" "renum_upds upds" s' arbitrary: upds s)
+  subgoal for upds s
+    apply simp
+    apply (subgoal_tac "s o vars_inv o renum_vars = s")
+     apply (solves \<open>auto intro: is_upds.intros simp: comp_def\<close>)
+    apply (rule ext)
+    apply (simp add: vars_inv_def the_inv_f_f[OF inj_renum_vars])
+    done
+
+  apply clarsimp
+  apply (erule is_upd_renumI')
+  apply (auto simp: vars_inv_def bij_f_the_inv_f[OF bij_renum_vars] comp_def intro!: is_upds.intros)
+  done
+
+lemma is_upds_renumI':
+  assumes "is_upds (s o vars_inv) (renum_upds f) s'"
+  shows "is_upds s f (s' o renum_vars)"
+  using is_upds_renumI[OF assms] by simp
+
+lemma is_upds_renumI'':
+  assumes "is_upds s (renum_upds ps) s'"
+  shows "is_upds (s o renum_vars) ps (s' o renum_vars)"
+  using assms
+  by (induction "renum_upds ps" s' arbitrary: ps) (auto 4 3 intro: is_upd_renumI'' is_upds.intros)
+
+lemma is_upds_concat_renumI'':
+  assumes "is_upds s (concat_map renum_upds ps) s'"
+  shows "is_upds (s o renum_vars) (concat ps) (s' o renum_vars)"
+  apply (rule is_upds_renumI'')
+  using assms by (simp add: map_concat)
+
+lemma bounded_renumD1:
+  assumes "bounded renum.bounds s'" "dom (s' \<circ> renum_vars) \<subseteq> var_set"
+  shows "bounded bounds (s' o renum_vars)"
+  using assms
+  by (intro bounded_renumD') (auto simp: vars_inv_def bij_f_the_inv_f[OF bij_renum_vars] comp_def)
+
+lemma renum_reset_append:
+  "renum_reset xs @ renum_reset ys = renum_reset (xs @ ys)"
+  unfolding renum_reset_def by simp
+
+lemma if_eq_distrib:
+  "(if i = j then f i a else f j b) = (f j (if i = j then a else b))"
+  by auto
+
+lemma dom_comp_eq_vimage:
+  "dom (s o f) = f -` dom s"
+  unfolding dom_def by auto
+
+lemma dom_comp_vars_inv_eqD:
+  assumes "dom s' = dom (s o vars_inv)"
+  shows "dom (s' o renum_vars) = dom s"
+  using assms inj_renum_vars surj_renum_vars unfolding vars_inv_def
+  by (subst (asm) dom_the_inv_comp) (auto simp: dom_comp_eq_vimage dest: injD)
+
+lemma sem_trans_upd_domD:
+  assumes "(L ! p, b, g', a, f, r, l1) \<in> trans (N p)" "p < n_ps"
+  shows "fst ` set f \<subseteq> var_set"
+proof -
+  from assms have "fst ` set f \<subseteq> var_set"
+    unfolding var_set_def
+    apply -
+    apply (rule semilattice_sup_class.sup.coboundedI2)
+    apply clarsimp
+    apply (inst_existentials "(fst \<circ> snd \<circ> snd \<circ> snd \<circ> snd) ` trans (N p)" p)
+      apply force+
+    done
+  then show ?thesis .
+qed
+
+lemma SilD:
+  fixes map_action
+  assumes "Sil a = map_act map_action a1"
+  obtains a' where "a1 = Sil a'" "a = map_action a'"
+  using assms by (cases a1) auto
+
+lemma InD:
+  fixes map_action
+  assumes "In a = map_act map_action a1"
+  obtains a' where "a1 = In a'" "a = map_action a'"
+  using assms by (cases a1) auto
+
+lemma OutD:
+  fixes map_action
+  assumes "Out a = map_act map_action a1"
+  obtains a' where "a1 = Out a'" "a = map_action a'"
+  using assms by (cases a1) auto
+
+lemma In_OutD:
+  assumes "In a = renum_act a1" "Out a = renum_act a2"
+  obtains a' where "a = renum_acts a'" "a1 = In a'" "a2 = Out a'"
+  using assms unfolding renum_act_def by (elim InD OutD) (auto simp: injD[OF inj_renum_acts])
+
+lemma renum_sem_broadcast_eq:
+  "renum.broadcast = renum_acts ` set broadcast"
+  unfolding renum.broadcast_def by simp
+
+lemma inj_eq_iff:
+  "f x = f y \<longleftrightarrow> x = y" if "inj f"
+  using that unfolding inj_def by auto
+
+lemma trans_sem_N_renum_broadcastI:
+  assumes
+    "\<forall>p\<in>set ps. (renum_states p (L ! p), bs p, gs p, In a, fs p, rs p, ls p) \<in> trans (renum.N p)"
+    "set ps \<subseteq> {0..<n_ps}"
+  obtains bs' gs' fs' rs' ls' a' where
+    "\<forall>p\<in>set ps. (L ! p, bs' p, gs' p, In a', fs' p, rs' p, ls' p) \<in> trans (N p)"
+    "\<forall>p\<in>set ps. bs p = renum_bexp (bs' p)"
+    "\<forall>p\<in>set ps. gs p = renum_cconstraint (gs' p)"
+    "ps \<noteq> [] \<longrightarrow> a = renum_acts a'"
+    "\<forall>p\<in>set ps. fs p = renum_upds (fs' p)"
+    "\<forall>p\<in>set ps. rs p = renum_reset (rs' p)"
+    "\<forall>p\<in>set ps. ls p = renum_states p (ls' p)"
+proof -
+  define t where
+    "t p \<equiv> SOME (b', g', a', f', r', l1). (L ! p, b', g', a', f', r', l1) \<in> trans (N p)
+    \<and> bs p = renum_bexp b' \<and> gs p = renum_cconstraint g' \<and> In a = renum_act a' \<and> fs p = renum_upds f'
+    \<and> rs p = renum_reset r' \<and> ls p = renum_states p l1" for p
+  define bs' where "bs' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> b'" for p
+  define gs' where "gs' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> g'" for p
+  define as' where "as' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> a'" for p
+  define fs' where "fs' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> f'" for p
+  define rs' where "rs' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> r'" for p
+  define ls' where "ls' p \<equiv> case t p of (b', g', a', f', r', l1) \<Rightarrow> l1" for p
+
+  have *: "case t p of (b', g', a', f', r', l1) \<Rightarrow>
+    (L ! p, b', g', a', f', r', l1) \<in> trans (N p)
+    \<and> bs p = renum_bexp b' \<and> gs p = renum_cconstraint g' \<and> In a = renum_act a' \<and> fs p = renum_upds f'
+    \<and> rs p = renum_reset r' \<and> ls p = renum_states p l1" if "p \<in> set ps" for p
+    using assms(1)
+    apply -
+    apply (drule bspec[OF _ that])
+    apply (drule trans_sem_N_renumI')
+    subgoal
+      using assms(2) that by auto
+    unfolding t_def by (untag, elims, fo_rule someI, auto)
+
+  show ?thesis
+  proof (cases "ps")
+    case Nil
+    then show ?thesis
+      by (auto intro: that)
+  next
+    case (Cons p ps')
+    then have "p \<in> set ps"
+      by auto
+    define a' where "a' = as' p"
+    have 1: "as' q = a'" and "In a = renum_act a'" if "q \<in> set ps" for q
+      unfolding as'_def a'_def using *[OF that] *[OF \<open>p \<in> set ps\<close>]
+      by (auto dest: injD[OF injective_functions(3)])
+    then obtain a1 where "a' = In a1" "a = renum_acts a1"
+      using \<open>p \<in> set ps\<close> unfolding renum_act_def by (auto elim!: InD)
+    then show ?thesis
+      apply -
+      apply (rule that[of bs' gs' a1 fs' rs' ls'])
+      unfolding bs'_def gs'_def fs'_def rs'_def ls'_def
+      by (solves \<open>(intros, frule *)?, auto simp: as'_def dest: 1\<close>)+
+  qed
+qed
+
+lemmas all_mono_rule = all_mono[THEN mp, OF impI, rotated]
+
+method prop_monos =
+  erule all_mono_rule
+  | erule (1) imp_mono_rule
+  | erule disj_mono[rule_format, rotated 2]
+
+lemma inv_sem_N_renum_broadcastI:
+  assumes
+"\<forall>pa<n_ps.
+  [renum_reset r @ concat (map rs ps)\<rightarrow>0]map_u u
+    \<turnstile> inv (renum.N pa)
+        ((fold (\<lambda>p L. L[p := ls p]) ps (map_index renum_states L)) [p := renum_states p l1] ! pa)"
+"\<forall>p\<in>set ps. rs p = renum_reset (rs' p)"
+"\<forall>p\<in>set ps. ls p = renum_states p (ls' p)"
+"\<forall>p\<in>set ps. ls' p \<in> (\<Union>(l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})"
+"l1 \<in> (\<Union>(l, b, g, a, r, u, l') \<in> trans (N p). {l, l'})"
+"L \<in> states"
+shows
+"\<forall>pa<n_ps.
+  [r @ concat (map rs' ps)\<rightarrow>0]u \<turnstile> inv (N pa) ((fold (\<lambda>p L. L[p := ls' p]) ps L) [p := l1] ! pa)"
+proof -
+  have [simp]: "renum_reset r @ concat (map rs ps) = renum_reset (r @ concat (map rs' ps))"
+    using assms(2) by (simp cong: map_cong add: map_concat renum_reset_def)
+  have [simp]: "length L = n_ps"
+    using \<open>L \<in> states\<close> by auto
+  have [simp]:
+    "((fold (\<lambda>p L. L[p := ls p]) ps (map_index renum_states L)) [p := renum_states p l1] ! q)
+    = renum_states q ((fold (\<lambda>p L. L[p := ls' p]) ps L) [p := l1] ! q)"
+    if "q < n_ps" for q
+    using assms(4-) that
+    apply (cases "p = q")
+     apply (simp add: fold_upds_aux_length)
+    apply (simp
+        add: assms(3) map_trans_broad_aux1[symmetric] fold_upds_aux_length
+        cong: fold_cong)
+    done
+  have "(fold (\<lambda>p L. L[p := ls' p]) ps L)[p := l1] ! q \<in> loc_set" if "q < n_ps" for q
+    using assms(4-)
+    apply (intro sem_states_loc_setD)
+    subgoal
+      using that by (simp add: n_ps_def)
+    apply (erule state_preservation_updI)
+    by (rule state_preservation_fold_updI)
+  then show ?thesis
+    using assms
+    by (auto dest: inv_renum_sem_D simp: renum_reset_map_u simp del: map_map_index set_map)
+qed
+
+method solve_triv =
+  assumption
+  | erule (1) bounded_renumD'; fail
+  | rule inv_renum_sem_D[OF _ _ sem_states_loc_setD]; simp; fail
+  | rule check_bexp_renumI; simp; fail
+  | rule map_u_renum_cconstraint_clock_valD; simp; fail
+  | rule is_upd_renumI is_upd_renumI'' is_upds_renumI' is_upds_renumI'' is_upds_concat_renumI'',
+    simp; fail
+  | simp; fail
+  | simp add:
+      vars_inv_def bij_f_the_inv_f[OF bij_renum_vars] renum_reset_map_u[symmetric] map_index_update
+      renum_reset_append;
+    fail
+  | subst nth_map, (simp; fail)+; fail
+  | (rule state_preservation_updI, blast)+, simp; fail
+
+lemma trans_sem_upd_domI:
+  assumes "(L ! p, b', g', a, f', r', l1) \<in> trans (N p)" "dom s = var_set" "p < n_ps"
+  shows "\<forall>(x, _)\<in>set (renum_upds f'). x \<in> dom (s o vars_inv)"
+  unfolding renum_upd_def using assms
+  by (auto simp: dom_comp_eq_vimage vars_inv_def the_inv_f_f[OF inj_renum_vars]
+           dest!: sem_trans_upd_domD)
+
+lemma step_single_renumI:
+  assumes
+    "step_u renum.sem
+      (map_index renum_states L) (s o vars_inv) (map_u u) a L' s' u'"
+    "L \<in> states" "dom s \<subseteq> var_set" "dom s = var_set"
+  shows "\<exists> a1 L1 s1 u1. step_u sem L s u a1 L1 s1 u1 \<and> renum_label a1 = a \<and>
+    L' = map_index renum_states L1 \<and> s' = s1 o vars_inv \<and> u' = map_u u1"
+  using assms(1-3)
+  supply [simp] = length_automata_eq_n_ps L_len
+  supply [simp del] = map_map_index set_map
+  apply (subst A_split)
+  apply (subst (asm) renum.A_split)
+  apply (simp only: renum_label_def label.map renum_n_ps_simp)
+
+  using [[goals_limit=2]]
+
+  apply (cases a)
+
+  \<comment> \<open>Delay\<close>
+  subgoal
+    apply (simp only:)
+    apply (erule step_u_elims')
+    apply intros
+    apply (rule step_u.intros)
+    preferT \<open>''nonnegative''\<close>
+      apply assumption
+    preferT \<open>''urgent''\<close> subgoal
+      using urgency_removed by - (tag, auto)
+    preferT \<open>''target invariant''\<close>
+      apply (solves \<open>tag, simp, prop_monos+, solve_triv+\<close>)
+    apply (tag, solve_triv)+
+    done
+
+  \<comment> \<open>Internal\<close>
+  subgoal for a'
+    apply (simp only:)
+    apply (erule step_u_elims')
+    subgoal premises prems[tagged] for l b g f r l' p
+    using prems apply tag
+    using sym[OF [[get_tagged \<open>''loc''\<close>]]]
+    usingT "TRANS _" \<open>''range''\<close>
+    apply simp
+    apply (drule (1) trans_sem_N_renumI', untag)
+    apply elims
+    unfolding renum_act_def
+    apply (rule SilD, assumption)
+
+    apply intros
+    apply (rule step_u.intros(2))
+
+    preferT \<open>TRANS ''silent''\<close> apply (tag, solve_triv)
+
+    preferT \<open>''committed''\<close>
+    apply (solves \<open>tag, prop_monos; auto simp: committed_renum_eq dest: injD[OF inj_renum_states]\<close>)
+
+    preferT \<open>''bexp''\<close> apply (tag, solve_triv)
+    preferT \<open>''guard''\<close> apply (tag, solve_triv)
+
+    preferT \<open>''target invariant''\<close>
+      apply (solves \<open>tag "''new valuation''" "''new loc''",
+        simp add: map_index_update[symmetric] renum_reset_map_u, prop_monos+,
+        rule inv_renum_sem_D[OF _ _ sem_states_loc_setD],
+        solve_triv+\<close>)
+
+    preferT \<open>''is_upd''\<close>
+      apply (tag, solve_triv)
+
+    preferT \<open>''bounded''\<close> subgoal
+      apply (tag \<open>''is_upd''\<close>, erule bounded_renumD1)
+      apply (drule is_upds_dom)
+      subgoal
+        apply (simp add: dom_comp_eq_vimage)
+        unfolding renum_upd_def
+        apply (clarsimp simp: vars_inv_def the_inv_f_f[OF inj_renum_vars] set_map)
+        apply (drule (1) sem_trans_upd_domD)
+        using assms(4)
+        apply auto
+        done
+      apply (drule dom_comp_vars_inv_eqD, simp)
+      done
+
+    apply (tag, solve_triv)+
+    usingT \<open>''new loc''\<close> apply solve_triv
+
+(* or: auto *)
+    apply (rule ext)
+    usingT \<open>''new valuation''\<close> apply solve_triv+
+    done
+  done
+
+  \<comment> \<open>Binary\<close>
+  subgoal for a'
+    apply (simp only:)
+    apply (erule step_u_elims')
+    subgoal premises prems[tagged] for l1 b1 g1 f1 r1 l1' p l2 b2 g2 f2 r2 l2' q s'a
+    using prems apply tag
+    apply (tag "RECV ''loc''" "SEND ''loc''")
+    apply (drule sym[of "map_index renum_states L ! _"])+
+    apply (tag "TRANS _" "RECV ''range''" "SEND ''range''")
+    apply simp
+    apply (drule (1) trans_sem_N_renumI')
+    apply (drule (1) trans_sem_N_renumI', untag)
+    apply elims
+    apply (rule In_OutD, assumption, assumption)
+
+    apply intros
+        apply (rule step_u.intros(3))
+                        preferT \<open>''not broadcast''\<close>
+                        apply (tag, simp add: renum_sem_broadcast_eq broadcast_def)
+    using inj_renum_acts
+    apply auto[1]
+    preferT \<open>TRANS ''in''\<close> apply (tag, solve_triv)
+    preferT \<open>TRANS ''out''\<close> apply (tag, solve_triv)
+    preferT \<open>''committed''\<close> subgoal
+      by (tag, auto simp: committed_renum_eq inj_eq_iff[OF inj_renum_states])
+    preferT \<open>''target invariant''\<close>
+      apply (
+        solves \<open>tag "''new valuation''" "''new loc''",
+        simp add: map_index_update[symmetric] renum_reset_map_u renum_reset_append,
+        prop_monos+,
+        rule inv_renum_sem_D[OF _ _ sem_states_loc_setD],
+        solve_triv+\<close>)
+
+    preferT \<open>''upd''\<close> apply (tag, solve_triv)
+    preferT \<open>''upd''\<close> apply (tag, solve_triv)
+
+    preferT \<open>''bounded''\<close> subgoal
+      usingT- \<open>''upd''\<close>
+      apply -
+      apply (tag, erule bounded_renumD1)
+      apply (drule is_upds_dom)
+      subgoal
+        using assms(4) by - (drule trans_sem_upd_domI; simp split: prod.splits)
+      apply (drule is_upds_dom)
+      subgoal
+        using assms(4) by - (drule trans_sem_upd_domI; simp split: prod.splits)
+      apply (simp, drule dom_comp_vars_inv_eqD, simp)
+      done
+    apply (tag, solve_triv)+
+    apply (tag "''new loc''", solve_triv)
+    apply (rule ext; solve_triv)
+    apply (tag "''new valuation''", solve_triv)
+    done
+  done
+
+  \<comment> \<open>Broadcast\<close>
+  subgoal for a'
+    apply (simp only:)
+    apply (erule step_u_elims')
+    subgoal premises prems[tagged] for l b g f r l' p ps bs gs fs rs ls' s'a
+    using prems apply (tag "SEND ''range''" "TRANS ''out''")
+    using sym[OF [[get_tagged \<open>SEND ''loc''\<close>]]]
+    apply simp
+    apply (frule (1) trans_sem_N_renumI')
+    apply elims
+    subgoal premises prems[tagged] for b' g' a'a f' r' l1
+      using prems(1-6) usingT- "TRANS ''in''" "SEL _" \<open>''renum action''\<close>
+      unfolding renum_act_def
+      apply -
+      apply (rule OutD, assumption)
+      apply (simp add: atLeastLessThan_upperD)
+      apply (erule(1) trans_sem_N_renum_broadcastI)
+      apply (subgoal_tac "map fs ps = map renum_upds (map fs' ps)")
+       defer
+       apply (simp; fail)
+
+      subgoal for a1 bs' gs' fs' rs' ls1 a2
+
+        apply intros
+            apply (rule step_u.intros(4)[where ps = ps])
+
+        preferT \<open>TRANS ''out''\<close> apply (tag \<open>TRANS ''orig''\<close>, solve_triv)
+
+        preferT \<open>''broadcast''\<close>
+          apply (tag,
+            clarsimp simp: renum_sem_broadcast_eq inj_eq_iff[OF inj_renum_acts] broadcast_def; fail)
+
+        preferT \<open>TRANS ''in''\<close>
+        apply (tag,
+            cases "ps = []"; simp add: atLeastLessThan_upperD inj_eq_iff[OF inj_renum_acts]; fail)
+
+        preferT \<open>''committed''\<close> subgoal
+          apply (tag, prop_monos)
+          apply (solves \<open>auto simp: committed_renum_eq inj_eq_iff[OF inj_renum_states]\<close>)
+          apply prop_monos
+          subgoal
+            by (auto simp: committed_renum_eq inj_eq_iff[OF inj_renum_states] atLeastLessThan_upperD)
+          subgoal premises prems
+            using \<open>L \<in> states\<close> prems(21) (* last premise only *)
+            by (auto simp: inj_eq_iff[OF inj_renum_states] committed_renum_eq)
+          done
+
+        preferT \<open>''bexp''\<close> apply (tag \<open>''renum bexp''\<close>, solve_triv)
+
+        preferT \<open>''bexp''\<close> apply (tag, erule Ball_mono, solve_triv; fail)
+
+        preferT \<open>''guard''\<close> apply (tag \<open>''renum guard''\<close>, solve_triv)
+
+        preferT \<open>''guard''\<close> apply (tag, erule Ball_mono, solve_triv)
+
+        preferT \<open>''maximal''\<close> subgoal
+          apply tag
+          apply simp
+          apply prop_monos+
+          apply clarify
+          apply (drule (1) trans_N_renumD)+
+          apply (simp add: renum_act_def)
+          apply (drule check_bexp_renumD)
+          apply (drule map_u_renum_cconstraint_clock_valI)
+          apply blast
+          done
+
+        preferT \<open>''target invariant''\<close>
+          apply (tag "''new loc''" "''new valuation''" "''renum reset''" "''renum loc''" "TRANS _")
+          apply (simp add: atLeastLessThan_upperD)
+          apply (solves \<open>erule (2) inv_sem_N_renum_broadcastI, blast, fast, solve_triv\<close>)
+
+        preferT \<open>''upds''\<close>
+          apply (tag "''renum upds''", simp only:, solve_triv)
+
+        preferT \<open>''bounded''\<close> subgoal
+          apply (tag "''upd''" "''upds''" "''renum upds''", simp only:)
+          apply (erule bounded_renumD1)
+          apply (drule is_upds_dom)
+          subgoal
+            using assms(4) usingT \<open>TRANS ''orig''\<close> by (intro trans_sem_upd_domI)
+
+          apply (drule is_upds_dom)
+          subgoal
+            using assms(4) usingT \<open>''upd''\<close> \<open>''upds''\<close>
+            apply (simp add: set_map)
+            apply (erule Ball_mono, drule trans_sem_upd_domI, assumption)
+             apply (simp add: atLeastLessThan_upperD set_map; fail)+
+            done
+          apply (simp, drule dom_comp_vars_inv_eqD, simp)
+          done
+
+        preferT \<open>''upd''\<close> apply (tag "''renum upds''", solve_triv)
+        
+        apply (tag, solve_triv)+
+
+        apply (tag "''new loc''" "''renum loc''",
+          simp add: map_trans_broad_aux1 map_index_update cong: fold_cong; fail)
+
+        apply (solves \<open>auto simp: vars_inv_def bij_f_the_inv_f[OF bij_renum_vars]\<close>)
+        apply (tag "''new valuation''" "''renum reset''",
+          simp add: renum_reset_map_u[symmetric] renum_reset_def map_concat comp_def cong: map_cong;
+          fail)
+        done
+      done
+    done
+    done
+  done
+
+lemma step_u_var_set_invariant:
+  assumes "step_u sem L s u a L' s' u'" "dom s = var_set"
+  shows "dom s' = var_set"
+  using assms dom_bounds_var_set by (auto 4 4 dest!: bounded_inv simp: bounded_def)
+
+lemmas step_u_invariants = states_inv step_u_var_set_invariant
+
+interpretation single: Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). \<exists> a. step_u sem L s u a L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). \<exists> a. step_u renum.sem L s u a L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = map_index renum_states L \<and> s' = s o vars_inv \<and> u' = map_u u"
+  "\<lambda> (L, s, u). L \<in> states \<and> dom s = var_set" "\<lambda>_. True"
+  apply standard
+  supply [simp del] = map_map_index set_map
+     apply clarsimp
+  subgoal for L s u L' s' u' a
+    by (drule (1) step_single_renumD, auto)
+  subgoal
+    by clarsimp (drule step_single_renumI[rotated]; blast)
+  subgoal
+    by clarsimp (intro conjI; elim step_u_invariants)
+  subgoal
+    .
+  done
+
+interpretation Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' renum.sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = map_index renum_states L \<and> s' = s o vars_inv \<and> u' = map_u u"
+  "\<lambda> (L, s, u). L \<in> states \<and> dom s = var_set" "\<lambda>_. True"
+proof -
+  define rsem where "rsem = renum.sem"
+  note step_single_renumD = step_single_renumD[folded rsem_def]
+  have "rsem \<turnstile> \<langle>map_index renum_states L, (s \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u\<rangle> \<rightarrow> \<langle>map_index renum_states L', (s' \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u'\<rangle>"
+    if "sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+      and "L \<in> states"
+      and "dom s = var_set"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> int option"
+      and u :: "'c \<Rightarrow> real"
+      and L' :: "'s list"
+      and s' :: "'x \<Rightarrow> int option"
+      and u' :: "'c \<Rightarrow> real"
+  proof -
+    from that(1) obtain L1 s1 u1 a where
+      "sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L1, s1, u1\<rangle>" "a \<noteq> Del" "sem \<turnstile> \<langle>L1, s1, u1\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      by (elim step_u'_elims)
+    with that(2-) show ?thesis
+      apply -
+      apply (rule step_u'.intros[where a = "renum_label a"])
+        apply (erule (1) step_single_renumD[where a = Del, unfolded renum_label_def, simplified], blast)
+       apply (cases a; (fast | simp add: renum_label_def))
+      apply (erule step_single_renumD)
+       apply (blast dest: step_u_invariants)+
+      done
+  qed
+  moreover have "\<exists>a aa b. sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>a, aa, b\<rangle> \<and> L' = map_index renum_states a \<and> s' = (aa \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars \<and> u' = map_u b"
+    if "L \<in> states"
+      and "dom s = var_set"
+      and "rsem \<turnstile> \<langle>map_index renum_states L, (s \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> int option"
+      and u :: "'c \<Rightarrow> real"
+      and L' :: "nat list"
+      and s' :: "nat \<Rightarrow> int option"
+      and u' :: "nat \<Rightarrow> real"
+  proof -
+    from that(3) obtain L1 s1 u1 a where
+      "rsem \<turnstile> \<langle>map_index renum_states  L, (s o vars_inv), map_u u\<rangle> \<rightarrow>Del \<langle>L1, s1, u1\<rangle>"
+      "a \<noteq> Del" "rsem \<turnstile> \<langle>L1, s1, u1\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      by (elim step_u'_elims)
+    with that(1,2) show ?thesis
+      apply -
+      apply (drule (1) step_single_renumI[folded rsem_def], blast)
+      apply safe
+      subgoal premises prems for a1 L1a s1a u1a
+      proof -
+        { fix a1a L1b s1b u1b
+          have "
+           L \<in> states \<Longrightarrow>
+           dom s = var_set \<Longrightarrow>
+           renum_label a1a \<noteq>
+           Simple_Network_Language.label.Del \<Longrightarrow>
+           sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a1 \<langle>L1a, s1a, u1a\<rangle> \<Longrightarrow>
+           renum_label a1 =
+           Simple_Network_Language.label.Del \<Longrightarrow>
+           L1 = map_index renum_states L1a \<Longrightarrow>
+           u1 = map_u u1a \<Longrightarrow>
+           s1 =
+           (s1a \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv)
+            renum_vars \<Longrightarrow>
+           sem \<turnstile> \<langle>L1a, s1a, u1a\<rangle> \<rightarrow>a1a \<langle>L1b, s1b, u1b\<rangle> \<Longrightarrow>
+           a = renum_label a1a \<Longrightarrow>
+           L' = map_index renum_states L1b \<Longrightarrow>
+           s' =
+           (s1b \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv)
+            renum_vars \<Longrightarrow>
+           u' = map_u u1b \<Longrightarrow>
+           a1 = Simple_Network_Language.label.Del"
+            by (cases a1; simp add: renum_label_def)
+        } note * = this
+        from prems show ?thesis
+          apply -
+        apply (drule step_single_renumI[folded rsem_def])
+          apply (blast dest: step_u_invariants)+
+        apply (subgoal_tac "a1 = Del")
+         apply clarsimp
+         apply intros
+            apply (erule step_u'.intros)
+        apply (auto intro: *)
+        done
+    qed
+    done
+  qed
+  moreover have "x1b \<in> states \<and> dom x1c = var_set"
+    if "x1 \<in> states"
+      and "dom x1a = var_set"
+      and "sem \<turnstile> \<langle>x1, x1a, x2a\<rangle> \<rightarrow> \<langle>x1b, x1c, x2c\<rangle>"
+    for x1 :: "'s list"
+      and x1a :: "'x \<Rightarrow> int option"
+      and x2a :: "'c \<Rightarrow> real"
+      and x1b :: "'s list"
+      and x1c :: "'x \<Rightarrow> int option"
+      and x2c :: "'c \<Rightarrow> real"
+    using that by (elim step_u'_elims) (auto 4 4 dest: step_u_invariants)
+  moreover note * = calculation
+  show "Bisimulation_Invariant
+     (\<lambda>(L, s, u) (L', s', u'). sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u'). renum.sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u').
+         L' = map_index renum_states L \<and>
+         s' = (s o vars_inv) \<and>
+         u' = map_u u)
+     (\<lambda>(L, s, u). L \<in> states \<and> dom s = var_set) (\<lambda>_. True)"
+    unfolding rsem_def[symmetric]
+    apply (standard; clarsimp split: prod.splits)
+    by (rule *; assumption)+
+qed
+
+interpretation Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' renum.sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = map_index renum_states L \<and> s' = s o vars_inv \<and> u' = map_u u"
+  "\<lambda> (L, s, u). L \<in> states \<and> dom s = var_set" "\<lambda>_. True"
+proof -
+  define rsem where "rsem = renum.sem"
+  note step_single_renumD = step_single_renumD[folded rsem_def]
+  have "rsem \<turnstile> \<langle>map_index renum_states L, (s \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u\<rangle> \<rightarrow> \<langle>map_index renum_states L', (s' \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u'\<rangle>"
+    if "sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+      and "L \<in> states"
+      and "dom s = var_set"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> int option"
+      and u :: "'c \<Rightarrow> real"
+      and L' :: "'s list"
+      and s' :: "'x \<Rightarrow> int option"
+      and u' :: "'c \<Rightarrow> real"
+  proof -
+    from that(1) obtain L1 s1 u1 a where
+      "sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L1, s1, u1\<rangle>" "a \<noteq> Del" "sem \<turnstile> \<langle>L1, s1, u1\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      by (elim step_u'_elims)
+    with that(2-) show ?thesis
+      apply -
+      apply (rule step_u'.intros[where a = "renum_label a"])
+        apply (erule (1) step_single_renumD[where a = Del, unfolded renum_label_def, simplified], blast)
+       apply (cases a; (fast | simp add: renum_label_def))
+      apply (erule step_single_renumD)
+       apply (blast dest: step_u_invariants)+
+      done
+  qed
+  moreover have "\<exists>a aa b. sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>a, aa, b\<rangle> \<and> L' = map_index renum_states a \<and> s' = (aa \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars \<and> u' = map_u b"
+    if "L \<in> states"
+      and "dom s = var_set"
+      and "rsem \<turnstile> \<langle>map_index renum_states L, (s \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv) renum_vars, map_u u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+    for L :: "'s list"
+      and s :: "'x \<Rightarrow> int option"
+      and u :: "'c \<Rightarrow> real"
+      and L' :: "nat list"
+      and s' :: "nat \<Rightarrow> int option"
+      and u' :: "nat \<Rightarrow> real"
+  proof -
+    from that(3) obtain L1 s1 u1 a where
+      "rsem \<turnstile> \<langle>map_index renum_states  L, (s o vars_inv), map_u u\<rangle> \<rightarrow>Del \<langle>L1, s1, u1\<rangle>"
+      "a \<noteq> Del" "rsem \<turnstile> \<langle>L1, s1, u1\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+      by (elim step_u'_elims)
+    with that(1,2) show ?thesis
+      apply -
+      apply (drule (1) step_single_renumI[folded rsem_def], blast)
+      apply safe
+      subgoal premises prems for a1 L1a s1a u1a
+      proof -
+        {fix a1a L1b s1b u1b
+          have "
+           L \<in> states \<Longrightarrow>
+           dom s = var_set \<Longrightarrow>
+           renum_label a1a \<noteq>
+           Simple_Network_Language.label.Del \<Longrightarrow>
+           sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a1 \<langle>L1a, s1a, u1a\<rangle> \<Longrightarrow>
+           renum_label a1 =
+           Simple_Network_Language.label.Del \<Longrightarrow>
+           L1 = map_index renum_states L1a \<Longrightarrow>
+           u1 = map_u u1a \<Longrightarrow>
+           s1 =
+           (s1a \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv)
+            renum_vars \<Longrightarrow>
+           sem \<turnstile> \<langle>L1a, s1a, u1a\<rangle> \<rightarrow>a1a \<langle>L1b, s1b, u1b\<rangle> \<Longrightarrow>
+           a = renum_label a1a \<Longrightarrow>
+           L' = map_index renum_states L1b \<Longrightarrow>
+           s' =
+           (s1b \<circ>\<circ> Simple_Network_Rename_Defs.vars_inv)
+            renum_vars \<Longrightarrow>
+           u' = map_u u1b \<Longrightarrow>
+           a1 = Simple_Network_Language.label.Del"
+           by (cases a1; simp add: renum_label_def)
+        } note ** = this
+        from prems show ?thesis
+          apply -
+          apply (drule step_single_renumI[folded rsem_def])
+            apply (blast dest: step_u_invariants)+
+          apply (subgoal_tac "a1 = Del")
+           apply clarsimp
+           apply intros
+              apply (erule step_u'.intros)
+          apply (auto intro: **)
+          done
+    qed
+    done
+  qed
+  moreover have "x1b \<in> states \<and> dom x1c = var_set"
+    if "x1 \<in> states"
+      and "dom x1a = var_set"
+      and "sem \<turnstile> \<langle>x1, x1a, x2a\<rangle> \<rightarrow> \<langle>x1b, x1c, x2c\<rangle>"
+    for x1 :: "'s list"
+      and x1a :: "'x \<Rightarrow> int option"
+      and x2a :: "'c \<Rightarrow> real"
+      and x1b :: "'s list"
+      and x1c :: "'x \<Rightarrow> int option"
+      and x2c :: "'c \<Rightarrow> real"
+    using that by (elim step_u'_elims) (auto 4 4 dest: step_u_invariants)
+  moreover note * = calculation
+  show "Bisimulation_Invariant
+     (\<lambda>(L, s, u) (L', s', u'). sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u'). renum.sem \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u').
+         L' = map_index renum_states L \<and>
+         s' = (s o vars_inv) \<and>
+         u' = map_u u)
+     (\<lambda>(L, s, u). L \<in> states \<and> dom s = var_set) (\<lambda>_. True)"
+    unfolding rsem_def[symmetric]
+    apply (standard; clarsimp split: prod.splits)
+    by (rule *; assumption)+
+qed
+
+lemmas renum_bisim = Bisimulation_Invariant_axioms
+
+end
+
+locale Simple_Network_Impl_Formula_real =
+  Simple_Network_Rename_real where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, real, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, real) cconstraint) list) list" +
+  fixes \<Phi> :: "(nat, 's, 'x, int) formula"
+    and s0 :: "('x \<times> int) list"
+    and L0 :: "'s list"
+begin
+
+definition \<Phi>' where
+  "\<Phi>' = map_formula renum_states renum_vars id \<Phi>"
+
+definition a0 where
+  "a0 = (L0, map_of s0, \<lambda>_. 0)"
+
+definition a0' where
+  "a0' = (map_index renum_states L0, map_of (map (\<lambda>(x, v). (renum_vars x, v)) s0), \<lambda>_. 0)"
+
+context
+  assumes L0_states: "L0 \<in> states"
+  assumes s0_dom: "fst ` set s0 = var_set" and s0_distinct: "distinct (map fst s0)"
+begin
+
+(* Refine to subset of var_set? *)
+lemma state_eq_aux:
+  assumes "x \<notin> renum_vars ` var_set"
+  shows "vars_inv x \<notin> var_set"
+  unfolding vars_inv_def
+proof safe
+  assume "the_inv renum_vars x \<in> var_set"
+  then have "renum_vars (the_inv renum_vars x) = x"
+    by (intro f_the_inv_f inj_renum_vars) (simp add: surj_renum_vars)
+  with assms \<open>_ \<in> var_set\<close> show False
+    by force
+qed
+
+lemma state_eq:
+  assumes "fst ` set s0 = var_set" "distinct (map fst s0)"
+  shows "map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0) = (map_of s0 \<circ>\<circ>\<circ> the_inv_into) UNIV renum_vars"
+    (is "?l = ?r")
+proof (rule ext)
+  fix x
+  show "?l x = ?r x"
+  proof (cases "x \<in> renum_vars ` fst ` set s0")
+    case True
+    then show ?thesis
+      apply clarsimp
+      apply (subst map_of_mapk_SomeI')
+      subgoal
+        using inj_renum_vars by (auto intro: inj_on_subset)
+       apply (rule map_of_is_SomeI, rule assms, assumption)
+      apply (simp add: the_inv_f_f[OF inj_renum_vars] s0_distinct)
+      done
+  next
+    case False
+    then have "vars_inv x \<notin> fst ` set s0"
+      using state_eq_aux assms(1) unfolding vars_inv_def by auto
+    with False show ?thesis
+      apply -
+      apply (frule map_of_NoneI)
+      apply (simp add: vars_inv_def)
+      apply (auto simp: map_of_eq_None_iff)
+      done
+  qed
+qed
+
+interpretation Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' renum.sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = map_index renum_states L \<and> s' = s o vars_inv \<and> u' = map_u u"
+  "\<lambda> (L, s, u). L \<in> states \<and> dom s = var_set" "\<lambda>_. True"
+  by (rule renum_bisim)
+
+lemma start_equiv:
+  "A_B.equiv' a0 a0'"
+  unfolding A_B.equiv'_def a0_def a0'_def
+  apply (clarsimp simp: vars_inv_def, intro conjI)
+  subgoal
+    by (intro state_eq s0_dom s0_distinct)
+  subgoal
+    unfolding map_u_def by auto
+  subgoal
+    by (rule L0_states)
+  subgoal
+    using s0_dom dom_map_of_conv_image_fst[of s0] by fastforce
+  done
+
+lemma check_sexp_equiv:
+  assumes "A_B.equiv' (L, s, u) (L', s', u')" "locs_of_sexp e \<subseteq> {0..<n_ps}"
+  shows
+  "check_sexp e L (the \<circ> s) \<longleftrightarrow>
+   check_sexp (map_sexp renum_states renum_vars id e) L' (the \<circ> s')"
+  using assms unfolding A_B.equiv'_def
+  by (induction e)
+     (simp add: inj_eq states_lengthD renum_states_inj vars_inv_def the_inv_f_f[OF inj_renum_vars])+
+
+lemma models_iff:
+  "sem,a0 \<Turnstile> \<Phi> = renum.sem,a0' \<Turnstile> \<Phi>'" if "locs_of_formula \<Phi> \<subseteq> {0..<n_ps}"
+proof -
+  have "rel_ctl_formula compatible (ctl_of \<Phi>) (ctl_of \<Phi>')"
+    using that unfolding \<Phi>'_def
+    by (cases \<Phi>; auto simp: check_sexp_equiv prop_of_def rel_fun_def)
+  with start_equiv show ?thesis
+    by (simp add: models_ctl_iff CTL_compatible[THEN rel_funD, symmetric])
+qed
+
+lemma has_deadlock_iff:
+  "has_deadlock sem a0 \<longleftrightarrow> has_deadlock renum.sem a0'"
+  unfolding has_deadlock_def using start_equiv by (intro deadlock_iff, unfold A_B.equiv'_def) auto
+
+end (* Context for assumptions *)
+
+end (* Simple Network Rename real *)
+
+
+lemma Bisimulation_Invariants_Bisimulation_Invariant:
+  assumes "Bisimulation_Invariants A B sim PA PA PB PB"
+  shows "Bisimulation_Invariant A B sim PA PB"
+proof -
+  interpret Bisimulation_Invariants A B sim PA PA PB PB
+    by (rule assms)
+  show ?thesis
+    by (standard; blast intro: A_B_step B_A_step)
+qed
+
+lemma Bisimulation_Invariants_Bisimulation_Invariant_iff:
+  "Bisimulation_Invariants A B sim PA PA PB PB \<longleftrightarrow> Bisimulation_Invariant A B sim PA PB"
+  using
+    Bisimulation_Invariants_Bisimulation_Invariant Bisimulation_Invariant_Bisimulation_Invariants
+  by blast
+
+lemmas Bisimulation_Invariant_composition =
+  Bisimulation_Invariant_Invariants_composition[
+    THEN Bisimulation_Invariants_Bisimulation_Invariant,
+    OF _ Bisimulation_Invariant_Bisimulation_Invariants]
+
+lemma Bisimulation_Invariant_refl:
+  "Bisimulation_Invariant A A (=) P P" if "\<And>a b. P a \<Longrightarrow> A a b \<Longrightarrow> P b"
+  by (rule Bisimulation_Invariant.intro) (auto intro: that)
+
+
+locale Simple_Network_Rename_int =
+  Simple_Network_Rename_Defs_int where automata = automata +
+  Simple_Network_Rename' where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list" +
+  assumes urgency_removed: "\<forall> (_, U, _, _) \<in> set automata. U = []"
+begin
+
+lemma n_ps_eq1:
+  "Prod_TA_Defs.n_ps
+        (set broadcast, map automaton_of (map conv_automaton automata),
+         map_of bounds') = n_ps"
+  unfolding n_ps_def Prod_TA_Defs.n_ps_def by simp
+
+lemma var_set_eq1:
+  "Prod_TA_Defs.var_set
+     (set broadcast, map automaton_of (map conv_automaton automata),
+      map_of bounds') = var_set"
+  unfolding sem_def sem_var_set_eq[symmetric] by simp
+
+lemma urgency_removed':
+  "\<forall>i<n_ps. urgent
+  (Prod_TA_Defs.N (set broadcast, map automaton_of (map conv_automaton automata), map_of bounds') i)
+  = {}"
+  unfolding urgent_def n_ps_def Prod_TA_Defs.n_ps_def Prod_TA_Defs.N_def
+  unfolding conv_automaton_def automaton_of_def
+  using urgency_removed by (fastforce dest: nth_mem split: prod.split)
+
+sublocale real: Simple_Network_Rename_real where automata = "map conv_automaton automata"
+  apply standard
+  unfolding n_ps_eq1 var_set_eq1
+  by (rule bij_renum_clocks renum_states_inj bij_renum_vars bounds'_var_set inj_renum_acts
+           urgency_removed')+
+
+lemma sem_unfold1:
+  "real.sem = sem"
+  unfolding sem_def by simp
+
+lemma var_set_eq:
+  "real.var_set = sem.var_set"
+  unfolding sem_unfold1[symmetric] ..
+
+lemma map_acconstraint_conv_ac_commute:
+  "map_acconstraint renum_clocks id (conv_ac ac) = conv_ac (map_acconstraint renum_clocks id ac)"
+  by (cases ac; simp)
+
+lemma map_ccconstraint_conv_cc_commute:
+  "renum_cconstraint (conv_cc g) = conv_cc (renum_cconstraint g)"
+  unfolding renum_cconstraint_def map_cconstraint_def by (simp add: map_acconstraint_conv_ac_commute)
+
+lemma rename_conv_automaton_commute:
+  "real.renum_automaton n (conv_automaton x) = conv_automaton (real.renum_automaton n x)"
+  unfolding real.renum_automaton_def conv_automaton_def
+  by (clarsimp split: prod.split simp: map_ccconstraint_conv_cc_commute)
+
+lemma sem_unfold2:
+  "real.renum.sem = renum.sem"
+  by (simp add: Simple_Network_Impl.sem_def rename_conv_automaton_commute)
+
+sublocale renum_bisim: Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' renum.sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = map_index renum_states L \<and> s' = s o vars_inv \<and> u' = map_u u"
+  "\<lambda> (L, s, u). L \<in> sem.states \<and> dom s = var_set" "\<lambda>_. True"
+  apply (rule Bisimulation_Invariant_sim_replace)
+   apply (rule Bisimulation_Invariant_composition)
+    apply (rule real.renum_bisim[unfolded sem_unfold1 sem_unfold2 sem_var_set_eq])
+   apply (rule Bisimulation_Invariant_refl)
+   apply auto
+  done
+
+lemmas renum_bisim = renum_bisim.Bisimulation_Invariant_axioms
+
+end
+
+locale Simple_Network_Rename_Start' =
+  Simple_Network_Rename' where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, 't, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, 't) cconstraint) list) list" +
+  fixes s0 :: "('x \<times> int) list"
+    and L0 :: "'s list"
+  assumes L0_states: "L0 \<in> states"
+  assumes s0_dom: "fst ` set s0 = var_set" and s0_distinct: "distinct (map fst s0)"
+begin
+
+end
+
+locale Simple_Network_Rename_Start_int =
+  Simple_Network_Rename_int where automata = automata +
+  Simple_Network_Rename_Start' where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list"
+begin
+
+definition a0 where
+  "a0 = (L0, map_of s0, \<lambda>_. 0)"
+
+definition a0' where
+  "a0' = (map_index renum_states L0, map_of (map (\<lambda>(x, v). (renum_vars x, v)) s0), \<lambda>_. 0)"
+
+(* Refine to subset of var_set? *)
+lemma state_eq_aux:
+  assumes "x \<notin> renum_vars ` var_set"
+  shows "vars_inv x \<notin> var_set"
+  unfolding vars_inv_def
+proof safe
+  assume "the_inv renum_vars x \<in> var_set"
+  then have "renum_vars (the_inv renum_vars x) = x"
+    by (intro f_the_inv_f real.inj_renum_vars) (simp add: real.surj_renum_vars)
+  with assms \<open>_ \<in> var_set\<close> show False
+    by force
+qed
+
+lemma state_eq:
+  assumes "fst ` set s0 = var_set" "distinct (map fst s0)"
+  shows "map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0) = (map_of s0 \<circ>\<circ>\<circ> the_inv_into) UNIV renum_vars"
+    (is "?l = ?r")
+proof (rule ext)
+  fix x
+  show "?l x = ?r x"
+  proof (cases "x \<in> renum_vars ` fst ` set s0")
+    case True
+    then show ?thesis
+      apply clarsimp
+      apply (subst map_of_mapk_SomeI')
+      subgoal
+        using real.inj_renum_vars by (auto intro: inj_on_subset)
+       apply (rule map_of_is_SomeI, rule assms, assumption)
+      apply (simp add: the_inv_f_f[OF real.inj_renum_vars] s0_distinct)
+      done
+  next
+    case False
+    then have "vars_inv x \<notin> fst ` set s0"
+      using state_eq_aux assms(1) unfolding vars_inv_def by auto
+    with False show ?thesis
+      apply -
+      apply (frule map_of_NoneI)
+      apply (simp add: vars_inv_def)
+      apply (auto simp: map_of_eq_None_iff)
+      done
+  qed
+qed
+
+lemma start_equiv:
+  "renum_bisim.A_B.equiv' a0 a0'"
+  unfolding renum_bisim.A_B.equiv'_def a0_def a0'_def
+  apply (clarsimp simp: vars_inv_def, intro conjI)
+  subgoal
+    by (intro state_eq s0_dom s0_distinct)
+  subgoal
+    unfolding map_u_def by auto
+  subgoal
+    unfolding sem_states_eq by (rule L0_states)
+  subgoal
+    using s0_dom dom_map_of_conv_image_fst[of s0] by fastforce
+  done
+
+lemma check_sexp_equiv:
+  assumes "renum_bisim.A_B.equiv' (L, s, u) (L', s', u')" "locs_of_sexp e \<subseteq> {0..<n_ps}"
+  shows
+  "check_sexp e L (the \<circ> s) \<longleftrightarrow>
+   check_sexp (map_sexp renum_states renum_vars id e) L' (the \<circ> s')"
+  using assms unfolding renum_bisim.A_B.equiv'_def
+  by (induction e)
+     (simp add:
+       inj_eq sem.states_lengthD renum_states_inj vars_inv_def the_inv_f_f[OF real.inj_renum_vars])+
+
+lemma check_sexp_compatible:
+  assumes "locs_of_sexp e \<subseteq> {0..<n_ps}"
+  shows "renum_bisim.compatible
+    (\<lambda>(L, s, u). check_sexp e L (the \<circ> s))
+    (\<lambda>(L', s', u'). check_sexp (map_sexp renum_states renum_vars id e) L' (the \<circ> s'))"
+  using check_sexp_equiv[OF _ assms] by auto
+
+lemma has_deadlock_iff:
+  "has_deadlock sem a0 \<longleftrightarrow> has_deadlock renum.sem a0'"
+  unfolding has_deadlock_def using start_equiv
+  by (intro renum_bisim.deadlock_iff, unfold renum_bisim.A_B.equiv'_def) auto
+
+lemma state_formula_compatible:
+  "(\<Union>x \<in> set_state_formula \<phi>. locs_of_sexp x) \<subseteq> {0..<n_ps} \<Longrightarrow>
+  rel_state_formula renum_bisim.compatible
+    (map_state_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<phi>)
+    (map_state_formula (\<lambda>P (L, s, _).
+      check_sexp (map_sexp (\<lambda>p. renum_states p) renum_vars id P) L (the o s))
+     \<phi>)" and path_formula_compatible:
+  "(\<Union>x \<in> set_path_formula \<psi>. locs_of_sexp x) \<subseteq> {0..<n_ps} \<Longrightarrow>
+  rel_path_formula renum_bisim.compatible
+    (map_path_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<psi>)
+    (map_path_formula (\<lambda>P (L, s, _).
+      check_sexp (map_sexp (\<lambda>p. renum_states p) renum_vars id P) L (the o s))
+     \<psi>)"
+   by (induction \<phi> and \<psi>) (auto simp: check_sexp_equiv prop_of_def rel_fun_def)
+
+lemmas models_state_compatible = renum_bisim.models_state_compatible[OF state_formula_compatible]
+
+end
+
+locale Simple_Network_Rename_Formula_int =
+  Simple_Network_Rename_Start_int where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list" +
+  fixes \<Phi> :: "(nat, 's, 'x, int) formula"
+begin
+
+definition \<Phi>' where
+  "\<Phi>' = map_formula renum_states renum_vars id \<Phi>"
+
+lemma models_iff:
+  "sem,a0 \<Turnstile> \<Phi> = renum.sem,a0' \<Turnstile> \<Phi>'" if "locs_of_formula \<Phi> \<subseteq> {0..<n_ps}"
+proof -
+  have "rel_ctl_formula renum_bisim.compatible (ctl_of \<Phi>) (ctl_of \<Phi>')"
+    using that unfolding \<Phi>'_def
+    by (cases \<Phi>; auto simp: check_sexp_equiv prop_of_def rel_fun_def)
+  with start_equiv show ?thesis
+    by (simp add: models_ctl_iff renum_bisim.CTL_compatible[THEN rel_funD, symmetric])
+qed
+
+end (* Simple Network Rename Formula int *)
+
+
+
+lemma vars_of_bexp_finite[finite_intros]:
+  "finite (vars_of_bexp (b::('a, 'b) bexp))"
+and vars_of_exp_finite[finite_intros]:
+  "finite (vars_of_exp (e::('a, 'b) exp))"
+  by (induction b and e) auto
+
+lemma set_bexp_vars_of_bexp:
+  "set_bexp (b::('a, 'b) bexp) = vars_of_bexp b"
+and set_exp_vars_of_exp:
+  "set_exp (e::('a, 'b) exp) = vars_of_exp e"
+  by (induction b and e) auto
+
+definition (in Prod_TA_Defs)
+  "act_set \<equiv> (\<Union> p \<in> {0..<n_ps}. \<Union> (l, e, g, a, _) \<in> trans (N p). act.set_act a) \<union> broadcast"
+
+lemma (in Simple_Network_Impl) act_set_compute:
+  "act_set =
+  \<Union> ((\<lambda>(_, _, t, _). \<Union> ((\<lambda>(l, e, g, a, _). act.set_act a) ` set t)) ` set automata) \<union> set broadcast"
+  unfolding act_set_def
+  apply (fo_rule arg_cong2)
+  apply (clarsimp simp: N_p_trans_eq n_ps_def act_set_def broadcast_def)
+  apply safe
+     apply (clarsimp simp: N_p_trans_eq n_ps_def act_set_def broadcast_def)
+     apply (drule nth_mem, erule bexI[rotated], force)
+    apply (drule mem_nth, force)
+  unfolding broadcast_def
+  apply simp+
+  done
+
+lemma set1_acconstraint_elim:
+  assumes "c \<in> set1_acconstraint ac"
+  obtains x where "(c, x) = constraint_pair ac"
+  using assms by (cases ac) auto
+
+lemma (in Simple_Network_Impl)
+  assumes "(x1, x2, T, I) \<in> set automata" "(l, b, g, a, f, r, l') \<in> set T"
+  shows clk_set'I1[intro]: "c \<in> set r \<Longrightarrow> c \<in> clk_set'"
+    and clk_set'I2[intro]: "ac \<in> set g \<Longrightarrow> c \<in> set1_acconstraint ac \<Longrightarrow> c \<in> clk_set'"
+    and loc_setI1[intro]: "l \<in> loc_set" and loc_setI2[intro]: "l' \<in> loc_set"
+    and act_setI[intro]: "a' \<in> set_act a \<Longrightarrow> a' \<in> act_set"
+    and var_setI1[intro]: "v \<in> set_bexp b \<Longrightarrow> v \<in> var_set"
+    and var_setI2[intro]: "(x, e) \<in> set f \<Longrightarrow> x \<in> var_set"
+    and var_setI3[intro]: "(x, e) \<in> set f \<Longrightarrow> v \<in> set_exp e \<Longrightarrow> v \<in> var_set"
+  using assms unfolding
+    loc_set_compute act_set_compute var_set_compute set_bexp_vars_of_bexp set_exp_vars_of_exp
+    clk_set'_def
+         apply -
+         apply force
+        apply (fastforce elim: set1_acconstraint_elim simp: clkp_set'_def collect_clock_pairs_def)
+       apply (simp; blast)+
+  done
+
+lemma (in Simple_Network_Impl) clk_set'I3[intro]:
+  assumes "(x1, x2, T, I) \<in> set automata"
+  shows "(l, g') \<in> set I \<Longrightarrow> ac \<in> set g' \<Longrightarrow> c \<in> set1_acconstraint ac \<Longrightarrow> c \<in> clk_set'"
+  using assms unfolding clk_set'_def clkp_set'_def collect_clock_pairs_def
+  by (force elim!: set1_acconstraint_elim)
+
+
+definition
+  "find_remove P = map_option (\<lambda>(xs, x, ys). (x, xs @ ys)) o List.extract P"
+
+fun merge_pairs :: "('a \<times> 'b list) list \<Rightarrow> ('a \<times> 'b list) list \<Rightarrow> ('a \<times> 'b list) list" where
+  "merge_pairs [] ys = ys" |
+  "merge_pairs ((k, v) # xs) ys = (case find_remove (\<lambda>(k', _). k' = k) ys of
+    None \<Rightarrow> (k, v) # merge_pairs xs ys
+  | Some ((_, v'), ys) \<Rightarrow> (k, v @ v') # merge_pairs xs ys
+  )"
+
+(*
+definition
+  "conv_urge urge \<equiv> \<lambda>(committed, urgent, trans, inv).
+    (committed,
+     [],
+     map (\<lambda>(l, b, g, a, f, r, l'). (l, b, g, a, f, urge # r, l')) trans,
+     map (\<lambda>(l, cc). (l, if l \<in> set urgent then acconstraint.LE urge 0 # cc else cc)) inv)"
+*)
+
+definition
+  "conv_urge urge \<equiv> \<lambda>(committed, urgent, trans, inv).
+    (committed,
+     [],
+     map (\<lambda>(l, b, g, a, f, r, l'). (l, b, g, a, f, urge # r, l')) trans,
+     merge_pairs (map (\<lambda>l. (l, [acconstraint.LE urge 0])) urgent) inv)"
+
+(* Generalized version. Move to library *)
+lemma map_of_distinct_upd2:
+  assumes "x \<notin> set (map fst xs)"
+  shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \<mapsto> y)"
+  using assms by (induction xs) auto
+
+(* Generalized version. Move to library *)
+lemma map_of_distinct_lookup:
+  assumes "x \<notin> set (map fst xs)"
+  shows "map_of (xs @ (x,y) # ys) x = Some y"
+  using map_of_distinct_upd2[OF assms] by auto
+
+lemma find_remove_SomeD:
+  assumes "find_remove P xs = Some (x, ys)"
+  shows "\<exists>as bs. xs = as @ x # bs \<and> ys = as @ bs \<and> (\<forall>x \<in> set as. \<not> P x) \<and> P x"
+  using assms unfolding find_remove_def by (auto dest: extract_SomeE)
+
+lemma find_map:
+  "find Q (map f xs) = map_option f (find P xs)" if "\<forall>x \<in> set xs. Q (f x) \<longleftrightarrow> P x"
+  using that by (induction xs) auto
+
+lemma find_remove_map:
+  "find_remove Q (map f xs) = map_option (\<lambda>(x, xs). (f x, map f xs)) (find_remove P xs)"
+  if "\<forall>x \<in> set xs. Q (f x) \<longleftrightarrow> P x"
+  using that
+  by (induction xs)
+     (auto simp: find_remove_def extract_Nil_code extract_Cons_code split: option.split)
+
+lemma map_merge_pairs2:
+  "map (\<lambda>(k, v). (g k, map f v)) (merge_pairs xs ys)
+  = merge_pairs (map (\<lambda>(k, v). (g k, map f v)) xs) (map (\<lambda>(k, v). (g k, map f v)) ys)"
+  if inj: "inj_on g (fst ` (set xs \<union> set ys))"
+proof -
+  have *:
+    "find_remove (\<lambda>(k', _). k' = g k) (map (\<lambda>(k, v). (g k, map f v)) ys)
+   = map_option
+      (\<lambda>((k, v), ys). ((g k, map f v), map (\<lambda>(k, v). (g k, map f v)) ys))
+      (find_remove (\<lambda>(k', _). k' = k) ys)"
+    if "inj_on g (fst ` (set xs \<union> set ys))" "k \<in> fst ` set xs" for k xs ys
+    using that
+    by (subst find_remove_map[where P = "(\<lambda>(k', _). k' = k)"])
+      (auto elim: inj_onD[rotated, where A = "fst ` (set xs \<union> set ys)"]
+        intro!: arg_cong2[where f = map_option])
+  show ?thesis
+    using inj
+  proof (induction xs arbitrary: ys)
+    case Nil
+    then show ?case
+      by simp
+  next
+    case (Cons x xs)
+    then show ?case
+      apply (clarsimp split: prod.split option.split simp: *[OF Cons(2)])
+      apply (subst Cons.IH)
+       apply (drule find_remove_SomeD)
+       apply auto
+      done
+  qed
+qed
+
+lemma map_merge_pairs:
+  "map (\<lambda>(k, v). (k, map f v)) (merge_pairs xs ys)
+  = merge_pairs (map (\<lambda>(k, v). (k, map f v)) xs) (map (\<lambda>(k, v). (k, map f v)) ys)"
+  using map_merge_pairs2[where g = id] by auto
+
+lemma map_map_commute:
+  "map f (map g xs) = map g (map f xs)" if "\<forall>x \<in> set xs. f (g x) = g (f x)"
+  using that by simp
+
+lemma conv_automaton_conv_urge_commute:
+  "conv_automaton (conv_urge c A) = conv_urge c (conv_automaton A)"
+  unfolding comp_def conv_automaton_def conv_urge_def
+  by (simp split: prod.split add: map_merge_pairs comp_def)
+
+lemma conv_automaton_conv_urge_commute':
+  "conv_automaton o conv_urge c = conv_urge c o conv_automaton"
+  unfolding comp_def conv_automaton_conv_urge_commute ..
+
+locale Simple_Network_Rename_Defs_int_urge =
+  Simple_Network_Rename_Defs_int where automata = automata for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list"
+  + fixes urge :: 'c
+
+lemma (in Simple_Network_Rename_Defs) conv_automaton_of:
+  "Simple_Network_Language.conv_A (automaton_of A) = automaton_of (conv_automaton A)"
+  unfolding automaton_of_def conv_automaton_def
+    Simple_Network_Language.conv_A_def
+  by (force
+      simp: default_map_of_alt_def map_of_map Simple_Network_Language.conv_t_def split: prod.splits
+     )
+
+locale Simple_Network_Rename =
+  Simple_Network_Rename_Defs_int_urge where automata = automata for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list" +
+  assumes renum_states_inj:
+    "\<forall>i<n_ps. \<forall>x\<in>loc_set. \<forall>y\<in>loc_set. renum_states i x = renum_states i y \<longrightarrow> x = y"
+  and renum_clocks_inj: "inj_on renum_clocks (insert urge clk_set')"
+  and renum_vars_inj:   "inj_on renum_vars var_set"
+  and renum_acts_inj: "inj_on renum_acts act_set"
+  and infinite_types:
+    "infinite (UNIV :: 'c set)" "infinite (UNIV :: 'x set)" "infinite (UNIV :: 's set)"
+    "infinite (UNIV :: 'a set)"
+  and bounds'_var_set: "fst ` set bounds' = var_set"
+  and loc_set_invs: "\<Union> ((\<lambda>g. fst ` set g) ` set (map (snd o snd o snd) automata)) \<subseteq> loc_set"
+  and loc_set_committed: "\<Union> ((set o fst) ` set automata) \<subseteq> loc_set"
+  and loc_set_urgent: "\<Union> ((set o (fst o snd)) ` set automata) \<subseteq> loc_set"
+  and urge_not_in_clk_set: "urge \<notin> clk_set'"
+begin
+
+lemma clk_set'_finite:
+  "finite clk_set'"
+  unfolding clk_set'_def unfolding clkp_set'_def by (intro finite_intros) auto
+
+lemmas [finite_intros] = trans_N_finite
+
+lemma set_act_finite[finite_intros]:
+  "finite (set_act a)"
+  by (cases a) auto
+
+lemma loc_set_finite:
+  "finite loc_set"
+  unfolding loc_set_def by (auto intro!: finite_intros)
+
+lemma var_set_finite[finite_intros]:
+  "finite var_set"
+  unfolding var_set_def by (auto intro!: finite_intros)
+
+lemma act_set_finite:
+  "finite act_set"
+  unfolding act_set_def broadcast_def by (auto intro!: finite_intros)
+
+lemma bij_extend_bij_renum_clocks:
+  "bij (extend_bij renum_clocks (insert urge clk_set'))"
+  by (intro extend_bij_bij renum_clocks_inj clk_set'_finite infinite_types finite.insertI) simp
+
+lemma renum_vars_bij_extends[simp]:
+  "extend_bij renum_vars var_set x = renum_vars x" if "x \<in> var_set"
+  by (intro extend_bij_extends[rule_format] renum_vars_inj var_set_finite infinite_types that)
+
+lemma bij_extend_bij_renum_states:
+  "bij (extend_bij renum_vars var_set)"
+  by (intro extend_bij_bij renum_vars_inj var_set_finite infinite_types) simp
+
+lemma inj_extend_bij_renum_states: "inj (extend_bij (renum_states p) loc_set)" if "p < n_ps"
+  using renum_states_inj infinite_types loc_set_finite \<open>p < n_ps\<close>
+  by (intro extend_bij_inj) (auto intro!: inj_onI)
+
+lemma renum_states_extend:
+  "extend_bij (renum_states p) loc_set l = renum_states p l" if "l \<in> loc_set" "p < n_ps"
+  using renum_states_inj infinite_types loc_set_finite \<open>p < n_ps\<close> \<open>l \<in> loc_set\<close>
+  by (intro extend_bij_extends[rule_format]) (auto intro!: inj_onI)
+
+lemma renum_acts_bij_extends[simp]:
+  "extend_bij renum_acts act_set x = renum_acts x" if "x \<in> act_set"
+  by (intro extend_bij_extends[rule_format] renum_acts_inj act_set_finite infinite_types that)
+
+lemma inj_extend_bij_renum_acts: "inj (extend_bij renum_acts act_set)"
+  using renum_acts_inj infinite_types act_set_finite by (intro extend_bij_inj) auto
+
+lemma constraint_clk_conv_ac:
+  "constraint_clk (conv_ac ac) = constraint_clk ac"
+  by (cases ac) auto
+
+interpretation urge: Prod_TA_sem_urge
+  "(set broadcast, map (Simple_Network_Language.conv_A o automaton_of) automata, map_of bounds')"
+  urge
+  apply (standard; clarsimp)
+  subgoal
+    using urge_not_in_clk_set
+    unfolding conv_automaton_of
+    unfolding automaton_of_def conv_automaton_def
+    apply (clarsimp split: prod.split_asm)
+    apply (drule default_map_of_in_listD)
+    unfolding clk_set'_def clkp_set'_def collect_clock_pairs_def
+    apply clarsimp
+    subgoal premises prems
+      using prems(1,7,8,10)
+      unfolding constraint_clk_conv_ac unfolding constraint_clk_constraint_pair
+      by force
+    done
+  subgoal
+    using urge_not_in_clk_set
+    unfolding conv_automaton_of
+    unfolding automaton_of_def conv_automaton_def
+    apply (clarsimp split: prod.split_asm)
+    unfolding clk_set'_def clkp_set'_def collect_clock_pairs_def
+    unfolding constraint_clk_conv_ac unfolding constraint_clk_constraint_pair
+    apply force
+    done
+
+  subgoal
+    using urge_not_in_clk_set
+    unfolding conv_automaton_of
+    unfolding automaton_of_def conv_automaton_def
+    apply (clarsimp split: prod.split_asm)
+    unfolding clk_set'_def clkp_set'_def collect_clock_pairs_def
+    apply force
+    done
+  done
+
+
+sublocale rename: Simple_Network_Rename_Defs_int
+  broadcast bounds'
+  "extend_bij renum_acts act_set"
+  "extend_bij renum_vars var_set"
+  "extend_bij renum_clocks (insert urge clk_set')"
+  "\<lambda>p. extend_bij (renum_states p) loc_set"
+  "map (conv_urge urge) automata"
+  .
+
+sublocale rename: Simple_Network_Rename_int
+  broadcast bounds'
+  "extend_bij renum_acts act_set"
+  "extend_bij renum_vars var_set"
+  "extend_bij renum_clocks (insert urge clk_set')"
+  "\<lambda>p. extend_bij (renum_states p) loc_set"
+  "map (conv_urge urge) automata"
+  apply (standard;
+      (intro allI impI bij_extend_bij_renum_clocks inj_extend_bij_renum_states
+        inj_extend_bij_renum_acts bij_extend_bij_renum_states bounds'_var_set)?)
+    apply (simp add: Prod_TA_Defs.n_ps_def; fail)
+  subgoal
+    unfolding bounds'_var_set rename.var_set_compute var_set_compute unfolding conv_urge_def
+    by (fo_rule arg_cong2; fastforce)
+  subgoal
+    unfolding conv_urge_def by auto
+  done
+
+definition
+  "renum_upd' = map (\<lambda>(x, upd). (renum_vars x, map_exp renum_vars upd))"
+
+definition
+  "renum_reset' = map renum_clocks"
+
+definition renum_automaton' where
+  "renum_automaton' i \<equiv> \<lambda>(committed, trans, inv). let
+    committed' = map (renum_states i) committed;
+    trans' = map (\<lambda>(l, g, a, upd, r, l').
+      (renum_states i l,
+      map_cconstraint renum_clocks id g, renum_act a, renum_upd' upd, map renum_clocks r, 
+      renum_states i l')
+    ) trans;
+    inv' = map (\<lambda>(l, g). (renum_states i l, map_cconstraint renum_clocks id g)) inv
+  in (committed', trans', inv')
+"
+
+lemmas renum_automaton'_alt_def = renum_automaton'_def[unfolded
+  renum_reset'_def renum_upd'_def map_cconstraint_def renum_act_def
+  ]
+
+lemma renum_automaton_eq:
+  "rename.renum_automaton p (automata ! p) = renum_automaton p (automata ! p)"
+  if "p < n_ps"
+proof -
+  have renum_clocks: "extend_bij renum_clocks (insert urge clk_set') c = renum_clocks c"
+    if "c \<in> clk_set'" for c
+    apply (rule extend_bij_extends[rule_format])
+       apply (rule renum_clocks_inj clk_set'_finite finite.insertI)+
+    using that infinite_types by auto
+  have renum_cconstraint: "rename.renum_cconstraint g = renum_cconstraint g"
+    if "\<Union> (set1_acconstraint ` (set g)) \<subseteq> clk_set'" for g
+    unfolding rename.renum_cconstraint_def renum_cconstraint_def map_cconstraint_def
+    apply (rule map_cong, rule HOL.refl)
+    apply (rule acconstraint.map_cong_pred, rule HOL.refl)
+    using that
+    apply (auto intro: renum_clocks[simplified] simp: pred_acconstraint_def)
+    done
+  show ?thesis
+    using that[folded length_automata_eq_n_ps]
+    unfolding rename.renum_automaton_def renum_automaton_def
+    apply (clarsimp split: prod.split)
+    apply safe
+    subgoal committed
+      using loc_set_committed
+      by (subst renum_states_extend) (auto 4 3 simp: n_ps_def dest!: nth_mem)
+    subgoal urgent
+      using loc_set_urgent
+      by (subst renum_states_extend) (auto 4 3 simp: n_ps_def dest!: nth_mem)
+    subgoal start_locs
+      by (subst renum_states_extend) (auto simp: n_ps_def dest!: nth_mem)
+    subgoal state
+      unfolding rename.renum_bexp_def renum_bexp_def
+      by (auto dest!: nth_mem intro!: renum_vars_bij_extends bexp.map_cong_pred simp: pred_bexp_def)
+    subgoal guards
+      by (auto dest!: nth_mem intro!: renum_cconstraint)
+    subgoal actions
+      unfolding rename.renum_act_def renum_act_def
+      by (auto simp: pred_act_def dest!: nth_mem intro!: renum_acts_bij_extends act.map_cong_pred)
+    subgoal upds
+      unfolding renum_upd_def rename.renum_upd_def rename.renum_exp_def renum_exp_def
+      by (auto dest!: nth_mem intro!: renum_vars_bij_extends exp.map_cong_pred simp: pred_exp_def)
+    subgoal clock_resets
+      unfolding rename.renum_reset_def renum_reset_def by (auto dest!: nth_mem intro!: renum_clocks)
+    subgoal dest_locs
+      by (subst renum_states_extend) (auto simp: n_ps_def dest!: nth_mem)
+    subgoal inv_locs
+      using loc_set_invs
+      by (subst renum_states_extend) (auto 4 4 simp: n_ps_def dest!: nth_mem)
+    subgoal renum_clocks
+      by (auto dest!: nth_mem intro!: renum_cconstraint)
+    done
+qed
+
+lemma renum_urge_automaton_eq1:
+  "renum_automaton p (conv_urge urge (automata ! p))
+  = conv_urge (renum_clocks urge) (renum_automaton p (automata ! p))" if "p < n_ps"
+  unfolding renum_automaton_def conv_urge_def
+  apply (simp split: prod.split)
+  apply safe
+  subgoal
+    unfolding renum_reset_def by simp
+  unfolding renum_cconstraint_def map_cconstraint_def
+  apply (subst map_merge_pairs2)
+  subgoal
+    using loc_set_urgent loc_set_invs \<open>p < n_ps\<close> unfolding inj_on_def n_ps_def
+    by (auto; force
+        elim!: renum_states_inj[unfolded n_ps_def, simplified, rule_format, rotated -1]
+        intro: nth_mem)
+  apply (fo_rule arg_cong2; simp)
+  done
+
+lemma renum_urge_automaton_eq2:
+  "rename.real.renum_automaton p (conv_urge urge (automata ! p))
+  = conv_urge
+      (extend_bij renum_clocks (insert urge clk_set') urge)
+      (rename.renum_automaton p (automata ! p))"
+  if "p < n_ps"
+  unfolding rename.renum_automaton_def conv_urge_def
+  apply (simp split: prod.split)
+  apply safe
+  subgoal
+    unfolding rename.renum_reset_def by simp
+  unfolding rename.renum_cconstraint_def map_cconstraint_def
+  apply (subst map_merge_pairs2)
+  subgoal
+    by (rule inj_extend_bij_renum_states that subset_UNIV inj_on_subset)+
+  apply (fo_rule arg_cong2)
+   apply simp+
+  done
+
+lemma renum_urge_automaton_eq:
+  "rename.real.renum_automaton p (conv_urge urge (automata ! p)) =
+   renum_automaton p (conv_urge urge (automata ! p))" if "p < n_ps"
+  using that
+  unfolding renum_urge_automaton_eq1[OF that] renum_urge_automaton_eq2[OF that]
+  apply (fo_rule arg_cong2)
+  subgoal
+    by (intro
+      extend_bij_extends[rule_format] renum_clocks_inj clk_set'_finite finite.insertI infinite_types
+      ) auto
+  by (rule renum_automaton_eq)
+
+lemma rename_N_eq_sem:
+  "rename.renum.sem =
+  Simple_Network_Language_Model_Checking.N
+    (map renum_acts broadcast)
+    (map_index renum_automaton (map (conv_urge urge) automata))
+    (map (\<lambda>(a,p). (renum_vars a,p)) bounds')
+  "
+  unfolding rename.renum.sem_def
+  unfolding Simple_Network_Language.conv_def
+  apply simp
+  apply (intro conjI)
+  subgoal
+    by (rule image_cong; intro HOL.refl renum_acts_bij_extends; simp add: act_set_def broadcast_def)
+  subgoal
+    apply (rule map_index_cong)
+    unfolding conv_automaton_of
+    apply safe
+    apply (fo_rule arg_cong)+
+    apply (erule renum_urge_automaton_eq[folded length_automata_eq_n_ps])
+    done
+  subgoal
+    using bounds'_var_set by - (fo_rule arg_cong, auto intro: renum_vars_bij_extends)
+  done
+
+lemma map_of_merge_pairs:
+  "map_of (merge_pairs xs ys) = (\<lambda>x.
+  (if x \<in> fst ` set xs \<and> x \<in> fst ` set ys then Some (the (map_of xs x) @ the (map_of ys x))
+   else if x \<in> fst ` set xs then map_of xs x
+   else map_of ys x))"
+proof -
+  have 1: False if "find_remove (\<lambda>(k', _). k' = k) ys = None" "(k, x) \<in> set ys" for k x ys
+    using that unfolding find_remove_def by (auto simp: extract_None_iff)
+  have 2: "map_of xs k = Some x" if
+    "find_remove (\<lambda>(k', _). k' = k) xs = Some ((k', x), ys)" for k k' x xs ys
+    using that
+    by (auto 4 4 split: option.split dest: map_of_SomeD find_remove_SomeD simp: map_add_def)
+  show ?thesis
+    apply (rule ext)
+    apply (induction xs arbitrary: ys)
+     apply (simp; fail)
+    apply (clarsimp split: if_split_asm option.split)
+    apply (auto 4 3 split: option.split simp: map_add_def dest: find_remove_SomeD 2 1)
+    done
+qed
+
+lemma default_map_of_merge_pairs:
+  "default_map_of [] (merge_pairs xs ys) = (\<lambda>x.
+  (if x \<in> fst ` set xs then the (map_of xs x) @ default_map_of [] ys x
+   else default_map_of [] ys x))"
+  unfolding default_map_of_alt_def map_of_merge_pairs
+  by (rule ext) (auto 4 3 dest: map_of_SomeD weak_map_of_SomeI split: if_split_asm)
+
+lemma automaton_of_conv_urge_commute:
+  "automaton_of (conv_urge urge A) = urge.add_reset (urge.add_inv (automaton_of A))"
+  unfolding conv_urge_def urge.add_reset_def urge.add_inv_def automaton_of_def
+  apply (clarsimp split: prod.splits)
+  apply (rule ext)
+  unfolding default_map_of_merge_pairs
+  apply (clarsimp, safe)
+  apply (clarsimp)
+  subgoal premises prems for _ urgent _ _ l
+  proof -
+    have *: "map_of (map (\<lambda>x. (x, [acconstraint.LE urge 0])) urgent)
+      = (\<lambda>l. if l \<in> set urgent then Some [acconstraint.LE urge 0] else None)"
+      using map_of_map_keys[where
+          m = "\<lambda>l. if l \<in> set urgent then Some [acconstraint.LE urge 0] else None"]
+      by (force cong: map_cong simp: dom_def)
+    from \<open>l \<in> set urgent\<close> show ?thesis
+      by (subst *) auto
+  qed
+  by auto
+
+lemma urge_commute:
+  "rename.sem = urge.A_urge"
+  unfolding rename.sem_def urge.A_urge_def
+  apply simp
+  unfolding conv_automaton_conv_urge_commute conv_automaton_of automaton_of_conv_urge_commute
+  by fast
+
+lemma conv_automaton_of':
+  "automaton_of \<circ> conv_automaton = Simple_Network_Language.conv_A o automaton_of"
+  unfolding comp_def conv_automaton_of ..
+
+sublocale urge_bisim: Bisimulation_Invariant
+  "\<lambda>(L, s, u) (L', s', u'). step_u' sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). step_u' rename.sem L s u L' s' u'"
+  "\<lambda>(L, s, u) (L', s', u'). L' = L \<and> s' = s \<and> u' = u(urge := 0)"
+  "\<lambda> (L, s, u). L \<in> states" "\<lambda>(L, s, u). True"
+  unfolding urge_commute sem_def conv_automaton_of' by (rule urge.urge_bisim[unfolded conv_states])
+
+lemma check_sexp_equiv:
+  assumes "urge_bisim.A_B.equiv' (L, s, u) (L', s', u')"
+  shows "check_sexp e L (the o s) = check_sexp e L' (the o s')"
+  using assms unfolding urge_bisim.A_B.equiv'_def by simp
+
+context
+  fixes a0 :: "'s list \<times> ('x \<Rightarrow> int option) \<times> ('c \<Rightarrow> real)"
+  assumes start: "fst a0 \<in> states" "snd (snd a0) urge = 0"
+begin
+
+lemma start_equiv: "urge_bisim.A_B.equiv' a0 a0"
+  using start unfolding urge_bisim.A_B.equiv'_def by auto
+
+lemma urge_models_iff:
+  "sem,a0 \<Turnstile> \<Phi> \<longleftrightarrow> rename.sem,a0 \<Turnstile> \<Phi>"
+proof -
+  have "rel_ctl_formula urge_bisim.compatible (ctl_of \<Phi>) (ctl_of \<Phi>)"
+    by (cases \<Phi>) (auto simp: prop_of_def rel_fun_def check_sexp_equiv)
+  with start_equiv show ?thesis
+    by (simp only: models_ctl_iff urge_bisim.CTL_compatible[THEN rel_funD, symmetric])
+qed
+
+lemma urge_has_deadlock_iff:
+  "has_deadlock sem a0 \<longleftrightarrow> has_deadlock rename.sem a0"
+  unfolding has_deadlock_def using start_equiv
+  by (intro urge_bisim.deadlock_iff, unfold urge_bisim.A_B.equiv'_def) auto
+
+lemma state_formula_compatible:
+  "(\<Union>x \<in> set_state_formula \<phi>. locs_of_sexp x) \<subseteq> {0..<n_ps} \<Longrightarrow>
+  rel_state_formula urge_bisim.compatible
+    (map_state_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<phi>)
+    (map_state_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<phi>)" and path_formula_compatible:
+  "(\<Union>x \<in> set_path_formula \<psi>. locs_of_sexp x) \<subseteq> {0..<n_ps} \<Longrightarrow>
+  rel_path_formula urge_bisim.compatible
+    (map_path_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<psi>)
+    (map_path_formula (\<lambda>P (L, s, _). check_sexp P L (the o s)) \<psi>)"
+   by (induction \<phi> and \<psi>) (auto simp: check_sexp_equiv prop_of_def rel_fun_def)
+
+lemmas urge_models_state_compatible =
+  urge_bisim.models_state_compatible[OF state_formula_compatible]
+
+end (* Start State *)
+
+end (* Simple_Network_Rename *)
+
+
+locale Simple_Network_Rename_Start =
+  Simple_Network_Rename where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list" +
+  fixes s0 :: "('x \<times> int) list"
+    and L0 :: "'s list"
+  assumes L0_states: "L0 \<in> states"
+      and s0_dom: "fst ` set s0 = var_set" and s0_distinct: "distinct (map fst s0)"
+begin
+
+lemma rename_n_ps_eq:
+  "rename.n_ps = n_ps"
+  unfolding rename.length_automata_eq_n_ps[symmetric] n_ps_def by simp
+
+lemma rename_states_eq:
+  "rename.states = states"
+  unfolding rename.states_def states_def rename_n_ps_eq
+  by (simp add: rename.N_eq[unfolded rename_n_ps_eq] N_eq n_ps_def del: map_map)
+     (auto simp: automaton_of_def conv_urge_def trans_def split: prod.splits)
+
+lemma state_eq:
+  "map_of (map (\<lambda>(x, y). (extend_bij renum_vars var_set x, y)) s0) =
+    map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0)"
+  using s0_dom by - (rule arg_cong, auto intro: renum_vars_bij_extends)
+
+lemma sexp_eq:
+  assumes
+    \<open>vars_of_sexp e \<subseteq> var_set\<close>
+    \<open>set2_sexp e \<subseteq> loc_set\<close>
+    \<open>locs_of_sexp e \<subseteq> {0..<n_ps}\<close>
+  shows \<open>map_sexp (\<lambda>p. extend_bij (renum_states p) loc_set) (extend_bij renum_vars var_set) id e =
+         map_sexp renum_states renum_vars id e\<close>
+  using assms by (induction e; clarsimp simp: renum_states_extend)
+
+lemma L0_dom:
+  "length L0 = n_ps" "set L0 \<subseteq> loc_set"
+  using L0_states by (auto intro!: states_loc_setD)
+
+definition
+  "a0 = (L0, map_of s0, \<lambda>_. 0)"
+
+definition
+  "a0' = (
+    map_index (\<lambda>p. renum_states p) L0,
+    map_of (map (\<lambda>(x, y). (renum_vars x, y)) s0),
+    \<lambda>_. 0)"
+
+sublocale rename: Simple_Network_Rename_Start_int
+  broadcast bounds'
+  "extend_bij renum_acts act_set"
+  "extend_bij renum_vars var_set"
+  "extend_bij renum_clocks (insert urge clk_set')"
+  "\<lambda>p. extend_bij (renum_states p) loc_set"
+  s0 L0
+  "map (conv_urge urge) automata"
+  apply (standard; (rule L0_states[folded rename_states_eq] s0_distinct)?)
+  unfolding s0_dom rename.bounds'_var_set[symmetric] bounds'_var_set ..
+
+lemma rename_a0_eq:
+  "rename.a0 = a0"
+  unfolding rename.a0_def a0_def ..
+
+lemma rename_a0'_eq:
+  "rename.a0' = a0'"
+  unfolding a0'_def rename.a0'_def
+  apply (clarsimp, rule conjI)
+  subgoal
+    using L0_dom by (auto intro!: map_index_cong renum_states_extend)
+  subgoal
+    unfolding vars_inv_def using s0_dom s0_distinct
+    by (auto simp: state_eq Simple_Network_Rename_Defs.vars_inv_def)
+  done
+
+lemma has_deadlock_iff:
+  "has_deadlock rename.renum.sem a0' \<longleftrightarrow> has_deadlock sem a0"
+proof -
+  have "has_deadlock sem a0 \<longleftrightarrow> has_deadlock rename.sem a0"
+    by (rule urge_has_deadlock_iff) (auto intro: L0_states simp: a0_def)
+  also have "\<dots> \<longleftrightarrow> has_deadlock rename.renum.sem a0'"
+    by (rule rename.has_deadlock_iff[unfolded rename_a0_eq rename_a0'_eq])
+  finally show ?thesis ..
+qed
+
+lemma N_eq_sem:
+  "Simple_Network_Language_Model_Checking.N broadcast automata bounds' = sem"
+  unfolding conv_alt_def sem_def
+  by safe (rule nth_equalityI; simp add: conv_N_eq N_eq sem_N_eq conv_automaton_of n_ps_def)
+
+lemma rename_N_eq_sem':
+  "Simple_Network_Language_Model_Checking.N
+    (map renum_acts broadcast)
+    (map_index renum_automaton automata)
+    (map (\<lambda>(a,p). (renum_vars a,p)) bounds')
+  = renum.sem"
+  unfolding renum.sem_def
+  unfolding renum.conv_alt_def
+  by safe (rule nth_equalityI; simp add: conv_N_eq N_eq sem_N_eq conv_automaton_of n_ps_def)
+
+lemmas has_deadlock_iff' =
+  has_deadlock_iff[unfolded rename_N_eq_sem, folded N_eq_sem, unfolded a0_def a0'_def]
+
+lemmas start_equiv = start_equiv[of a0, unfolded a0_def, simplified, OF L0_states]
+
+lemmas urge_models_state_compatible =
+  urge_models_state_compatible[THEN rel_funD, OF _ _ _ start_equiv,
+    of a0, unfolded a0_def, simplified, OF L0_states]
+
+lemmas rename_models_state_compatible =
+  rename.models_state_compatible[THEN rel_funD, OF _ rename.start_equiv,
+    unfolded rename_a0_eq a0_def rename_n_ps_eq rename_a0'_eq a0'_def]
+
+lemmas models_state_compatible =
+  transp_on_equality[THEN transpD, OF urge_models_state_compatible rename_models_state_compatible]
+
+lemmas models_state_compatible' = models_state_compatible[unfolded rename_N_eq_sem, folded N_eq_sem]
+
+end
+
+locale Simple_Network_Rename_Formula =
+  Simple_Network_Rename_Start where automata = automata
+  for automata ::
+    "('s list \<times> 's list
+      \<times> (('a :: countable) act, 's, 'c, int, 'x :: countable, int) transition list
+      \<times> (('s :: countable) \<times> ('c :: countable, int) cconstraint) list) list" +
+  fixes \<Phi> :: "(nat, 's, 'x, int) formula"
+  assumes formula_dom:
+    "set2_formula \<Phi> \<subseteq> loc_set"
+    "locs_of_formula \<Phi> \<subseteq> {0..<n_ps}"
+    "vars_of_formula \<Phi> \<subseteq> var_set"
+begin
+
+sublocale rename: Simple_Network_Rename_Formula_int
+  broadcast bounds'
+  "extend_bij renum_acts act_set"
+  "extend_bij renum_vars var_set"
+  "extend_bij renum_clocks (insert urge clk_set')"
+  "\<lambda>p. extend_bij (renum_states p) loc_set"
+  s0 L0
+  "map (conv_urge urge) automata"
+  apply (standard; (rule L0_states[folded rename_states_eq] s0_distinct)?)
+  unfolding s0_dom rename.bounds'_var_set[symmetric] bounds'_var_set .
+
+definition
+  "\<Phi>' = map_formula (\<lambda>p. renum_states p) renum_vars id \<Phi>"
+
+lemma rename_\<Phi>'_eq:
+  "rename.\<Phi>' = \<Phi>'"
+  using formula_dom unfolding rename.\<Phi>'_def \<Phi>'_def by (induction \<Phi>; clarsimp simp: sexp_eq)
+
+lemma models_iff1:
+  "rename.renum.sem,a0' \<Turnstile> \<Phi>' \<longleftrightarrow> rename.sem,a0 \<Turnstile> \<Phi>"
+  by (intro rename.models_iff[unfolded rename_a0_eq rename_a0'_eq rename_\<Phi>'_eq rename_n_ps_eq]
+       formula_dom sym)
+
+lemma models_iff2:
+  "rename.sem,a0 \<Turnstile> \<Phi> \<longleftrightarrow> sem,a0 \<Turnstile> \<Phi>"
+  by (rule sym, intro urge_models_iff formula_dom) (auto intro: formula_dom L0_states simp: a0_def)
+
+lemma models_iff:
+  "rename.renum.sem,a0' \<Turnstile> \<Phi>' \<longleftrightarrow> sem,a0 \<Turnstile> \<Phi>"
+  unfolding models_iff1 models_iff2 ..
+
+lemmas models_iff' =
+  models_iff[unfolded rename_N_eq_sem, folded N_eq_sem, unfolded a0_def a0'_def \<Phi>'_def]
+
+end (* Simple_Network_Rename_Formula *)
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/TA_Equivalences.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/TA_Equivalences.thy
@@ -0,0 +1,64 @@
+theory TA_Equivalences
+  imports
+    Timed_Automata.Timed_Automata
+    Munta_Base.Normalized_Zone_Semantics_Impl_Refine
+begin
+
+locale TA_Equiv =
+  fixes A B :: "('a, 'c, 't :: time, 's) ta"
+  fixes S :: "'s set"
+  assumes state_set_trans_of: "state_set (trans_of A) \<subseteq> S"
+  assumes invs: "\<forall> l \<in> S. inv_of A l = inv_of B l"
+  assumes trans_eq: "trans_of A = trans_of B"
+begin
+
+lemma delay1:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l',u'\<rangle>" if "l \<in> S" "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l',u'\<rangle>"
+  using that invs by (auto elim!: step_t.cases)
+
+lemma action1:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>a \<langle>l',u'\<rangle>" if "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow>a \<langle>l',u'\<rangle>"
+proof -
+  from that have "l \<in> S" "l' \<in> S"
+    using state_set_trans_of by (auto elim!: step_a.cases simp: trans_eq state_set_def)
+  with that invs show ?thesis
+    by (auto elim!: step_a.cases simp: trans_eq intro: step_a.intros)
+qed
+
+lemma delay2:
+  "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l',u'\<rangle>" if "l \<in> S" "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>d \<langle>l',u'\<rangle>"
+  using that invs by (auto elim!: step_t.cases)
+
+lemma action2:
+  "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow>a \<langle>l',u'\<rangle>" if "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>a \<langle>l',u'\<rangle>"
+proof -
+  from that have "l \<in> S" "l' \<in> S"
+    using state_set_trans_of by (auto elim!: step_a.cases simp: trans_eq[symmetric] state_set_def)
+  with that invs show ?thesis
+    by (auto elim!: step_a.cases simp: trans_eq [symmetric] intro: step_a.intros)
+qed
+
+lemma step1:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>" if "l \<in> S" "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>"
+  using that(2,1) by cases (auto dest: action1 delay1)
+
+lemma step2:
+  "B \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>" if "l \<in> S" "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>"
+  using that(2,1) by cases (auto dest: action2 delay2)
+
+lemma S_inv:
+  "l' \<in> S" if "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>" "l \<in> S"
+  using that state_set_trans_of
+  by (auto elim!: step.cases step_a.cases step_t.cases simp: state_set_def)
+
+interpretation interp: Bisimulation_Invariant
+  "\<lambda> (l, u) (l', u'). A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>"
+  "\<lambda> (l, u) (l', u'). B \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l',u'\<rangle>"
+  "\<lambda> lu lu'. lu' = lu" "\<lambda> (l, u). l \<in> S" "\<lambda> (l, u). l \<in> S"
+  by standard (auto dest: step1 step2 intro: S_inv)
+
+
+
+end (* TA Equiv *)
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/TA_Impl_Misc2.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/TA_Impl_Misc2.thy
@@ -0,0 +1,280 @@
+theory TA_Impl_Misc2
+  imports
+    Munta_Base.TA_Impl_Misc
+    "HOL-Library.Sublist"
+    "List-Index.List_Index"
+    Automatic_Refinement.Misc
+begin
+
+lemma mem_nth:
+  "x \<in> set xs \<Longrightarrow> \<exists> i < length xs. xs ! i = x"
+  by (metis index_less_size_conv nth_index)
+
+lemma union_subsetI:
+  "A \<union> B \<subseteq> C \<union> D" if "A \<subseteq> C" "B \<subseteq> D"
+  using that by blast
+
+lemma map_eq_imageD:
+  "f ` set xs = set ys" if "map f xs = ys"
+  using that by auto
+
+lemma if_contract:
+  "(if a then x else if b then x else y) = (if a \<or> b then x else y)" for a x b y
+  by (rule SMT.z3_rule)
+
+paragraph \<open>More intros\<close>
+
+named_theorems more_intros
+named_theorems more_elims
+lemmas [more_intros] =
+  image_eqI[rotated] CollectI subsetI
+
+lemmas [more_elims] =
+  CollectE
+
+paragraph \<open>Finiteness\<close>
+
+lemma finite_prodI:
+  "finite {(a,b). P a \<and> Q b}" if "finite {a. P a}" "finite {a. Q a}"
+  using that by simp
+
+lemma finite_prodI3:
+  "finite {(a,b,c). P a \<and> Q b \<and> Q1 c}"
+  if "finite {a. P a}" "finite {a. Q a}" "finite {a. Q1 a}"
+  using that by simp
+
+lemma finite_prodI4:
+  "finite {(a,b,c,d). P a \<and> Q b \<and> Q1 c \<and> Q2 d}"
+  if "finite {a. P a}" "finite {a. Q a}" "finite {a. Q1 a}" "finite {a. Q2 a}"
+  using that by simp
+
+lemma finite_prodI5:
+  "finite {(a,b,c,d,e). P a \<and> Q b \<and> Q1 c \<and> Q2 d \<and> Q3 e}"
+  if "finite {a. P a}" "finite {a. Q a}" "finite {a. Q1 a}" "finite {a. Q2 a}" "finite {a. Q3 a}"
+  using that by simp
+
+named_theorems finite_intros
+named_theorems more_finite_intros
+
+lemmas [finite_intros] =
+  finite_UnI finite_Union finite_imageI
+  finite_lists_length_eq finite_lists_length_le
+  finite_subset_distinct distinct_finite_set
+
+lemmas [more_finite_intros] =
+  finite_prodI finite_prodI3 finite_prodI4 finite_prodI5
+
+paragraph \<open>Lists\<close>
+
+lemma fold_evD2:
+  assumes
+    "P y (fold f xs acc)" "\<not> P y acc"
+    "\<And> acc x. \<not> P y acc \<Longrightarrow> Q acc \<Longrightarrow> P y (f x acc) \<Longrightarrow> x \<in> set xs \<Longrightarrow> x = y"
+    "Q acc" "\<And> acc x. Q acc \<Longrightarrow> Q (f x acc)"
+    "\<And> acc x. \<not> P y acc \<Longrightarrow> Q acc \<Longrightarrow> P y (f x acc) \<Longrightarrow> R y"
+  shows "\<exists> ys zs. xs = ys @ y # zs \<and> \<not> P y (fold f ys acc) \<and> P y (f y (fold f ys acc)) \<and> R y"
+proof -
+  from fold_evD'[OF assms(2,1)] obtain x ys zs where *:
+    "xs = ys @ x # zs" "\<not> P y (fold f ys acc)" "P y (f x (fold f ys acc))"
+    by auto
+  moreover from assms(4-) have "Q (fold f ys acc)" by (auto intro: fold_acc_preserv)
+  moreover from \<open>xs = _\<close> have "x \<in> set xs"
+    by auto
+  ultimately show ?thesis using assms(3,6) by auto
+qed
+
+lemmas fold_evD2' = fold_evD2[where R = "\<lambda> _. True", simplified]
+
+lemma filter_map_map_filter:
+  "filter P (map f xs) = List.map_filter (\<lambda> x. let y = f x in if P y then Some y else None) xs"
+  by (induction xs; simp add: Let_def List.map_filter_simps)
+
+lemma distinct_map_filterI:
+  "distinct (List.map_filter f xs)"
+  if "\<forall>x \<in> set xs. \<forall>y \<in> set xs. \<forall>a. f x = Some a \<and> f y = Some a \<longrightarrow> x = y" "distinct xs"
+  using that by (induction xs) (auto simp: map_filter_simps set_map_filter split: option.split)
+
+lemma filter_eqI:
+  assumes
+    "subseq ys xs" "\<forall>x \<in> set ys. P x"
+    "\<forall>zs. subseq zs xs \<and> length zs > length ys \<longrightarrow> (\<exists> x \<in> set zs. \<not> P x)"
+  shows "filter P xs = ys"
+  using assms
+proof (induction xs arbitrary: ys rule: list.induct)
+  case Nil
+  then show ?case
+    by - (cases ys; simp)
+next
+  case (Cons x xs)
+  show ?case
+  proof (cases "P x")
+    case True
+    show ?thesis
+    proof (cases ys)
+      case Nil
+      have "subseq [x] (x # xs)"
+        by auto
+      with Cons.prems Nil \<open>P x\<close> show ?thesis
+        by fastforce
+    next
+      case (Cons y ys')
+      have "x = y"
+      proof (rule ccontr)
+        assume "x \<noteq> y"
+        with \<open>subseq ys (x # xs)\<close> \<open>ys = _\<close> have "subseq (x # ys) (x # xs)"
+          by simp
+        with Cons.prems(2-) \<open>P x\<close> show False
+          by fastforce
+      qed
+      have "\<exists>x\<in>set zs. \<not> P x" if "subseq zs xs" and "length ys' < length zs" for zs
+      proof -
+        from \<open>subseq zs xs\<close> have "subseq (x # zs) (x # xs)"
+          by simp
+        with \<open>length ys' < length zs\<close> Cons.prems(3) \<open>ys = _\<close> have "\<exists>x\<in>set (x # zs). \<not> P x"
+          by (intro Cons.prems(3)[rule_format]; simp)
+        with \<open>P x\<close> show ?thesis
+          by auto
+      qed
+      with Cons.prems \<open>P x\<close> \<open>ys = _\<close> \<open>x = y\<close> show ?thesis
+        by (auto intro!: Cons.IH)
+    qed
+  next
+    case False
+    with Cons.prems show ?thesis
+      by (cases ys) (auto split: if_split_asm intro!: Cons.IH)
+  qed
+qed
+
+lemma filter_greatest_subseqD:
+  "\<exists> x \<in> set zs. \<not> P x" if "subseq zs xs" "length zs > length (filter P xs)"
+  using that by (metis filter_id_conv not_subseq_length subseq_filter)
+
+lemma filter_eq_iff_greatest_subseq:
+  "filter P xs = ys \<longleftrightarrow>
+  subseq ys xs \<and> (\<forall>x \<in> set ys. P x) \<and>
+  (\<forall>zs. subseq zs xs \<and> length zs > length ys \<longrightarrow> (\<exists> x \<in> set zs. \<not> P x))"
+  using filter_greatest_subseqD filter_eqI by auto
+
+lemma subseq_subsetD:
+  "set xs \<subseteq> set ys" if "subseq xs ys"
+  using that
+  by (intro subsetI) (unfold subseq_singleton_left[symmetric], erule subseq_order.trans)
+
+lemma subseq_distinct:
+  "distinct xs" if "distinct ys" "subseq xs ys"
+  using subseqs_distinctD that by simp
+
+lemma finite_subseqs:
+  "finite {xs. subseq xs ys}" (is "finite ?S")
+proof -
+  have "?S \<subseteq> {xs. set xs \<subseteq> set ys \<and> length xs \<le> length ys}"
+    using not_subseq_length by (force dest: subseq_subsetD)
+  also have "finite \<dots>"
+    by (auto intro: finite_lists_length_le)
+  finally (finite_subset) show ?thesis .
+qed
+
+lemma filter_distinct_eqI:
+  assumes
+    "subseq ys xs" "\<forall>x \<in> set ys. P x" "\<forall>x \<in> set xs. x \<notin> set ys \<longrightarrow> \<not> P x" "distinct xs"
+  shows "filter P xs = ys"
+proof (intro filter_eqI, safe)
+  fix zs assume prems: "subseq zs xs" "length ys < length zs"
+  obtain x where "x \<in> set zs" "x \<notin> set ys"
+  proof (atomize_elim, rule ccontr)
+    assume "\<nexists>x. x \<in> set zs \<and> x \<notin> set ys"
+    then have "set zs \<subseteq> set ys"
+      by auto
+    moreover from prems assms have "distinct zs" "distinct ys"
+      by (blast intro: subseq_distinct)+
+    ultimately show False
+      using \<open>length ys < length zs\<close>
+      by (auto dest: card_mono[rotated] simp: distinct_card[symmetric])
+  qed
+  with prems assms show "\<exists>x\<in>set zs. \<not> P x"
+    by (auto 4 3 dest: subseq_subsetD)
+qed (use assms in blast)+
+
+lemma subseq_sorted_wrt:
+  "sorted_wrt R xs" if "sorted_wrt R ys" "subseq xs ys"
+  using that
+  by (induction xs arbitrary: ys)
+     (auto 0 4 dest: subseq_subsetD list_emb_ConsD subseq_Cons' simp: sorted_wrt_append)
+
+lemma subseq_sorted:
+  "sorted xs" if "sorted ys" "subseq xs ys"
+  using that by (rule subseq_sorted_wrt)
+
+lemma sorted_distinct_subset_subseqI:
+  assumes "sorted xs" "distinct xs" "sorted ys" "set xs \<subseteq> set ys"
+  shows "subseq xs ys"
+  using assms
+proof (induction ys arbitrary: xs)
+  case Nil
+  then show ?case
+    by simp
+next
+  case (Cons y ys xs)
+  from Cons.prems show ?case
+    by (cases xs; simp) (safe; rule Cons.IH; auto 4 4)
+qed
+
+lemma sorted_distinct_subseq_iff:
+  assumes "sorted ys" "distinct ys"
+  shows "subseq xs ys \<longleftrightarrow> (sorted xs \<and> distinct xs \<and> set xs \<subseteq> set ys)"
+  using assms
+  by safe
+     (erule
+       subseq_subsetD[THEN subsetD] sorted_distinct_subset_subseqI subseq_distinct subseq_sorted;
+       assumption
+     )+
+
+lemma subseq_mapE:
+  assumes "subseq xs (map f ys)"
+  obtains xs' where "subseq xs' ys" "map f xs' = xs"
+  using assms
+  by (induct x1 \<equiv> xs x2 \<equiv> "map f ys" arbitrary: xs ys rule: list_emb.induct)
+     (auto, metis map_consI(1) subseq_Cons2)
+
+lemma list_all2_map_fst_aux:
+  assumes "list_all2 (\<lambda>x y. x \<in> Pair y ` (zs y)) xs ys"
+  shows "list_all2 (=) (map fst xs) ys"
+  using assms by (smt fstI imageE list.rel_mono_strong list_all2_map1)
+
+lemma list_all2_fst_aux:
+  "map fst xs = ys" if "list_all2 (\<lambda>x y. fst x = y) xs ys"
+  using that by (induction) auto
+
+text \<open>Stronger version of @{thm Map.map_of_mapk_SomeI}\<close>
+theorem map_of_mapk_SomeI':
+  assumes "inj_on f (fst ` set t)"
+    and "map_of t k = Some x"
+  shows "map_of (map (\<lambda>(k, y). (f k, g y)) t) (f k) = Some (g x)"
+  using assms
+  apply (induction t)
+   apply (solves simp)
+  apply (clarsimp split: if_split_asm)
+  apply (metis DiffI imageI img_fst map_of_SomeD singletonD)
+  done
+
+theorem map_of_mapk_SomeI:
+  assumes "inj f"
+    and "map_of t k = Some x"
+  shows "map_of (map (\<lambda>(k, y). (f k, g y)) t) (f k) = Some (g x)"
+  using assms by - (rule map_of_mapk_SomeI', erule inj_on_subset, auto)
+
+lemma list_all2_map_eq_iff:
+  "list_all2 (\<lambda>x y. f x = g y) xs ys \<longleftrightarrow> map f xs = map g ys"
+proof
+  assume "list_all2 (\<lambda>x y. f x = g y) xs ys"
+  then show "map f xs = map g ys"
+    by induction auto
+next
+  assume "map f xs = map g ys"
+  then have "length xs = length ys"
+    by (rule map_eq_imp_length_eq)
+  then show "list_all2 (\<lambda>x y. f x = g y) xs ys"
+    using \<open>map f xs = _\<close> by (induction rule: list_induct2; simp)
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Simple_Networks/TA_More2.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Simple_Networks/TA_More2.thy
@@ -0,0 +1,9 @@
+theory TA_More2
+  imports Munta_Base.TA_More
+begin
+
+lemma collect_clock_pairs_concat:
+  "collect_clock_pairs (concat xxs) = (\<Union> x \<in> set xxs. collect_clock_pairs x)"
+  unfolding collect_clock_pairs_def by auto
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl.thy
@@ -0,0 +1,4276 @@
+theory Normalized_Zone_Semantics_Impl
+  imports
+    Timed_Automata.Normalized_Zone_Semantics
+    Worklist_Algorithms.Worklist_Locales
+    TA_DBM_Operations_Impl
+    Floyd_Warshall.FW_Code
+    Munta_Base.TA_More
+begin
+
+unbundle no_library_syntax
+
+section \<open>Implementation of Reachability Checking\<close>
+
+hide_const D
+
+no_notation Extended_Nat.infinity_class.infinity ("\<infinity>")
+
+definition default_ceiling where
+  "default_ceiling A = (
+    let M = (\<lambda> c. {m . (c, m) \<in> Timed_Automata.clkp_set A}) in
+      (\<lambda> x. if M x = {} then 0 else nat (Max (M x))))"
+
+subsection \<open>Outdated Material\<close>
+
+\<comment> \<open>Unused\<close>
+definition default_ceiling_real where
+  "default_ceiling_real A = (
+    let M = (\<lambda> c. {m . (c, m) \<in> Timed_Automata.clkp_set A}) in
+      (\<lambda> x. if M x = {} then 0 else nat (floor (Max (M x)) + 1)))"
+
+\<comment> \<open>This is for automata carrying real time annotations\<close>
+\<comment> \<open>Unused\<close>
+lemma standard_abstraction_real:
+  assumes
+    "finite (Timed_Automata.clkp_set A)" "finite (Timed_Automata.collect_clkvt (trans_of A))"
+    "\<forall>(_,m::real) \<in> Timed_Automata.clkp_set A. m \<in> \<nat>"
+  shows "Timed_Automata.valid_abstraction A (clk_set A) (default_ceiling_real A)"
+proof -
+  from assms have 1: "finite (clk_set A)" by auto
+  have 2: "Timed_Automata.collect_clkvt (trans_of A) \<subseteq> clk_set A" by auto
+  from assms obtain L where L: "distinct L" "set L = Timed_Automata.clkp_set A"
+    by (meson finite_distinct_list)
+  let ?M = "\<lambda> c. {m . (c, m) \<in> Timed_Automata.clkp_set A}"
+  let ?X = "clk_set A"
+  let ?m = "map_of L"
+  let ?k = "\<lambda> x. if ?M x = {} then 0 else nat (floor (Max (?M x)) + 1)"
+  { fix c m assume A: "(c, m) \<in> Timed_Automata.clkp_set A"
+    from assms(1) have "finite (snd ` Timed_Automata.clkp_set A)" by auto
+    moreover have "?M c \<subseteq> (snd ` Timed_Automata.clkp_set A)" by force
+    ultimately have fin: "finite (?M c)" by (blast intro: finite_subset)
+    then have "Max (?M c) \<in> {m . (c, m) \<in> Timed_Automata.clkp_set A}" using Max_in A by auto
+    with assms(3) have "Max (?M c) \<in> \<nat>" by auto
+    then have "floor (Max (?M c)) = Max (?M c)" by (metis Nats_cases floor_of_nat of_int_of_nat_eq)
+    with A have *: "?k c = Max (?M c) + 1"
+    proof auto
+      fix n :: int and x :: real
+      assume "Max {m. (c, m) \<in> Timed_Automata.clkp_set A} = real_of_int n"
+      then have "real_of_int (n + 1) \<in> \<nat>"
+        using \<open>Max {m. (c, m) \<in> Timed_Automata.clkp_set A} \<in> \<nat>\<close> by auto
+      then show "real (nat (n + 1)) = real_of_int n + 1"
+        by (metis Nats_cases ceiling_of_int nat_int of_int_1 of_int_add of_int_of_nat_eq)
+    qed
+    from fin A have "Max (?M c) \<ge> m" by auto
+    moreover from A assms(3) have "m \<in> \<nat>" by auto
+    ultimately have "m \<le> ?k c" "m \<in> \<nat>" "c \<in> clk_set A" using A * by force+
+  }
+  then have "\<forall>(x, m) \<in> Timed_Automata.clkp_set A. m \<le> ?k x \<and> x \<in> clk_set A \<and> m \<in> \<nat>" by blast
+  with 1 2 have "Timed_Automata.valid_abstraction A ?X ?k" by - (standard, assumption+)
+  then show ?thesis unfolding default_ceiling_real_def by auto
+qed
+
+\<comment> \<open>Unused\<close>
+\<comment> \<open>This is for automata carrying int time annotations\<close>
+lemma standard_abstraction_int:
+  assumes
+    "finite (Timed_Automata.clkp_set A)" "finite (Timed_Automata.collect_clkvt (trans_of A))"
+    "\<forall>(_,m::int) \<in> Timed_Automata.clkp_set A. m \<in> \<nat>"
+  and "clk_set A \<subseteq> X" "finite X"
+  shows "Timed_Automata.valid_abstraction A X (default_ceiling A)"
+proof -
+  from \<open>_ \<subseteq> X\<close> have 2: "Timed_Automata.collect_clkvt (trans_of A) \<subseteq> X" by auto
+  from assms obtain L where L: "distinct L" "set L = Timed_Automata.clkp_set A"
+    by (meson finite_distinct_list)
+  let ?M = "\<lambda> c. {m . (c, m) \<in> Timed_Automata.clkp_set A}"
+  let ?X = "clk_set A"
+  let ?m = "map_of L"
+  let ?k = "\<lambda> x. if ?M x = {} then 0 else nat (Max (?M x))"
+  { fix c m assume A: "(c, m) \<in> Timed_Automata.clkp_set A"
+    from assms(1) have "finite (snd ` Timed_Automata.clkp_set A)" by auto
+    moreover have "?M c \<subseteq> (snd ` Timed_Automata.clkp_set A)" by force
+    ultimately have fin: "finite (?M c)" by (blast intro: finite_subset)
+    then have "Max (?M c) \<in> {m . (c, m) \<in> Timed_Automata.clkp_set A}" using Max_in A by auto
+    with assms(3) have "Max (?M c) \<in> \<nat>" by auto
+    with A have "?k c = nat (Max (?M c))" by auto
+    from fin A have "Max (?M c) \<ge> m" by auto
+    moreover from A assms(3) have "m \<in> \<nat>" by auto
+    ultimately have "m \<le> ?k c" "m \<in> \<nat>" "c \<in> X" using A \<open>clk_set A \<subseteq> X\<close> by force+
+  }
+  then have "\<forall>(x, m) \<in> Timed_Automata.clkp_set A. m \<le> ?k x \<and> x \<in> X \<and> m \<in> \<nat>" by blast
+  with \<open>finite X\<close> 2 have "Timed_Automata.valid_abstraction A X ?k" by - (standard; assumption)
+  then show ?thesis unfolding default_ceiling_def by auto
+qed
+
+\<comment> \<open>Outdated: Not specific enough for implementation.\<close>
+definition default_numbering where
+  "default_numbering A = (SOME v. bij_betw v A {1..card A} \<and>
+  (\<forall> c \<in> A. v c > 0) \<and>
+  (\<forall> c. c \<notin> A \<longrightarrow> v c > card A))"
+
+\<comment> \<open>Outdated: Not specific enough for implementation.\<close>
+lemma default_numbering:
+  assumes "finite S"
+  defines "v \<equiv> default_numbering S"
+  defines "n \<equiv> card S"
+  shows "bij_betw v S {1..n}" (is "?A")
+  and "\<forall> c \<in> S. v c > 0" (is "?B")
+  and "\<forall> c. c \<notin> S \<longrightarrow> v c > n" (is "?C")
+proof -
+  from standard_numbering[OF \<open>finite S\<close>] obtain v' and n' :: nat where v':
+    "bij_betw v' S {1..n'} \<and> (\<forall> c \<in> S. v' c > 0) \<and> (\<forall> c. c \<notin> S \<longrightarrow> v' c > n')"
+  by metis
+  moreover from this(1) \<open>finite S\<close> have "n' = n" unfolding n_def by (auto simp: bij_betw_same_card)
+  ultimately have "?A \<and> ?B \<and> ?C" unfolding v_def default_numbering_def
+  by - (drule someI[where x = v']; simp add: n_def)
+  then show ?A ?B ?C by auto
+qed
+
+lemma finite_ta_Regions'_int:
+  fixes A :: "('a, 'c, int, 's) ta"
+  defines "x \<equiv> SOME x. x \<notin> clk_set A"
+  assumes "finite_ta A"
+  shows "Regions' (clk_set A) (default_numbering (clk_set A)) (card (clk_set A)) x"
+proof -
+  from assms obtain x' where x': "x' \<notin> clk_set A" unfolding finite_ta_def by auto
+  then have x: "x \<notin> clk_set A" unfolding x_def by (rule someI)
+  from \<open>finite_ta A\<close> have "finite (clk_set A)" unfolding finite_ta_def by auto
+  with default_numbering[of "clk_set A"] assms have
+            "bij_betw (default_numbering (clk_set A)) (clk_set A) {1..(card (clk_set A))}"
+            "\<forall>c\<in>clk_set A. 0 < (default_numbering (clk_set A)) c"
+            "\<forall>c. c \<notin> clk_set A \<longrightarrow> (card (clk_set A)) < (default_numbering (clk_set A)) c"
+  by auto
+  then show ?thesis using x assms unfolding finite_ta_def by - (standard, auto)
+qed
+
+lemma finite_ta_RegionsD_int:
+  fixes A :: "('a, 'c, int, 's) ta"
+  assumes "finite_ta A"
+  defines "S \<equiv> clk_set A"
+  defines "v \<equiv> default_numbering S"
+  defines "n \<equiv> card S"
+  defines "x \<equiv> SOME x. x \<notin> S"
+  defines "k \<equiv> default_ceiling A"
+  shows
+    "Regions' (clk_set A) v n x" "Timed_Automata.valid_abstraction A (clk_set A) k"
+    "global_clock_numbering A v n"
+proof -
+  from standard_abstraction_int assms have k:
+    "Timed_Automata.valid_abstraction A (clk_set A) k"
+  unfolding finite_ta_def by blast
+  from finite_ta_Regions'_int[OF \<open>finite_ta A\<close>] have *: "Regions' (clk_set A) v n x" unfolding assms .
+  then interpret interp: Regions' "clk_set A" k v n x .
+  from interp.clock_numbering have "global_clock_numbering A v n" by blast
+  with * k show
+    "Regions' (clk_set A) v n x" "Timed_Automata.valid_abstraction A (clk_set A) k"
+    "global_clock_numbering A v n"
+  .
+qed
+
+subsection \<open>Misc\<close>
+
+lemma finite_project_snd:
+  assumes "finite S" "\<And> a. finite (f a ` T)"
+  shows "finite ((\<lambda> (a, b) . f a b) ` (S \<times> T))" (is "finite ?S")
+proof -
+  have "?S = {x . \<exists> a b. a \<in> S \<and> (b \<in> T \<and> x = f a b)}" by auto
+  then show ?thesis
+    using [[simproc add: finite_Collect]]
+    by (auto simp: assms)
+qed
+
+lemma list_all_upt:
+  fixes a b i :: nat
+  shows "list_all (\<lambda> x. x < b) [a..<b]"
+unfolding list_all_iff by auto
+
+lemma norm_k_cong:
+  assumes "\<forall> i \<le> n. k i = k' i"
+  shows "norm M k n = norm M k' n"
+using assms unfolding norm_def by fastforce
+
+lemma norm_upd_norm'':
+  fixes k :: "nat list"
+  assumes "length k \<ge> Suc n"
+  shows "curry (norm_upd M k n) = norm (curry M) (\<lambda> i. real (k ! i)) n"
+ apply (simp add: curry_def norm_upd_norm)
+ apply (rule norm_k_cong)
+using assms by simp
+
+lemma normalized_integral_dbms_finite':
+  assumes "length k = Suc n"
+  shows "finite {norm_upd M (k :: nat list) n | M. dbm_default (curry M) n}" (is "finite ?S")
+proof -
+  let ?k = "(\<lambda> i. k ! i)"
+  have "norm_upd M k n = uncurry (norm (curry M) ?k n)" for M
+    using assms by (subst norm_upd_norm, subst norm_k_cong) auto
+  then have
+    "?S \<subseteq> uncurry `
+      {norm (curry M) (\<lambda>i. k ! i) n | M. dbm_default (curry M) n}"
+    by auto
+  moreover have "finite \<dots>"
+    by (rule finite_imageI, rule finite_subset[rotated])
+       (rule normalized_integral_dbms_finite[where n = n and k = ?k], blast)
+  ultimately show ?thesis
+    by (auto intro: finite_subset)
+qed
+
+lemma And_commute:
+  "And A B = And B A"
+by (auto intro: min.commute)
+
+lemma FW'_diag_preservation:
+  assumes "\<forall> i \<le> n. M (i, i) \<le> 0"
+  shows "\<forall> i \<le> n. (FW' M n) (i, i) \<le> 0"
+using assms FW_diag_preservation[of n "curry M"] unfolding FW'_def by auto
+
+lemma FW_neg_diag_preservation:
+  "M i i < 0 \<Longrightarrow> i \<le> n \<Longrightarrow> (FW M n) i i < 0"
+using fw_mono[of i n i M n] by auto
+
+lemma FW'_neg_diag_preservation:
+  assumes "M (i, i) < 0" "i \<le> n"
+  shows "(FW' M n) (i, i) < 0"
+using assms FW_neg_diag_preservation[of "curry M"] unfolding FW'_def by auto
+
+lemma norm_empty_diag_preservation_int:
+  fixes k :: "nat \<Rightarrow> nat"
+  assumes "i \<le> n"
+  assumes "M i i < Le 0"
+  shows "norm M k n i i < Le 0"
+  using assms unfolding norm_def norm_diag_def by (force simp: Let_def less dest: dbm_lt_trans)
+
+lemma norm_diag_preservation_int:
+  fixes k :: "nat \<Rightarrow> nat"
+  assumes "i \<le> n"
+  assumes "M i i \<le> Le 0"
+  shows "norm M k n i i \<le> Le 0"
+  using assms unfolding norm_def norm_diag_def
+  by (force simp: Let_def less_eq dbm_le_def dest: dbm_lt_trans)
+
+lemma And_diag1:
+  assumes "A i i \<le> 0"
+  shows "(And A B) i i \<le> 0"
+using assms by (auto split: split_min)
+
+lemma And_diag2:
+  assumes "B i i \<le> 0"
+  shows "(And A B) i i \<le> 0"
+using assms by (auto split: split_min)
+
+lemma abstra_upd_diag_preservation:
+  assumes "i \<le> n" "constraint_clk ac \<noteq> 0"
+  shows "(abstra_upd ac M) (i, i) = M (i, i)"
+using assms by (cases ac) auto
+
+lemma abstr_upd_diag_preservation:
+  assumes "i \<le> n" "\<forall> c \<in> collect_clks cc. c \<noteq> 0"
+  shows "(abstr_upd cc M) (i, i) = M (i, i)"
+using assms unfolding abstr_upd_def
+by (induction cc arbitrary: M) (auto simp: abstra_upd_diag_preservation)
+
+lemma abstr_upd_diag_preservation':
+  assumes "\<forall> i \<le> n. M (i, i) \<le> 0" "\<forall> c \<in> collect_clks cc. c \<noteq> 0"
+  shows "\<forall> i \<le> n. (abstr_upd cc M) (i, i) \<le> 0"
+using assms unfolding abstr_upd_def
+by (induction cc arbitrary: M) (auto simp: abstra_upd_diag_preservation)
+
+lemma up_diag_preservation:
+  assumes "M i i \<le> 0"
+  shows "(up M) i i \<le> 0"
+  using assms unfolding up_def by (auto split: split_min)
+
+lemma reset_diag_preservation:
+  assumes "M i i \<le> 0" "d \<le> 0"
+  shows "reset M n k d i i \<le> 0"
+  using assms min_le_iff_disj unfolding neutral reset_def by auto
+
+lemma reset'_diag_preservation:
+  assumes "\<forall> i \<le> n. M i i \<le> 0" "d \<le> 0"
+  shows "\<forall> i \<le> n. reset' M n cs v d i i \<le> 0"
+  using assms
+  by (induction cs) (auto intro: reset_diag_preservation)
+
+subsection \<open>Implementation Semantics\<close>
+
+lemma FW'_out_of_bounds1:
+  assumes "i' > n"
+  shows "(FW' M n) (i', j') = M (i', j')"
+unfolding FW'_def using FW_out_of_bounds1[OF assms, where M = "curry M"] by auto
+
+lemma FW'_out_of_bounds2:
+  assumes "j' > n"
+  shows "(FW' M n) (i', j') = M (i', j')"
+unfolding FW'_def using FW_out_of_bounds2[OF assms, where M = "curry M"] by auto
+
+inductive step_impl ::
+  "('a, nat, 't :: linordered_ab_group_add, 's) ta \<Rightarrow> 's \<Rightarrow> 't DBM'
+    \<Rightarrow> nat \<Rightarrow> 'a action \<Rightarrow> 's \<Rightarrow> 't DBM' \<Rightarrow> bool"
+("_ \<turnstile>I \<langle>_, _\<rangle> \<leadsto>_,_ \<langle>_, _\<rangle>" [61,61,61] 61)
+where
+  step_t_impl:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l, FW' (abstr_upd (inv_of A l) (up_canonical_upd D n)) n\<rangle>" |
+  step_a_impl:
+    "A \<turnstile> l \<longrightarrow>g,a,r l'
+    \<Longrightarrow> A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l', FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g D) n) n r 0)) n\<rangle>"
+
+inductive_cases step_impl_cases: "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>"
+
+declare step_impl.intros[intro]
+
+lemma FW'_default:
+  assumes "dbm_default (curry M) n"
+  shows "dbm_default (curry (FW' M n)) n"
+using assms by (simp add: FW'_out_of_bounds1 FW'_out_of_bounds2)
+
+lemma abstr_upd_default:
+  assumes "dbm_default (curry M) n" "\<forall>c\<in>collect_clks cc. c \<le> n"
+  shows "dbm_default (curry (abstr_upd cc M)) n"
+using assms by (auto simp: abstr_upd_out_of_bounds1 abstr_upd_out_of_bounds2)
+
+lemma up_canonical_upd_default:
+  assumes "dbm_default (curry M) n"
+  shows "dbm_default (curry (up_canonical_upd M n)) n"
+using assms by (auto simp: up_canonical_out_of_bounds1 up_canonical_out_of_bounds2)
+
+lemma reset'_upd_default:
+  assumes "dbm_default (curry M) n" "\<forall>c\<in>set cs. c \<le> n"
+  shows "dbm_default (curry (reset'_upd M n cs d)) n"
+using assms by (auto simp: reset'_upd_out_of_bounds1 reset'_upd_out_of_bounds2)
+
+lemma FW'_int_preservation:
+  assumes "dbm_int (curry M) n"
+  shows "dbm_int (curry (FW' M n)) n"
+using FW_int_preservation[OF assms] unfolding FW'_def curry_def by auto
+
+lemma step_impl_norm_dbm_default_dbm_int:
+  assumes "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>" "dbm_default (curry D) n" "dbm_int (curry D) n"
+    "\<forall> c \<in> clk_set A. c \<le> n \<and> c \<noteq> 0" "\<forall> (_, d)  \<in> Timed_Automata.clkp_set A. d \<in> \<int>"
+  shows "dbm_default (curry D') n \<and> dbm_int (curry D') n"
+  using assms
+ apply (cases rule: step_impl.cases)
+
+  subgoal \<comment> \<open>Step is a time delay step\<close>
+  apply standard
+  apply standard
+  apply standard
+      apply safe[]
+
+    apply (simp add: FW'_out_of_bounds1)
+    apply (subst abstr_upd_out_of_bounds1[where n = n])
+    using collect_clks_inv_clk_set[of A] apply fastforce
+    apply assumption
+    apply (simp add: up_canonical_out_of_bounds1 FW'_out_of_bounds1; fail)
+
+    apply standard
+    apply safe[]
+    apply (simp add: FW'_out_of_bounds2)
+    apply (subst abstr_upd_out_of_bounds2[where n = n])
+    using collect_clks_inv_clk_set[of A] apply fastforce
+    apply assumption
+    apply (simp add: up_canonical_out_of_bounds2 FW'_out_of_bounds2; fail)
+
+    apply (simp only:)
+    apply (rule FW'_int_preservation)
+    apply (rule abstr_upd_int_preservation)
+    defer
+    apply (rule up_canonical_upd_int_preservation)
+    apply (simp add: Timed_Automata.clkp_set_def Timed_Automata.collect_clki_def; fast)+
+  done
+
+  subgoal for g a r \<comment> \<open>Step is an action step\<close>
+  apply standard
+  apply standard
+  apply standard
+  apply safe[]
+
+    apply (simp add: FW'_out_of_bounds1)
+    apply (subst abstr_upd_out_of_bounds1[where n = n])
+    using collect_clks_inv_clk_set[of A] apply fastforce
+    apply assumption
+    apply (subst reset'_upd_out_of_bounds1[where n = n])
+    apply (fastforce simp: Timed_Automata.collect_clkvt_def)
+    apply assumption
+    apply (simp add: FW'_out_of_bounds1)
+    apply (subst abstr_upd_out_of_bounds1[where n = n])
+    using collect_clocks_clk_set apply fast
+    apply assumption
+    apply (simp; fail)
+
+    apply safe[]
+    apply (simp add: FW'_out_of_bounds2)
+    apply (subst abstr_upd_out_of_bounds2[where n = n])
+    using collect_clks_inv_clk_set[of A] apply fastforce
+    apply assumption
+    apply (subst reset'_upd_out_of_bounds2[where n = n])
+    apply (fastforce simp: Timed_Automata.collect_clkvt_def)
+    apply assumption
+    apply (simp add: FW'_out_of_bounds2)
+    apply (subst abstr_upd_out_of_bounds2[where n = n])
+    using collect_clocks_clk_set apply fast
+    apply assumption
+    apply (simp; fail)
+
+    apply (simp only:)
+    apply (rule FW'_int_preservation)
+    apply (rule abstr_upd_int_preservation)
+    defer
+    apply (rule reset'_upd_int_preservation)
+    apply (rule FW'_int_preservation)
+    apply (rule abstr_upd_int_preservation)
+    apply (simp add: Timed_Automata.clkp_set_def Timed_Automata.collect_clkt_def; fast)
+    apply assumption
+    apply simp
+    apply (auto dest!: reset_clk_set; fail)
+    apply (simp add: Timed_Automata.clkp_set_def Timed_Automata.collect_clki_def; fast)
+ done
+done
+
+inductive steps_impl ::
+  "('a, nat, 't, 's) ta \<Rightarrow> 's \<Rightarrow> ('t :: linordered_ab_group_add) DBM'
+  \<Rightarrow> ('s \<Rightarrow> 't list) \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 't DBM' \<Rightarrow> bool"
+("_ \<turnstile>I \<langle>_, _\<rangle> \<leadsto>_,_* \<langle>_, _\<rangle>" [61,61,61,61] 61)
+where
+  refl: "A \<turnstile>I \<langle>l, Z\<rangle> \<leadsto>k,n* \<langle>l, Z\<rangle>" |
+  step: "A \<turnstile>I \<langle>l, Z\<rangle> \<leadsto>k,n* \<langle>l', Z'\<rangle> \<Longrightarrow> A \<turnstile>I \<langle>l', Z'\<rangle> \<leadsto>n,\<tau> \<langle>l'', Z''\<rangle>
+        \<Longrightarrow> A \<turnstile>I \<langle>l'', Z''\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l''', Z'''\<rangle>
+        \<Longrightarrow> A \<turnstile>I \<langle>l, Z\<rangle> \<leadsto>k,n* \<langle>l''', FW' (norm_upd Z''' (k l''') n) n\<rangle>"
+
+lemma steps_impl_induct[consumes 1, case_names refl step]:
+  fixes
+    A ::
+      "('a \<times>
+        (nat, 'b :: linordered_ab_group_add) acconstraint list \<times>
+        'c \<times> nat list \<times> 'a) set \<times>
+       ('a \<Rightarrow> (nat, 'b) acconstraint list)"
+
+  assumes "A \<turnstile>I \<langle>x2, x3\<rangle> \<leadsto>k,n* \<langle>x6, x7\<rangle>"
+    and "\<And>l Z. P A l Z k n l Z"
+    and
+    "\<And>l Z l' Z' l'' Z'' a l''' Z'''.
+        A \<turnstile>I \<langle>l, Z\<rangle> \<leadsto>k,n* \<langle>l', Z'\<rangle> \<Longrightarrow>
+        P A l Z k n l' Z' \<Longrightarrow>
+        A \<turnstile>I \<langle>l', Z'\<rangle> \<leadsto>n,\<tau> \<langle>l'', Z''\<rangle> \<Longrightarrow>
+        A \<turnstile>I \<langle>l'', Z''\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l''', Z'''\<rangle> \<Longrightarrow>
+        P A l Z k n l''' (FW' (norm_upd Z''' (k l''') n) n)"
+  shows "P A x2 x3 k n x6 x7"
+  using assms by (induction A\<equiv>A x2 x3 k \<equiv> k n \<equiv> n x6 x7; blast)
+
+lemma steps_impl_norm_dbm_default_dbm_int:
+  assumes "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>k,n* \<langle>l', D'\<rangle>"
+    and "dbm_default (curry D) n"
+    and "dbm_int (curry D) n"
+    and "\<forall>c\<in>clk_set A. c \<le> n \<and> c \<noteq> 0"
+    and "\<forall>(_, d)\<in>Timed_Automata.clkp_set A. d \<in> \<int>"
+    and "\<forall> l. \<forall>c\<in>set (k l). c \<in> \<int>"
+    and "\<forall> l. length (k l) = Suc n"
+  shows "l' = l \<and> D' = D \<or> (\<exists>M. D' = FW' (norm_upd M (k l') n) n \<and>
+             ((\<forall>i>n. \<forall>j. curry M i j = 0) \<and> (\<forall>j>n. \<forall>i. curry M i j = 0)) \<and> dbm_int (curry M) n)"
+using assms proof (induction)
+  case refl then show ?case by auto
+next
+  case (step A l Z k n l' Z' l'' Z'' a l''' Z''')
+  then have "
+    Z' = Z \<or>
+    (\<exists>M. Z' = FW' (norm_upd M (k l') n) n \<and>
+         ((\<forall>i>n. \<forall>j. curry M i j = 0) \<and> (\<forall>j>n. \<forall>i. curry M i j = 0)) \<and> dbm_int (curry M) n)"
+  by auto
+  then show ?case
+  proof (standard, goal_cases)
+    case 1
+    with step.prems step.hyps show ?thesis
+      apply simp
+      apply (drule step_impl_norm_dbm_default_dbm_int; simp)
+      apply (drule step_impl_norm_dbm_default_dbm_int; simp)
+      by metis
+  next
+    case 2
+    then obtain M where M:
+      "Z' = FW' (norm_upd M (k l') n) n" "dbm_default (curry M) n" "dbm_int (curry M) n"
+    by auto
+    with step.prems(3-) step.hyps show ?case
+     apply -
+     apply (drule step_impl_norm_dbm_default_dbm_int; simp)
+      using FW'_default[OF norm_upd_default, where M2 = M] apply force
+      using FW'_int_preservation[OF norm_upd_int_preservation, where M2 = M] apply force
+      apply (drule step_impl_norm_dbm_default_dbm_int; simp)
+      by metis
+  qed
+qed
+
+(* Potential naming conflict *)
+definition valid_dbm where "valid_dbm M n \<equiv> dbm_int M n \<and> (\<forall> i \<le> n. M 0 i \<le> 0)"
+
+lemma valid_dbm_pos:
+  assumes "valid_dbm M n"
+  shows "[M]v,n \<subseteq> {u. \<forall> c. v c \<le> n \<longrightarrow> u c \<ge> 0}"
+using dbm_positive assms unfolding valid_dbm_def unfolding DBM_zone_repr_def by fast
+
+definition "init_dbm = (\<lambda> (x, y). Le 0)"
+
+
+definition n_eq ("_ =_ _" [51,51,51] 50) where
+  "n_eq M n M' \<equiv> \<forall> i \<le> n. \<forall> j \<le> n. M i j = M' i j"
+
+lemma canonical_eq_upto:
+  fixes A B :: "real DBM"
+  assumes
+    "clock_numbering' v n" "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+    "canonical A n" "canonical B n"
+    "[A]v,n \<noteq> {}" "[A]v,n = [B]v,n"
+    "\<forall> i \<le> n. A i i = 0" "\<forall> i \<le> n. B i i = 0"
+  shows "A =n B"
+unfolding n_eq_def
+using assms
+apply -
+apply standard
+apply standard
+apply standard
+apply standard
+subgoal for i j
+  apply (cases "i = j")
+  apply fastforce
+  apply (rule order.antisym)
+  apply (rule DBM_canonical_subset_le; auto)
+  apply (rule DBM_canonical_subset_le; auto)
+done
+done
+
+lemma up_canonical_upd_up_canonical':
+  shows "curry (up_canonical_upd M n) =n up_canonical (curry M)"
+by (auto simp: n_eq_def intro: up_canonical_upd_up_canonical)
+
+lemma And_eqI':
+  assumes "A =n A'" "B =n B'"
+  shows "And A B =n (And A' B')"
+using assms unfolding n_eq_def by auto
+
+lemma n_eq_subst:
+  assumes "A =n B"
+  shows "(A =n C) = (B =n C)"
+using assms unfolding n_eq_def by auto
+
+lemma reset'''_reset'_upd'':
+  assumes "\<forall>c\<in>set cs. c \<noteq> 0"
+  shows "(curry (reset'_upd M n cs d)) =n (reset''' (curry M) n cs d)"
+using reset'''_reset'_upd'[OF assms] unfolding n_eq_def by auto
+
+lemma norm_eq_upto:
+  assumes "A =n B"
+  shows "norm A k n =n norm B k n"
+using assms unfolding n_eq_def by (auto simp: norm_def)
+
+
+
+lemma FW'_FW:
+  "curry (FW' M n) = FW (curry M) n"
+unfolding FW'_def by auto
+
+lemma And_eqI:
+  assumes "[A]v,n = [A1]v,n" "[B]v,n = [B1]v,n"
+  shows "[And A B]v,n = [And A1 B1]v,n"
+using assms by (simp only: And_correct[symmetric])
+
+lemma DBM_zone_repr_up_eqI:
+  assumes "clock_numbering' v n" "[A]v,n = [B]v,n"
+  shows "[up A]v,n = [up B]v,n"
+using assms DBM_up_complete'[where v = v] DBM_up_sound'[OF assms(1)] by fastforce
+
+lemma reset'_correct:
+  assumes "\<forall>c. 0 < v c \<and> (\<forall>x y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+    and "\<forall>c\<in>set cs. v c \<le> n"
+    and "\<forall>k\<le>n. 0 < k \<longrightarrow> (\<exists>c. v c = k)"
+  shows "{[cs\<rightarrow>d]u | u. u \<in> [M]v,n} = [reset' M n cs v d]v,n"
+proof safe
+  fix u
+  assume "u \<in> [M]v,n"
+  with DBM_reset'_complete[OF _ assms(1,2)] show
+    "[cs\<rightarrow>d]u \<in> [reset' M n cs v d]v,n"
+  by (auto simp: DBM_zone_repr_def)
+next
+  fix u
+  assume "u \<in> [reset' M n cs v d]v,n"
+  show "\<exists>u'. u = [cs\<rightarrow>d]u' \<and> u' \<in> [M]v,n"
+  proof (cases "[M]v,n = {}")
+    case True
+    with \<open>u \<in> _\<close> DBM_reset'_empty[OF assms(3,1,2)] show ?thesis by auto
+  next
+    case False
+    from DBM_reset'_sound[OF assms(3,1,2) \<open>u \<in> _\<close>] obtain ts where
+      "set_clocks cs ts u \<in> [M]v,n"
+    by blast
+    from DBM_reset'_resets'[OF assms(1,2)] \<open>u \<in> _\<close> \<open>_ \<noteq> {}\<close>
+    have "u = [cs \<rightarrow> d](set_clocks cs ts u)" by (fastforce simp: reset_set DBM_zone_repr_def)
+    with \<open>set_clocks _ _ _ \<in> _\<close> show ?thesis by auto
+  qed
+qed
+
+lemma DBM_zone_repr_reset'_eqI:
+  assumes "\<forall>c. 0 < v c \<and> (\<forall>x y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+    and "\<forall>c\<in>set cs. v c \<le> n"
+    and "\<forall>k\<le>n. 0 < k \<longrightarrow> (\<exists>c. v c = k)"
+    and "[A]v,n = [B]v,n"
+  shows "[reset' A n cs v d]v,n = [reset' B n cs v d]v,n"
+using assms(4) reset'_correct[OF assms(1-3)] by blast
+
+lemma up_canonical_neg_diag:
+  assumes "M i i < 0"
+  shows "(up_canonical M) i i < 0"
+using assms unfolding up_canonical_def by auto
+
+lemma up_neg_diag:
+  assumes "M i i < 0"
+  shows "(up M) i i < 0"
+using assms unfolding up_def by (auto split: split_min)
+
+lemma reset''_neg_diag:
+  fixes v :: "'c \<Rightarrow> nat"
+  assumes "\<forall> c. v c > 0" "M i i < 0" "i \<le> n"
+  shows "(reset'' M n cs v d) i i < 0"
+using reset''_diag_preservation[OF assms(1), where M = M and n = n] assms(2-) by auto
+
+lemma FW_canonical':
+  assumes "\<forall> i \<le> n. (FW M n) i i \<ge> 0"
+  shows "canonical (FW M n) n"
+  using FW_neg_cycle_detect assms
+  unfolding cycle_free_diag_equiv
+  by - (rule fw_canonical[unfolded cycle_free_diag_equiv]; fastforce)
+
+lemma FW_neg_diag_equiv:
+  assumes diag: "\<exists> i \<le> n. (FW A n) i i < 0"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and cn: "\<forall> c. v c > 0"
+      and equiv: "[A]v,n = [B]v,n"
+  shows "\<exists> i \<le> n. (FW B n) i i < 0"
+proof -
+  from assms obtain i where "(FW A n) i i < 0" "i \<le> n" by force
+  with neg_diag_empty[OF surj] FW_zone_equiv[OF surj] equiv have "[B]v,n = {}" by fastforce
+  with FW_zone_equiv[OF surj] have "[FW B n]v,n = {}" by auto
+  then obtain i where
+    "(FW B n) i i < 0" "i \<le> n"
+  using FW_detects_empty_zone[OF surj, where M = B, folded neutral] cn
+  by auto
+  then show ?thesis by auto
+qed
+
+lemma FW_dbm_zone_repr_eqI2:
+  assumes f_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (f M) i i < 0"
+      and g_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (g M) i i < 0"
+      and canonical:
+      "\<And> A B. canonical A n \<Longrightarrow> canonical B n \<Longrightarrow> \<forall> i \<le> n. A i i = 0 \<Longrightarrow> \<forall> i \<le> n. B i i = 0
+      \<Longrightarrow> [A]v,n = [B]v,n
+      \<Longrightarrow> [f A]v,n = [g B]v,n"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and cn: "\<forall> c. v c > 0"
+      and equiv: "[A]v,n = [B]v,n"
+      and diag: "\<forall> i \<le> n. A i i \<le> 0" "\<forall> i \<le> n. B i i \<le> 0"
+  shows "[f (FW A n)]v,n = [g (FW B n)]v,n"
+proof (cases "\<forall> i \<le> n. (FW A n) i i \<ge> 0")
+  case True
+  with FW_neg_diag_equiv[OF _ surj cn equiv[symmetric]] have *: "\<forall>i\<le>n. 0 \<le> (FW B n) i i" by fastforce
+  with True FW_diag_preservation[where M = A, OF diag(1)] FW_diag_preservation[where M = B, OF diag(2)]
+        FW_zone_equiv[OF surj, of A] FW_zone_equiv[OF surj, of B] equiv
+  show ?thesis by - (rule canonical[OF FW_canonical'[OF True] FW_canonical'[OF *]]; fastforce)
+next
+  case False
+  then obtain i where "(FW A n) i i < 0" "i \<le> n" by force
+  moreover with FW_neg_diag_equiv[OF _ surj cn equiv] obtain j where
+    "(FW B n) j j < 0" "j \<le> n"
+  using FW_detects_empty_zone[OF surj, where M = B, folded neutral] cn by auto
+  ultimately have "f (FW A n) i i < 0" "g (FW B n) j j < 0" using f_diag g_diag by auto
+  with \<open>i \<le> n\<close> \<open>j \<le> n\<close> neg_diag_empty[OF surj] show ?thesis by auto
+qed
+
+lemma FW_dbm_zone_repr_eqI:
+  assumes f_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (f M) i i < 0"
+      and g_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (g M) i i < 0"
+      and canonical: "\<And> M. canonical M n \<Longrightarrow> [f M]v,n = [g M]v,n"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+  shows "[f (FW M n)]v,n = [g (FW M n)]v,n"
+proof (cases "\<forall> i \<le> n. (FW M n) i i \<ge> 0")
+  case True
+  from canonical[OF FW_canonical'[OF True]] show ?thesis .
+next
+  case False
+  then obtain i where "(FW M n) i i < 0" "i \<le> n" by force
+  with f_diag g_diag have "f (FW M n) i i < 0" "g (FW M n) i i < 0" by auto
+  with \<open>i \<le> n\<close> neg_diag_empty[OF surj] show ?thesis by auto
+qed
+
+lemma FW_dbm_zone_repr_eqI':
+  assumes f_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (f M) i i < 0"
+      and g_diag: "\<And> M i. i \<le> n \<Longrightarrow> M i i < 0 \<Longrightarrow> (g M) i i < 0"
+      and canonical: "\<And> M. canonical M n \<Longrightarrow> \<forall> i \<le> n. M i i = 0 \<Longrightarrow> [f M]v,n = [g M]v,n"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and diag: "\<forall> i \<le> n. M i i \<le> 0"
+  shows "[f (FW M n)]v,n = [g (FW M n)]v,n"
+proof (cases "\<forall> i \<le> n. (FW M n) i i \<ge> 0")
+  case True
+  with FW_diag_preservation[where M = M, OF diag] canonical[OF FW_canonical'[OF True]] show ?thesis
+  by fastforce
+next
+  case False
+  then obtain i where "(FW M n) i i < 0" "i \<le> n" by force
+  with f_diag g_diag have "f (FW M n) i i < 0" "g (FW M n) i i < 0" by auto
+  with \<open>i \<le> n\<close> neg_diag_empty[OF surj] show ?thesis by auto
+qed
+
+
+subsubsection \<open>Transfer Proofs\<close>
+
+lemma conv_dbm_entry_mono:
+  assumes "a \<le> b"
+  shows "map_DBMEntry real_of_int a \<le> map_DBMEntry real_of_int b"
+using assms by (cases a; cases b) (auto simp: less_eq dbm_le_def elim!: dbm_lt.cases)
+
+lemma conv_dbm_entry_mono_strict:
+  assumes "a < b"
+  shows "map_DBMEntry real_of_int a < map_DBMEntry real_of_int b"
+using assms by (cases a; cases b) (auto simp: less elim!: dbm_lt.cases)
+
+(* Begin lifting syntax *)
+context
+  includes lifting_syntax
+begin
+
+definition "ri = (\<lambda> a b. real_of_int b = a)"
+
+abbreviation "acri \<equiv> rel_acconstraint (=) ri"
+
+abbreviation "acri' n \<equiv> rel_acconstraint (eq_onp (\<lambda> x. x < Suc n)) ri"
+
+abbreviation
+  "RI n \<equiv> (rel_prod (eq_onp (\<lambda> x. x < Suc n)) (eq_onp (\<lambda> x. x < Suc n)) ===> rel_DBMEntry ri)"
+
+lemma rel_DBMEntry_map_DBMEntry_ri [simp, intro]:
+  "rel_DBMEntry ri (map_DBMEntry real_of_int x) x"
+by (cases x) (auto simp: ri_def)
+
+lemma RI_fun_upd[transfer_rule]:
+  "(RI n ===> (=) ===> rel_DBMEntry ri ===> RI n) fun_upd fun_upd"
+unfolding rel_fun_def eq_onp_def by auto
+
+lemma min_ri_transfer[transfer_rule]:
+  "(rel_DBMEntry ri ===> rel_DBMEntry ri ===> rel_DBMEntry ri) min min"
+unfolding rel_fun_def
+  apply (simp split: split_min)
+  apply safe
+
+  subgoal for x y x' y'
+    apply (drule not_le_imp_less)
+    apply (drule conv_dbm_entry_mono_strict)
+  by (cases x; cases x'; cases y; cases y'; auto simp: ri_def; fail)
+
+  subgoal for x y x' y'
+    apply (drule not_le_imp_less)
+    apply (drule conv_dbm_entry_mono)
+  by (cases x; cases x'; cases y; cases y'; auto simp: ri_def; fail)
+
+done
+
+lemma ri_neg[transfer_rule, intro]:
+  "ri a b \<Longrightarrow> ri (-a) (-b)"
+unfolding ri_def by auto
+
+lemma abstra_upd_RI[transfer_rule]:
+  "(acri' n ===> RI n ===> RI n) abstra_upd abstra_upd"
+ apply rule
+ apply rule
+ subgoal for x y _ _
+  apply (cases x; cases y)
+ using min_ri_transfer unfolding eq_onp_def rel_fun_def by (auto dest: ri_neg)
+done
+
+lemma abstr_upd_RI[transfer_rule]:
+  "(list_all2 (acri' n) ===> RI n ===> RI n) abstr_upd abstr_upd"
+unfolding abstr_upd_def by transfer_prover
+
+lemma uminus_RI[transfer_rule]:
+  "(ri ===> ri) uminus uminus"
+unfolding ri_def by auto
+
+lemma add_RI[transfer_rule]:
+  "(ri ===> ri ===> ri) ((+) ) (+)"
+unfolding ri_def rel_fun_def by auto
+
+lemma add_rel_DBMEntry_transfer[transfer_rule]:
+  assumes R: "(A ===> B ===> C) (+) (+)"
+  shows "(rel_DBMEntry A ===> rel_DBMEntry B ===> rel_DBMEntry C) (+) (+)"
+  using R unfolding rel_fun_def[abs_def] apply safe
+  subgoal for x1 y1 x2 y2
+    by (cases x1; cases x2; cases y1; cases y2; simp add: add)
+  done
+
+lemma add_DBMEntry_RI[transfer_rule]:
+  "(rel_DBMEntry ri ===> rel_DBMEntry ri ===> rel_DBMEntry ri) ((+) ) (+)"
+  by transfer_prover
+
+lemma norm_upper_RI[transfer_rule]:
+  "(rel_DBMEntry ri ===> ri ===> rel_DBMEntry ri) norm_upper norm_upper"
+  apply (intro rel_funI)
+  subgoal for x y
+    by (cases x; cases y; fastforce simp: ri_def less[symmetric])
+  done
+
+lemma norm_lower_RI[transfer_rule]:
+  "(rel_DBMEntry ri ===> ri ===> rel_DBMEntry ri) norm_lower norm_lower"
+  apply (intro rel_funI)
+  subgoal for x y
+    by (cases x; cases y; fastforce simp: ri_def less[symmetric])
+  done
+
+lemma norm_lower_RI':
+  "(rel_DBMEntry ri ===> (=) ===> rel_DBMEntry ri) norm_lower norm_lower"
+  apply (intro rel_funI)
+  subgoal for x y
+    by (cases x; cases y; fastforce simp: ri_def less[symmetric])
+  done
+
+lemma norm_diag_RI[transfer_rule]:
+  "(rel_DBMEntry ri ===> rel_DBMEntry ri) norm_diag norm_diag"
+  unfolding norm_diag_def
+  apply (intro rel_funI)
+  subgoal for x y
+    by (cases x; cases y; fastforce simp: ri_def less[symmetric])
+  done
+
+lemma zero_RI[transfer_rule]:
+  "ri 0 0"
+by (simp add: ri_def)
+
+lemma nth_transfer[transfer_rule]:
+  fixes n :: nat
+  shows "((\<lambda> x y. list_all2 A x y \<and> length x = n) ===> eq_onp (\<lambda> x. x < n) ===> A) (!) (!)"
+by (auto simp: eq_onp_def ri_def rel_fun_def dest: list_all2_nthD)
+
+lemma nth_RI:
+  fixes n :: nat
+  shows "((\<lambda> x y. list_all2 ri x y \<and> length x = n) ===> eq_onp (\<lambda> x. x < n) ===> ri) (!) (!)"
+by (auto simp: eq_onp_def ri_def rel_fun_def dest: list_all2_nthD)
+
+lemma nth_RI':
+  fixes n :: nat
+  shows "((\<lambda> x y. list_all2 ri x y \<and> length x = n) ===> (\<lambda> x y. x = y \<and> x < n) ===> ri) (!) (!)"
+by (auto simp: ri_def rel_fun_def dest: list_all2_nthD)
+
+lemma weakening:
+  assumes "\<And> x y. B x y \<Longrightarrow> A x y" "(A ===> C) f g"
+  shows "(B ===> C) f g"
+using assms by (simp add: rel_fun_def)
+
+lemma weakening':
+  assumes "\<And> x y. B x y \<Longrightarrow> A x y" "(C ===> B) f g"
+  shows "(C ===> A) f g"
+using assms by (simp add: rel_fun_def)
+
+lemma eq_onp_Suc:
+  fixes n :: nat
+  shows "(eq_onp (\<lambda> x. x = n) ===> eq_onp (\<lambda> x. x = Suc n)) Suc Suc"
+unfolding rel_fun_def eq_onp_def by auto
+
+lemma upt_transfer_upper_bound[transfer_rule]:
+  "((=) ===> eq_onp (\<lambda> x. x = n) ===> list_all2 (eq_onp (\<lambda> x. x < n))) upt upt"
+unfolding rel_fun_def eq_onp_def apply clarsimp
+ apply (subst list.rel_eq_onp[unfolded eq_onp_def])
+unfolding list_all_iff by auto
+
+lemma zero_nat_transfer:
+  "eq_onp (\<lambda> x. x < Suc n) 0 0"
+by (simp add: eq_onp_def)
+
+lemma [transfer_rule]:
+  "bi_unique (rel_prod (eq_onp (\<lambda>x. x < Suc n)) (eq_onp (\<lambda>x. x < Suc n)))"
+unfolding bi_unique_def eq_onp_def by auto
+
+lemma [transfer_rule]:
+  "(eq_onp P ===> (=) ===> (=)) (+) (+)"
+unfolding eq_onp_def rel_fun_def by auto
+
+lemma up_canonical_upd_RI2[transfer_rule]:
+  "(RI n ===> (eq_onp (\<lambda> x. x < Suc n)) ===> RI n) up_canonical_upd up_canonical_upd"
+unfolding up_canonical_upd_def[abs_def] by transfer_prover
+
+lemma up_canonical_upd_RI[transfer_rule]:
+  "(RI n ===> (eq_onp (\<lambda> x. x = n)) ===> RI n) up_canonical_upd up_canonical_upd"
+unfolding up_canonical_upd_def[abs_def] by transfer_prover
+
+lemma up_canonical_upd_RI3[transfer_rule]:
+  "((rel_prod (=) (=) ===>
+   rel_DBMEntry (=)) ===> (eq_onp (\<lambda> x. x = n)) ===> (rel_prod (=) (=) ===>
+   rel_DBMEntry (=))) up_canonical_upd up_canonical_upd"
+unfolding up_canonical_upd_def[abs_def] by transfer_prover
+
+lemma eq_transfer:
+  "(eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> (=)) (=) (=)"
+  by (intro rel_funI; simp add: eq_onp_def)
+
+lemma norm_upd_norm:
+  "norm_upd = (\<lambda>M k n (i, j). norm (curry M) (\<lambda>i. k ! i) n i j)"
+  unfolding norm_upd_norm[symmetric] by simp
+
+lemma less_transfer:
+  "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> (=)) (<) (<)"
+  by (intro rel_funI; simp add: eq_onp_def)
+
+lemma less_eq_transfer:
+  "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x = n) ===> (=)) (\<le>) (\<le>)"
+  by (intro rel_funI; simp add: eq_onp_def)
+
+lemma norm_upd_transfer[transfer_rule]:
+  fixes n :: nat
+  notes eq_onp_Suc[of n, transfer_rule] zero_nat_transfer[transfer_rule] eq_transfer[transfer_rule]
+    less_transfer[transfer_rule] less_eq_transfer[transfer_rule]
+  shows
+    "(RI n ===> (\<lambda>x y. list_all2 ri x y \<and> length x = Suc n) ===> eq_onp (\<lambda>x. x = n)  ===> RI n)
+    norm_upd norm_upd"
+    unfolding norm_upd_norm norm_def by transfer_prover
+
+lemma dbm_entry_val_ri:
+  assumes "rel_DBMEntry ri e e'" "dbm_entry_val u c d e"
+  shows "dbm_entry_val u c d (map_DBMEntry real_of_int e')"
+using assms by (cases e; cases e') (auto simp: ri_def)
+
+lemma dbm_entry_val_ir:
+  assumes "rel_DBMEntry ri e e'" "dbm_entry_val u c d (map_DBMEntry real_of_int e')"
+  shows "dbm_entry_val u c d e"
+using assms by (cases e; cases e') (auto simp: ri_def)
+
+lemma bi_unique_eq_onp_less_Suc[transfer_rule]:
+  "bi_unique (eq_onp (\<lambda>x. x < Suc n))"
+by (simp add: eq_onp_def bi_unique_def)
+
+lemma fw_upd_transfer[transfer_rule]:
+ "((eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)
+  ===> eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n)
+  ===> (eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri))
+  fw_upd fw_upd"
+  unfolding fw_upd_def Floyd_Warshall.upd_def by transfer_prover
+
+lemma fw_upd_transfer'[transfer_rule]:
+ "(((=) ===> (=) ===> rel_DBMEntry ri)
+  ===> (=) ===> (=) ===> (=)
+  ===> ((=) ===> (=) ===> rel_DBMEntry ri))
+  fw_upd fw_upd"
+  unfolding fw_upd_def Floyd_Warshall.upd_def by transfer_prover
+
+lemma fwi_RI_transfer_aux:
+  assumes
+    "((\<lambda>x y. x < Suc n \<and> x = y) ===> (\<lambda>x y. x < Suc n \<and> x = y) ===> rel_DBMEntry ri) M M'"
+    "k < Suc n" "i < Suc n" "j < Suc n"
+  shows
+  "((\<lambda>x y. x < Suc n \<and> x = y) ===> (\<lambda>x y. x < Suc n \<and> x = y) ===> rel_DBMEntry ri)
+   (fwi M n k i j) (fwi M' n k i j)"
+using assms
+apply (induction _ "(i, j)" arbitrary: i j
+    rule: wf_induct[of "less_than <*lex*> less_than"]
+  )
+  apply (auto; fail)
+ subgoal for i j
+ apply (cases i; cases j; auto simp add: fw_upd_out_of_bounds2)
+ unfolding eq_onp_def[symmetric]
+ apply (drule rel_funD[OF fw_upd_transfer[of n]])
+ apply (auto simp: eq_onp_def dest: rel_funD; fail)
+
+ subgoal premises prems for n'
+ proof -
+  from prems have
+    "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> rel_DBMEntry ri)
+        (fwi M n k 0 n') (fwi M' n k 0 n')"
+  by auto
+  then show ?thesis
+   apply -
+   apply (drule rel_funD[OF fw_upd_transfer[of n]])
+   apply (drule rel_funD[where x = k and y = k])
+   apply (simp add: eq_onp_def \<open>k < Suc n\<close>; fail)
+   apply (drule rel_funD[where x = 0 and y = 0])
+   apply (simp add: eq_onp_def; fail)
+   apply (drule rel_funD[where x = "Suc n'" and y = "Suc n'"])
+   using prems apply (simp add: eq_onp_def; fail)
+   apply assumption
+  done
+ qed
+
+ subgoal premises prems for n'
+ proof -
+  from prems have
+    "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> rel_DBMEntry ri)
+        (fwi M n k n' n) (fwi M' n k n' n)"
+  by auto
+  then show ?thesis
+   apply -
+   apply (drule rel_funD[OF fw_upd_transfer[of n]])
+   apply (drule rel_funD[where x = k and y = k])
+   apply (simp add: eq_onp_def \<open>k < Suc n\<close>; fail)
+   apply (drule rel_funD[where x = "Suc n'" and y = "Suc n'"])
+   using prems apply (simp add: eq_onp_def; fail)
+   apply (drule rel_funD[where x = 0 and y = 0])
+   using prems apply (simp add: eq_onp_def; fail)
+   apply assumption
+  done
+ qed
+
+ subgoal premises prems for i j
+ proof -
+  from prems have
+    "(eq_onp (\<lambda>x. x < Suc n) ===> eq_onp (\<lambda>x. x < Suc n) ===> rel_DBMEntry ri)
+        (fwi M n k (Suc i) j) (fwi M' n k (Suc i) j)"
+  by auto
+  then show ?thesis
+   apply -
+   apply (drule rel_funD[OF fw_upd_transfer[of n]])
+   apply (drule rel_funD[where x = k and y = k])
+   apply (simp add: eq_onp_def \<open>k < Suc n\<close>; fail)
+   apply (drule rel_funD[where x = "Suc i" and y = "Suc i"])
+   using prems apply (simp add: eq_onp_def; fail)
+   apply (drule rel_funD[where x = "Suc j" and y = "Suc j"])
+   using prems apply (simp add: eq_onp_def; fail)
+   apply assumption
+  done
+ qed
+done
+done
+
+lemma fw_RI_transfer_aux:
+  assumes
+    "((\<lambda>x y. x < Suc n \<and> x = y) ===> (\<lambda>x y. x < Suc n \<and> x = y) ===> rel_DBMEntry ri) M M'"
+    "k < Suc n"
+  shows
+    "((\<lambda>x y. x < Suc n \<and> x = y) ===> (\<lambda>x y. x < Suc n \<and> x = y) ===> rel_DBMEntry ri)
+   (fw M n k) (fw M' n k)"
+  using assms
+  by (induction k) (auto intro: fwi_RI_transfer_aux)
+
+lemma fwi_RI_transfer[transfer_rule]:
+  "((eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)
+  ===> eq_onp (\<lambda> x. x = n) ===> eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n)
+  ===> eq_onp (\<lambda> x. x < Suc n) ===> (eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n)
+  ===> rel_DBMEntry ri)) fwi fwi"
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+by (auto intro: fwi_RI_transfer_aux simp: eq_onp_def)
+
+lemma fw_RI_transfer[transfer_rule]:
+  "((eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)
+  ===> eq_onp (\<lambda> x. x = n) ===> eq_onp (\<lambda> x. x < Suc n)
+  ===> (eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri))
+  fw fw"
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+ apply (rule rel_funI)
+by (auto intro: fw_RI_transfer_aux simp: eq_onp_def)
+
+lemma FW_RI_transfer[transfer_rule]:
+  "((eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)
+  ===> eq_onp (\<lambda> x. x = n)
+  ===> (eq_onp (\<lambda> x. x < Suc n) ===> eq_onp (\<lambda> x. x < Suc n) ===> rel_DBMEntry ri)) FW FW"
+by standard+ (drule rel_funD[OF fw_RI_transfer]; auto simp: rel_fun_def eq_onp_def)
+
+lemma FW_RI_transfer'[transfer_rule]:
+  "(RI n ===> eq_onp (\<lambda> x. x = n) ===> RI n) FW' FW'"
+using FW_RI_transfer[of n] unfolding FW'_def uncurry_def[abs_def] rel_fun_def by auto
+
+definition RI_I :: "nat \<Rightarrow> (nat, real, 's) invassn \<Rightarrow> (nat, int, 's) invassn \<Rightarrow> bool" where
+  "RI_I n \<equiv> ((=) ===> list_all2 (acri' n))"
+
+definition
+  "RI_T n \<equiv> rel_prod (=) (rel_prod (list_all2 (acri' n)) (rel_prod (=) (rel_prod (list_all2 (eq_onp (\<lambda> x. x < Suc n))) (=))))"
+
+definition RI_A :: "nat \<Rightarrow> ('a, nat, real, 's) ta \<Rightarrow> ('a, nat, int, 's) ta \<Rightarrow> bool" where
+  "RI_A n \<equiv> rel_prod (rel_set (RI_T n)) (RI_I n)"
+
+lemma inv_of_transfer [transfer_rule]:
+  "(RI_A n ===> RI_I n) inv_of inv_of"
+unfolding RI_A_def inv_of_def by transfer_prover
+
+lemma FW'_rsp:
+  "((=) ===> (=) ===> (=)) FW' FW'"
+unfolding rel_fun_def by auto
+
+lemma [transfer_rule]:
+  "(list_all2 (eq_onp (\<lambda> x. x < Suc n))) [1..n] [1..n]"
+unfolding eq_onp_def
+apply (rule list_all2I)
+(* apply (smt atLeastAtMost_iff mem_Collect_eq of_nat_Suc prod.simps(2) set_upto set_zip set_zip_rightD) *)
+apply auto
+apply (metis ab_semigroup_add_class.add.commute atLeastAtMost_iff in_set_zipE int_le_real_less of_int_less_iff of_int_of_nat_eq of_nat_Suc set_upto)
+by (simp add: zip_same)
+
+lemmas [transfer_rule] = zero_nat_transfer
+
+definition
+  "reset_canonical_upd' n M = reset_canonical_upd M n"
+
+lemma [transfer_rule]:
+  "(eq_onp (\<lambda>x. x < int (Suc n)) ===> eq_onp (\<lambda>x. x < Suc n)) nat nat"
+unfolding eq_onp_def rel_fun_def by auto
+
+lemma reset_canonical_upd_RI_aux:
+  "(RI n ===> eq_onp (\<lambda> x. x < Suc n) ===> ri ===> RI n)
+  (reset_canonical_upd' n) (reset_canonical_upd' n)"
+unfolding reset_canonical_upd'_def[abs_def] reset_canonical_upd_def[abs_def] by transfer_prover
+
+lemma reset_canonical_upd_RI[transfer_rule]:
+  "(RI n ===> eq_onp (\<lambda> x. x = n) ===> eq_onp (\<lambda> x. x < Suc n) ===> ri ===> RI n)
+  reset_canonical_upd reset_canonical_upd"
+using reset_canonical_upd_RI_aux unfolding reset_canonical_upd'_def[abs_def] rel_fun_def eq_onp_def
+by auto
+
+lemma reset'_upd_RI[transfer_rule]:
+  "(RI n ===> eq_onp (\<lambda> x. x = n) ===> list_all2 (eq_onp (\<lambda> x. x < Suc n)) ===> ri ===> RI n)
+  reset'_upd reset'_upd"
+unfolding reset'_upd_def[abs_def] by transfer_prover
+
+(* Information hiding for transfer prover *)
+definition "ri_len n = (\<lambda> x y. list_all2 ri x y \<and> length x = Suc n)"
+
+lemma RI_complete:
+  assumes lifts: "RI n D M" "RI_A n A' A"
+      and step: "A' \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l,M\<rangle> \<leadsto>n,a \<langle>l',M'\<rangle> \<and> RI n D' M'"
+using step proof cases
+  case prems: step_t_impl
+  let ?M' =
+    "FW' (abstr_upd (inv_of A l) (up_canonical_upd M n)) n"
+
+  note [transfer_rule] = lifts inv_of_transfer[unfolded RI_I_def]
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  have "RI n D' ?M'" unfolding prems by transfer_prover
+  with \<open>a = \<tau>\<close> \<open>l' = l\<close> show ?thesis by auto
+next
+  case prems: (step_a_impl g' a r')
+  obtain T I T' I' where "A = (T, I)" and "A' = (T', I')" by force
+  with lifts(2) prems(3) obtain g r where
+    "(rel_prod (=) (rel_prod (list_all2 (acri' n)) (rel_prod (=) (rel_prod (list_all2 (eq_onp (\<lambda> x. x < Suc n))) (=)))))
+     (l, g', a, r', l') (l, g, a, r, l')"
+    "(l, g, a, r, l') \<in> T"
+  unfolding RI_A_def RI_T_def rel_set_def trans_of_def by (cases rule: rel_prod.cases) fastforce
+  with \<open>A = _\<close> have g':
+    "list_all2 (acri' n) g' g" "A \<turnstile> l \<longrightarrow>g,a,r l'" "(list_all2 (eq_onp (\<lambda> x. x < Suc n))) r' r"
+  unfolding trans_of_def by (auto dest: rel_prod.cases)
+  from this(3) have "r' = r" unfolding eq_onp_def by (simp add: list_all2_eq list_all2_mono)
+
+  let ?M' = "FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g M) n) n r 0)) n"
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  note [transfer_rule] = g'[unfolded \<open>r' = r\<close>] lifts inv_of_transfer[unfolded RI_I_def]
+  have "RI n D' ?M'" unfolding prems \<open>r' = r\<close> by transfer_prover
+  with g' prems(1) show ?thesis unfolding \<open>r' = r\<close> by auto
+qed
+
+lemma IR_complete:
+  assumes lifts: "RI n D M" "RI_A n A' A"
+      and step: "A \<turnstile>I \<langle>l,M\<rangle> \<leadsto>n,a \<langle>l',M'\<rangle>"
+  shows "\<exists> D'. A' \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle> \<and> RI n D' M'"
+using step proof cases
+  case prems: step_t_impl
+  let ?D' =
+    "FW' (abstr_upd (inv_of A' l) (up_canonical_upd D n)) n"
+
+  note [transfer_rule] = lifts inv_of_transfer[unfolded RI_I_def]
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  have "RI n ?D' M'" unfolding prems by transfer_prover
+  with \<open>l' = l\<close> \<open>a = \<tau>\<close> show ?thesis by auto
+next
+  case prems: (step_a_impl g a r)
+  obtain T I T' I' where "A = (T, I)" and "A' = (T', I')" by force
+  with lifts(2) prems(3) obtain g' r' where
+    "(rel_prod (=) (rel_prod (list_all2 (acri' n)) (rel_prod (=) (rel_prod (list_all2 (eq_onp (\<lambda> x. x < Suc n))) (=)))))
+     (l, g', a, r', l') (l, g, a, r, l')"
+    "(l, g', a, r', l') \<in> T'"
+  unfolding RI_A_def RI_T_def rel_set_def trans_of_def by (cases rule: rel_prod.cases) fastforce
+  with \<open>A' = _\<close> have g':
+    "list_all2 (acri' n) g' g" "A' \<turnstile> l \<longrightarrow>g',a,r' l'" "(list_all2 (eq_onp (\<lambda> x. x < Suc n))) r' r"
+  unfolding trans_of_def by (auto dest: rel_prod.cases)
+  from this(3) have "r' = r" unfolding eq_onp_def by (simp add: list_all2_eq list_all2_mono)
+
+  let ?D' = "FW' (abstr_upd (inv_of A' l') (reset'_upd (FW' (abstr_upd g' D) n) n r 0)) n"
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  note [transfer_rule] = g'[unfolded \<open>r' = r\<close>] lifts inv_of_transfer[unfolded RI_I_def]
+  have "RI n ?D' M'" unfolding prems by transfer_prover
+  with g' prems(1) show ?thesis unfolding \<open>r' = r\<close> by auto
+qed
+
+lemma RI_norm_step:
+  assumes lifts: "RI n D M" "list_all2 ri k' k"
+      and len: "length k' = Suc n"
+  shows "RI n (FW' (norm_upd (FW' D n) k' n) n) (FW' (norm_upd (FW' M n) k n) n)"
+proof -
+  note [transfer_rule] = lifts norm_upd_transfer[folded ri_len_def]
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  have [transfer_rule]: "(ri_len n) k' k" using len lifts by (simp add: ri_len_def eq_onp_def)
+
+  show ?thesis by transfer_prover
+qed
+
+end (* End of lifting syntax *)
+
+
+
+subsubsection \<open>Semantic Equivalence\<close>
+
+lemma delay_step_impl_correct:
+  assumes "canonical (curry D) n" (* atm unused, would need for optimized variant without full FW *)
+          "clock_numbering' v n" "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> c > 0 \<and> v c \<le> n"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+  assumes D_inv: "D_inv = abstr (inv_of A l) (\<lambda>i j. \<infinity>) v"
+  shows
+  "[curry (abstr_upd (inv_of A l) (up_canonical_upd D n))]v,n =
+  [And (up (curry D)) D_inv]v,n"
+  apply (subst abstr_upd_abstr')
+   defer
+   apply (subst abstr_abstr'[symmetric])
+    defer
+  unfolding D_inv
+    apply (subst And_abstr[symmetric])
+      defer
+      defer
+      apply (rule And_eqI)
+       apply (subst DBM_up_to_equiv[folded n_eq_def, OF up_canonical_upd_up_canonical'])
+  using assms by (fastforce intro!: up_canonical_equiv_up)+
+
+lemma action_step_impl_correct:
+  assumes "canonical (curry D) n" (* atm unused, would need for optimized variant without full FW *)
+    "clock_numbering' v n" "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> c > 0 \<and> v c \<le> n"
+    "\<forall>c\<in>collect_clks g. v c = c \<and> c > 0 \<and> v c \<le> n"
+    "\<forall>c\<in> set r. v c = c \<and> c > 0 \<and> v c \<le> n"
+    "\<forall> i \<le> n. D (i, i) \<le> 0"
+    and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+  shows
+    "[curry (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g D) n) n r 0))]v,n =
+   [And (reset' (And (curry D) (abstr g (\<lambda>i j. \<infinity>) v)) n r v 0)
+                               (abstr (inv_of A l') (\<lambda>i j. \<infinity>) v)]v,n"
+  apply (subst abstr_upd_abstr', use assms in fastforce)
+  apply (subst abstr_abstr'[symmetric, where v = v], use assms in fastforce)
+  apply (subst And_abstr[symmetric], use assms in fastforce, use assms in fastforce)
+  apply (rule And_eqI[rotated], rule HOL.refl)
+  apply (subst DBM_up_to_equiv[folded n_eq_def, OF reset'''_reset'_upd''],
+      use assms in fastforce)
+  apply (subst reset''_reset'''[symmetric, where v = v], use assms in fastforce)
+  apply (subst FW'_FW)
+  apply (subst FW_dbm_zone_repr_eqI'[where g = "\<lambda> M. reset' M n r v 0"])
+       apply (rule reset''_neg_diag; fastforce simp: assms(2))
+      apply (erule DBM_reset'_neg_diag_preservation',
+        assumption, use assms(2) in fastforce, use assms in fastforce)
+     apply (erule reset'_reset''_equiv[symmetric]; use assms in fastforce)
+  using assms apply fastforce
+  subgoal
+  proof -
+    show "\<forall>i\<le>n. curry (abstr_upd g D) i i \<le> 0"
+      using assms(4) abstr_upd_diag_preservation'[OF assms(6)] by fastforce
+  qed
+  apply (rule DBM_zone_repr_reset'_eqI,
+      use assms in fastforce, use assms in fastforce, use assms in fastforce)
+  apply (subst FW_zone_equiv[symmetric], use assms in fastforce)
+  apply (subst abstr_upd_abstr', use assms in fastforce)
+  apply (subst abstr_abstr'[symmetric, where v = v], use assms in fastforce)
+  apply (rule And_abstr[symmetric]; use assms in fastforce)
+  done
+
+lemma norm_impl_correct:
+  fixes k :: "nat list"
+  assumes (* atm unused, would need for optimized variant without full FW *)
+          "clock_numbering' v n"
+          "\<forall> i \<le> n. D (i, i) \<le> 0"
+          "\<forall> i \<le> n. M i i \<le> 0"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and k: "Suc n \<le> length k"
+      and equiv: "[curry D]v,n = [M]v,n"
+  shows
+    "[curry (FW' (norm_upd (FW' D n) k n) n)]v,n = [norm (FW M n) (\<lambda> i. k ! i) n]v,n"
+ apply (subst FW'_FW)
+ apply (subst FW_zone_equiv[symmetric, OF surj])
+ apply (subst norm_upd_norm'')
+  apply (simp add: k; fail)
+ apply (subst FW'_FW)
+ subgoal
+  apply (rule FW_dbm_zone_repr_eqI2)
+  apply (rule norm_empty_diag_preservation_real[folded neutral, unfolded comp_def]; assumption)
+  apply (rule norm_empty_diag_preservation_real[folded neutral, unfolded comp_def]; assumption)
+  subgoal
+   apply (rule DBM_up_to_equiv[folded n_eq_def])
+   apply (rule norm_eq_upto)
+   apply (rule canonical_eq_upto)
+   apply (rule assms)
+   apply (rule assms)
+   apply assumption
+   apply assumption
+   using \<open>clock_numbering' v n\<close>
+   apply - apply (rule cyc_free_not_empty[OF canonical_cyc_free])
+  by simp+
+  using assms by fastforce+
+done
+
+lemma norm_action_step_impl_correct:
+  fixes k :: "nat list"
+  assumes "canonical (curry D) n" (* atm unused, would need for optimized variant without full FW *)
+          "clock_numbering' v n" "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> c > 0 \<and> v c \<le> n"
+          "\<forall>c\<in>collect_clks g. v c = c \<and> c > 0 \<and> v c \<le> n"
+          "\<forall>c\<in> set r. v c = c \<and> c > 0 \<and> v c \<le> n"
+          "\<forall> i \<le> n. D (i, i) \<le> 0"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and k: "Suc n \<le> length k"
+  shows
+  "[curry (FW' (norm_upd (FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g D) n) n r 0)) n) k n) n)]v,n =
+   [norm (FW(And (reset' (And (curry D) (abstr g (\<lambda>i j. \<infinity>) v)) n r v 0)
+                               (abstr (inv_of A l') (\<lambda>i j. \<infinity>) v)) n) (\<lambda> i. k ! i) n]v,n"
+ apply (rule norm_impl_correct)
+ apply (rule assms)
+
+ subgoal
+  apply (rule abstr_upd_diag_preservation')
+  apply safe[]
+  apply (subst reset'_upd_diag_preservation)
+  using assms(5) apply fastforce
+  apply assumption
+  apply (simp add: FW'_def)
+  apply (rule FW_diag_preservation[rule_format])
+  apply (simp add: curry_def)
+  apply (rule abstr_upd_diag_preservation'[rule_format, where n = n])
+  using assms(6) apply fastforce
+  using assms(4) apply fastforce
+  apply assumption
+  apply assumption
+ using assms(3) by fastforce
+
+ subgoal
+  apply safe[]
+  apply (rule And_diag1)
+  apply (rule DBM_reset'_diag_preservation[rule_format])
+  apply (rule And_diag1)
+  using assms(6) apply simp
+  using assms(2) apply simp
+  using assms(5) apply metis
+ by assumption
+
+ apply (rule surj; fail)
+ apply (rule k; fail)
+ apply (rule action_step_impl_correct; rule assms)
+done
+
+lemma norm_delay_step_impl_correct:
+  fixes k :: "nat list"
+  assumes "canonical (curry D) n" (* atm unused, would need for optimized variant without full FW *)
+          "clock_numbering' v n" "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> c > 0 \<and> v c \<le> n"
+          "\<forall> i \<le> n. D (i, i) \<le> 0"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and k: "Suc n \<le> length k"
+  assumes D_inv: "D_inv = abstr (inv_of A l) (\<lambda>i j. \<infinity>) v"
+  shows
+  "[curry (FW' (norm_upd (FW' (abstr_upd (inv_of A l) (up_canonical_upd D n)) n) k n) n)]v,n =
+  [norm (FW(And (up (curry D)) D_inv) n) (\<lambda> i. k ! i) n]v,n"
+ apply (rule norm_impl_correct)
+ apply (rule assms)
+
+ subgoal
+  apply (rule abstr_upd_diag_preservation')
+  apply safe[]
+  apply (subst up_canonical_upd_diag_preservation)
+  using assms(3,4) by fastforce+
+
+ subgoal
+  apply safe[]
+  apply (rule And_diag1)
+  apply (rule up_diag_preservation)
+  using assms(4) apply fastforce
+ done
+
+ apply (rule surj; fail)
+ apply (rule k; fail)
+ apply (rule delay_step_impl_correct; rule assms; fail)
+
+done
+
+lemma step_impl_sound:
+  assumes step: "A \<turnstile>I \<langle>l,M\<rangle> \<leadsto>n,a \<langle>l',M'\<rangle>"
+  assumes canonical: "canonical (curry M) n"
+  assumes numbering: "global_clock_numbering A v n" "\<forall> c \<in> clk_set A. v c = c"
+  assumes diag: "\<forall>i\<le>n. M (i, i) \<le> 0"
+  shows "\<exists> D. A \<turnstile> \<langle>l,curry M\<rangle> \<leadsto>v,n,a \<langle>l',D\<rangle> \<and> [curry M']v,n = [D]v,n"
+proof -
+  have *:
+    "\<forall>c. 0 < v c \<and> (\<forall>x y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+    "\<forall>k\<le>n. 0 < k \<longrightarrow> (\<exists>c. v c = k)" "\<forall>c\<in>clk_set A. v c \<le> n"
+  using numbering by metis+
+  from step show ?thesis
+  proof cases
+    case prems: step_t_impl
+    from *(1,3) numbering(2) collect_clks_inv_clk_set have
+      "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> 0 < v c \<and> v c \<le> n"
+    by fast
+    then have **:
+      "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> 0 < c \<and> v c \<le> n"
+    by fastforce
+    let ?D_inv = "abstr (inv_of A l) (\<lambda>i j. \<infinity>) v"
+    from prems delay_step_impl_correct[OF canonical *(1) ** *(2)] have
+      "[curry M']v,n = [And (up (curry M)) ?D_inv]v,n"
+      by (simp add: FW'_FW FW_zone_equiv[OF *(2), symmetric])
+    moreover have
+      "A \<turnstile> \<langle>l,curry M\<rangle> \<leadsto>v,n,a \<langle>l',And (up (curry M)) ?D_inv\<rangle>"
+      unfolding \<open>a = _\<close> \<open>l' = l\<close> by blast
+    ultimately show ?thesis by blast
+  next
+    case prems: (step_a_impl g a' r)
+    let ?M =
+      "And (reset' (And (curry M) (abstr g (\<lambda>i j. \<infinity>) v)) n r v 0) (abstr (inv_of A l') (\<lambda>i j. \<infinity>) v)"
+    from prems *(1,3) numbering(2) collect_clks_inv_clk_set collect_clocks_clk_set reset_clk_set have
+      "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> 0 < v c \<and> v c \<le> n"
+      "\<forall>c\<in>collect_clks g. v c = c \<and> 0 < v c \<and> v c \<le> n"
+      "\<forall>c\<in>set r. v c = c \<and> 0 < v c \<and> v c \<le> n"
+    by fast+
+    then have **:
+      "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> 0 < c \<and> v c \<le> n"
+      "\<forall>c\<in>collect_clks g. v c = c \<and> 0 < c \<and> v c \<le> n"
+      "\<forall>c\<in>set r. v c = c \<and> 0 < c \<and> v c \<le> n"
+      by fastforce+
+    from prems action_step_impl_correct[OF canonical *(1) ** diag *(2)] have
+      "[curry M']v,n = [?M]v,n"
+      by (simp add: FW'_FW FW_zone_equiv[OF *(2), symmetric])
+    moreover have
+      "A \<turnstile> \<langle>l,curry M\<rangle> \<leadsto>v,n,a \<langle>l',?M\<rangle>" unfolding \<open>a = _\<close> by - (intro step_z_norm step_a_z_dbm prems)
+    ultimately show ?thesis by auto
+  qed
+qed
+
+lemma step_impl_complete:
+  assumes step: "A \<turnstile> \<langle>l,curry D\<rangle> \<leadsto>v,n,a \<langle>l',curry D'\<rangle>"
+  assumes canonical: "canonical (curry D) n"
+  assumes numbering: "global_clock_numbering A v n" "\<forall> c \<in> clk_set A. v c = c"
+  assumes diag: "\<forall>i\<le>n. D (i, i) \<le> 0"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',M'\<rangle> \<and> [curry M']v,n = [curry D']v,n"
+proof -
+  have *:
+    "\<forall>c. 0 < v c \<and> (\<forall>x y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+    "\<forall>k\<le>n. 0 < k \<longrightarrow> (\<exists>c. v c = k)" "\<forall>c\<in>clk_set A. v c \<le> n"
+  using numbering by metis+
+  from step show ?thesis
+   apply cases
+  proof goal_cases
+    case prems: (1 D_inv)
+    from *(1,3) numbering(2) collect_clks_inv_clk_set have
+      "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> 0 < v c \<and> v c \<le> n"
+    by fast
+    then have **:
+      "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> 0 < c \<and> v c \<le> n"
+    by fastforce
+    let ?D_inv = "abstr (inv_of A l) (\<lambda>i j. \<infinity>) v"
+    let ?M' = "FW' (abstr_upd (inv_of A l) (up_canonical_upd D n)) n"
+    have
+      "[curry D']v,n = [And (up (curry D)) ?D_inv]v,n"
+      by (simp only: prems)
+    also from prems delay_step_impl_correct[OF canonical *(1) ** *(2)] have
+      "\<dots> = [curry ?M']v,n"
+      by (simp only: FW'_FW FW_zone_equiv[OF *(2), symmetric])
+    finally moreover have "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',?M'\<rangle>" by (auto simp only: \<open>l' = l\<close> \<open>a = _\<close>)
+    ultimately show ?thesis by auto
+  next
+    case prems: (2 g a' r)
+    let ?M =
+      "FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g D) n) n r 0)) n"
+    from prems *(1,3) numbering(2) collect_clks_inv_clk_set collect_clocks_clk_set reset_clk_set have
+      "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> 0 < v c \<and> v c \<le> n"
+      "\<forall>c\<in>collect_clks g. v c = c \<and> 0 < v c \<and> v c \<le> n"
+      "\<forall>c\<in>set r. v c = c \<and> 0 < v c \<and> v c \<le> n"
+    by fast+
+    then have **:
+      "\<forall>c\<in>collect_clks (inv_of A l'). v c = c \<and> 0 < c \<and> v c \<le> n"
+      "\<forall>c\<in>collect_clks g. v c = c \<and> 0 < c \<and> v c \<le> n"
+      "\<forall>c\<in>set r. v c = c \<and> 0 < c \<and> v c \<le> n"
+    by fastforce+
+    from prems action_step_impl_correct[OF canonical *(1) ** diag *(2)] have
+      "[curry D']v,n = [curry ?M]v,n"
+    by (simp only: FW'_FW FW_zone_equiv[OF *(2), symmetric])
+    moreover have "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',?M\<rangle>" using prems by auto
+    ultimately show ?thesis by auto
+  qed
+qed
+
+
+subsection \<open>Reachability Checker\<close>
+
+abbreviation conv_M :: "int DBM' \<Rightarrow> real DBM'" where "conv_M \<equiv> (o) (map_DBMEntry real_of_int)"
+
+abbreviation conv_ac :: "('a, int) acconstraint \<Rightarrow> ('a, real) acconstraint" where
+  "conv_ac \<equiv> map_acconstraint id real_of_int"
+
+abbreviation conv_cc :: "('a, int) cconstraint \<Rightarrow> ('a, real) cconstraint" where
+  "conv_cc \<equiv> map (map_acconstraint id real_of_int)"
+
+abbreviation "conv_t \<equiv> \<lambda> (l,g,a,r,l'). (l,conv_cc g,a,r,l')"
+
+definition "conv_A \<equiv> \<lambda> (T, I). (conv_t ` T, conv_cc o I)"
+
+lemma RI_zone_equiv:
+  assumes "RI n M M'"
+  shows "[curry M]v,n = [curry (conv_M M')]v,n"
+using assms unfolding DBM_zone_repr_def DBM_val_bounded_def rel_fun_def eq_onp_def
+ apply clarsimp
+ apply safe
+ apply (cases "M (0, 0)"; cases "M' (0, 0)"; fastforce simp: dbm_le_def ri_def; fail)
+ subgoal for _ c by (force intro: dbm_entry_val_ri[of "M (0, v c)"])
+ subgoal for _ c by (force intro: dbm_entry_val_ri[of "M (v c, 0)"])
+ subgoal for _ c1 c2 by (force intro: dbm_entry_val_ri[of "M (v c1, v c2)"])
+ apply (cases "M (0, 0)"; cases "M' (0, 0)"; fastforce simp: dbm_le_def ri_def; fail)
+ subgoal for _c by (rule dbm_entry_val_ir[of "M (0, v c)"]; auto)
+ subgoal for _ c by (rule dbm_entry_val_ir[of "M (v c, 0)"]; auto)
+ subgoal for _ c1 c2 by (rule dbm_entry_val_ir[of "M (v c1, v c2)"]; auto)
+ done
+
+lemma collect_clkvt_conv_A:
+  "collect_clkvt (trans_of A) = collect_clkvt (trans_of (conv_A A))"
+proof -
+  obtain T I where "A = (T, I)" by force
+  have "collect_clkvt (trans_of A) = \<Union>((set o fst \<circ> snd \<circ> snd \<circ> snd) ` T)"
+  unfolding collect_clkvt_alt_def \<open>A = _\<close> trans_of_def by simp
+  also have
+    "\<dots>= \<Union>((set o fst \<circ> snd \<circ> snd \<circ> snd) ` (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` T)"
+  by auto
+  also have "\<dots> = collect_clkvt ((\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` T)"
+  unfolding collect_clkvt_alt_def[symmetric] ..
+  also have "\<dots> = collect_clkvt (trans_of (conv_A A))" unfolding \<open>A = _\<close> trans_of_def
+    by (simp add: conv_A_def)
+  finally show ?thesis .
+qed
+
+lemma conv_ac_clock_id:
+  "constraint_pair (conv_ac ac) = (\<lambda> (a, b). (a, real_of_int b)) (constraint_pair ac)"
+by (cases ac) auto
+
+lemma collect_clock_pairs_conv_cc:
+  "collect_clock_pairs (map conv_ac cc) = (\<lambda> (a, b). (a, real_of_int b)) ` collect_clock_pairs cc"
+by (auto simp: conv_ac_clock_id collect_clock_pairs_def) (metis conv_ac_clock_id image_eqI prod.simps(2))
+
+lemma collect_clock_pairs_conv_cc':
+  fixes S :: "('a, int) acconstraint list set"
+  shows
+    "(collect_clock_pairs ` map conv_ac ` S)
+    = (((`) (\<lambda> (a, b). (a, real_of_int b))) ` collect_clock_pairs ` S)"
+ apply safe
+ apply (auto simp: collect_clock_pairs_conv_cc; fail)
+by (auto simp: collect_clock_pairs_conv_cc[symmetric])
+
+lemma collect_clock_pairs_conv_cc'':
+  fixes S :: "('a, int) acconstraint list set"
+  shows "(\<Union>x\<in>collect_clock_pairs ` map conv_ac ` S. x) = (\<Union>x\<in>collect_clock_pairs ` S. (\<lambda> (a, b). (a, real_of_int b)) ` x)"
+by (simp add: collect_clock_pairs_conv_cc')
+
+lemma clkp_set_conv_A:
+  "clkp_set (conv_A A) l = (\<lambda> (a, b). (a, real_of_int b)) ` clkp_set A l"
+  unfolding clkp_set_def collect_clki_def collect_clkt_alt_def inv_of_def trans_of_def conv_A_def
+  apply (simp only: image_Un image_Union split: prod.split)
+  by (auto simp: collect_clock_pairs_conv_cc' collect_clock_pairs_conv_cc[symmetric])
+
+lemma ta_clkp_set_conv_A:
+  "Timed_Automata.clkp_set (conv_A A) = (\<lambda> (a, b). (a, real_of_int b)) ` Timed_Automata.clkp_set A"
+ apply (simp split: prod.split add: conv_A_def)
+ unfolding
+   Timed_Automata.clkp_set_def ta_collect_clki_alt_def ta_collect_clkt_alt_def inv_of_def trans_of_def
+ apply (simp only: image_Un image_Union)
+ apply (subst collect_clock_pairs_conv_cc''[symmetric])
+ apply (subst collect_clock_pairs_conv_cc''[symmetric])
+ by fastforce
+
+lemma clkp_set_conv_A':
+  "fst ` Timed_Automata.clkp_set A = fst ` Timed_Automata.clkp_set (conv_A A)"
+by (fastforce simp: ta_clkp_set_conv_A)
+
+lemma clk_set_conv_A:
+  "clk_set (conv_A A) = clk_set A"
+  unfolding collect_clkvt_conv_A clkp_set_conv_A' ..
+
+lemma global_clock_numbering_conv:
+  assumes "global_clock_numbering A v n"
+  shows "global_clock_numbering (conv_A A) v n"
+using assms clk_set_conv_A by metis
+
+lemma real_of_int_nat:
+  assumes "a \<in> \<nat>"
+  shows "real_of_int a \<in> \<nat>"
+using assms by (metis Nats_cases of_int_of_nat_eq of_nat_in_Nats)
+
+lemma valid_abstraction_conv':
+  assumes "Timed_Automata.valid_abstraction A X' k"
+  shows "Timed_Automata.valid_abstraction (conv_A A) X' (\<lambda> x. real (k x))"
+  using assms
+  apply cases
+  apply (rule Timed_Automata.valid_abstraction.intros)
+    apply (auto intro: real_of_int_nat simp: ta_clkp_set_conv_A; fail)
+  using collect_clkvt_conv_A apply fast
+  by assumption
+
+lemma valid_abstraction_conv:
+  assumes "valid_abstraction A X' k"
+  shows "valid_abstraction (conv_A A) X' (\<lambda> l x. real (k l x))"
+  using assms
+  apply cases
+  apply (rule valid_abstraction.intros)
+     apply (auto 4 3 simp: clkp_set_conv_A intro: real_of_int_nat; fail)
+  using collect_clkvt_conv_A apply fast
+  by (auto split: prod.split_asm simp: trans_of_def conv_A_def)
+
+text \<open>Misc\<close>
+
+lemma map_nth:
+  fixes m :: nat
+  assumes "i \<le> m"
+  shows "map f [0..<Suc m] ! i = f i"
+  using assms
+by (metis add.left_neutral diff_zero le_imp_less_Suc nth_map_upt)
+
+lemma ints_real_of_int_ex:
+  assumes "z \<in> \<int>"
+  shows "\<exists>n. z = real_of_int n"
+using assms Ints_cases by auto
+
+lemma collect_clock_pairs_trans_clkp_set:
+  assumes "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  shows "collect_clock_pairs g \<subseteq> Timed_Automata.clkp_set A"
+using assms unfolding Timed_Automata.clkp_set_def Timed_Automata.collect_clkt_def by force
+
+lemma collect_clock_pairs_inv_clkp_set:
+  "collect_clock_pairs (inv_of A l) \<subseteq> Timed_Automata.clkp_set A"
+unfolding Timed_Automata.clkp_set_def Timed_Automata.collect_clki_def by force
+
+lemma finite_collect_clock_pairs[intro, simp]:
+  "finite (collect_clock_pairs x)"
+unfolding collect_clock_pairs_def by auto
+
+lemma finite_collect_clks[intro, simp]:
+  "finite (collect_clks x)"
+unfolding collect_clks_def by auto
+
+lemma check_diag_subset:
+  assumes "\<not> check_diag n D" "dbm_subset n D M"
+  shows "\<not> check_diag n M"
+using assms unfolding dbm_subset_def check_diag_def pointwise_cmp_def by fastforce
+
+(* Unused *)
+lemma dbm_int_dbm_default_convD:
+  assumes "dbm_int M n" "dbm_default M n"
+  shows "\<exists> M'. curry (conv_M M') = M"
+proof -
+  let ?unconv = "(o) (map_DBMEntry floor)"
+  let ?M' = "?unconv (uncurry M)"
+  show ?thesis
+   apply (rule exI[where x = ?M'])
+   using assms apply (intro ext)
+   apply clarsimp
+   subgoal for i j
+   by (cases "M i j"; cases "i > n"; cases "j > n";
+       fastforce dest!: leI intro: ints_real_of_int_ex simp: neutral
+     )
+  done
+qed
+
+(* Unused *)
+lemma dbm_int_all_convD:
+  assumes "dbm_int_all M"
+  shows "\<exists> M'. curry (conv_M M') = M"
+proof -
+  let ?unconv = "(o) (map_DBMEntry floor)"
+  let ?M' = "?unconv (uncurry M)"
+  show ?thesis
+   apply (rule exI[where x = ?M'])
+   using assms apply (intro ext)
+   apply clarsimp
+   subgoal for i j
+    apply (cases "M i j")
+    apply (auto intro!: ints_real_of_int_ex simp: neutral)
+    subgoal premises prems for d
+    proof -
+      from prems(2) have "M i j \<noteq> \<infinity>" by auto
+      with prems show ?thesis by fastforce
+    qed
+    subgoal premises prems for d
+    proof -
+       from prems(2) have "M i j \<noteq> \<infinity>" by auto
+       with prems show ?thesis by fastforce
+    qed
+    done
+  done
+qed
+
+lemma acri'_conv_ac:
+  assumes "fst (constraint_pair ac) < Suc n"
+  shows "acri' n (conv_ac ac) ac"
+using assms
+by (cases ac) (auto simp: ri_def eq_onp_def)
+
+lemma ri_conv_cc:
+  assumes "\<forall> c \<in> fst ` collect_clock_pairs cc. c < Suc n"
+  shows "(list_all2 (acri' n)) (conv_cc cc) cc"
+using assms
+apply -
+apply (rule list_all2I)
+apply safe
+subgoal premises prems for a b
+proof -
+  from prems(2) have "a = conv_ac b" by (metis in_set_zip nth_map prod.sel(1) prod.sel(2))
+  moreover from prems(1,2) have
+    "fst (constraint_pair b) < Suc n"
+  unfolding collect_clock_pairs_def by simp (meson set_zip_rightD)
+  ultimately show ?thesis by (simp add: acri'_conv_ac)
+qed
+by simp
+
+lemma canonical_conv_aux:
+  assumes "a \<le> b + c"
+  shows "map_DBMEntry real_of_int a \<le> map_DBMEntry real_of_int b + map_DBMEntry real_of_int c"
+  using assms unfolding less_eq add dbm_le_def
+  by (cases a; cases b; cases c) (auto elim!: dbm_lt.cases)
+
+lemma canonical_conv_aux_rev:
+  assumes "map_DBMEntry real_of_int a \<le> map_DBMEntry real_of_int b + map_DBMEntry real_of_int c"
+  shows "a \<le> b + c"
+  using assms unfolding less_eq add dbm_le_def
+  by (cases a; cases b; cases c) (auto elim!: dbm_lt.cases)
+
+(* Unused *)
+lemma canonical_conv:
+  assumes "canonical (curry M) n"
+  shows "canonical (curry (conv_M M)) n"
+  using assms by (auto intro: canonical_conv_aux)
+
+lemma canonical_conv_rev:
+  assumes "canonical (curry (conv_M M)) n"
+  shows "canonical (curry M) n"
+  using assms by (auto intro: canonical_conv_aux_rev)
+
+lemma canonical_RI_aux1:
+  assumes "rel_DBMEntry ri a1 b1" "rel_DBMEntry ri a2 b2" "rel_DBMEntry ri a3 b3" "a1 \<le> a2 + a3"
+  shows "b1 \<le> b2 + b3"
+  using assms unfolding ri_def less_eq add dbm_le_def
+  by (cases a1; cases a2; cases a3; cases b1; cases b2; cases b3) (auto elim!: dbm_lt.cases)
+
+lemma canonical_RI_aux2:
+  assumes "rel_DBMEntry ri a1 b1" "rel_DBMEntry ri a2 b2" "rel_DBMEntry ri a3 b3" "b1 \<le> b2 + b3"
+  shows "a1 \<le> a2 + a3"
+  using assms unfolding ri_def less_eq add dbm_le_def
+  by (cases a1; cases a2; cases a3; cases b1; cases b2; cases b3) (auto elim!: dbm_lt.cases)
+
+lemma canonical_RI:
+  assumes "RI n D M"
+  shows "canonical (curry D) n = canonical (curry M) n"
+  using assms unfolding rel_fun_def eq_onp_def
+  apply safe
+  subgoal for i j k
+    by (rule canonical_RI_aux1[of "D (i, k)" _ "D (i, j)" _ "D (j, k)"]; auto)
+  subgoal for i j k
+    by (rule canonical_RI_aux2[of _ "M (i, k)" _ "M (i, j)" _ "M (j, k)"]; auto)
+  done
+
+lemma RI_conv_M[transfer_rule]:
+  "(RI n) (conv_M M) M"
+  by (auto simp: rel_fun_def DBMEntry.rel_map(1) ri_def eq_onp_def DBMEntry.rel_refl)
+
+lemma canonical_conv_M_FW':
+  "canonical (curry (conv_M (FW' M n))) n = canonical (curry (FW' (conv_M M) n)) n"
+proof -
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  note [transfer_rule] = RI_conv_M
+  have 1: "RI n (FW' (conv_M M) n) (FW' M n)" by transfer_prover
+  have 2: "RI n (conv_M (FW' M n)) (FW' M n)" by (rule RI_conv_M)
+  from canonical_RI[OF 1] canonical_RI[OF 2] show ?thesis by simp
+qed
+
+lemma diag_conv:
+  assumes "\<forall> i \<le> n. (curry M) i i \<le> 0"
+  shows "\<forall> i \<le> n. (curry (conv_M M)) i i \<le> 0"
+  using assms by (auto simp: neutral dest!: conv_dbm_entry_mono)
+
+lemma map_DBMEntry_int_const:
+  assumes "get_const (map_DBMEntry real_of_int a) \<in> \<int>"
+  shows "get_const a \<in> \<int>"
+  using assms by (cases a; auto simp: Ints_def)
+
+lemma map_DBMEntry_not_inf:
+  fixes a :: "_ DBMEntry"
+  assumes "a \<noteq> \<infinity>"
+  shows "map_DBMEntry real_of_int a \<noteq> \<infinity>"
+  using assms by (cases a; auto simp: Ints_def)
+
+lemma dbm_int_conv_rev:
+  assumes "dbm_int (curry (conv_M Z)) n"
+  shows "dbm_int (curry Z) n"
+  using assms by (auto intro: map_DBMEntry_int_const dest: map_DBMEntry_not_inf)
+
+lemma neutral_RI:
+  assumes "rel_DBMEntry ri a b"
+  shows "a \<ge> 0 \<longleftrightarrow> b \<ge> 0"
+  using assms by (cases a; cases b; auto simp: neutral ri_def less_eq dbm_le_def elim!: dbm_lt.cases)
+
+lemma diag_RI:
+  assumes "RI n D M" "i \<le> n"
+  shows "D (i, i) \<ge> 0 \<longleftrightarrow> M (i, i) \<ge> 0"
+  using neutral_RI assms unfolding rel_fun_def eq_onp_def by auto
+
+lemma diag_conv_M:
+  assumes "i \<le> n"
+  shows "curry (conv_M (FW' M n)) i i \<ge> 0 \<longleftrightarrow> curry (FW' (conv_M M) n) i i \<ge> 0"
+proof -
+  have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+  note [transfer_rule] = RI_conv_M
+  have 1: "RI n (FW' (conv_M M) n) (FW' M n)" by transfer_prover
+  have 2: "RI n (conv_M (FW' M n)) (FW' M n)" by (rule RI_conv_M)
+  from \<open>i \<le> n\<close> diag_RI[OF 1] diag_RI[OF 2] show ?thesis by simp
+qed
+
+lemma conv_dbm_entry_mono_rev:
+  assumes "map_DBMEntry real_of_int a \<le> map_DBMEntry real_of_int b"
+  shows "a \<le> b"
+  using assms
+  apply (cases a; cases b)
+          apply (auto simp: less_eq dbm_le_def)
+     apply (cases rule: dbm_lt.cases, auto; fail)+
+  done
+
+lemma conv_dbm_entry_mono_strict_rev:
+  assumes "map_DBMEntry real_of_int a < map_DBMEntry real_of_int b"
+  shows "a < b"
+  using assms
+  apply (cases a; cases b)
+          apply (auto simp: less)
+     apply (cases rule: dbm_lt.cases, auto; fail)+
+  done
+
+lemma diag_conv_rev:
+  assumes "\<forall> i \<le> n. (curry (conv_M M)) i i \<le> 0"
+  shows "\<forall> i \<le> n. (curry M) i i \<le> 0"
+  using assms by (simp add: conv_dbm_entry_mono_rev neutral)
+
+lemma dbm_subset_conv:
+  assumes "dbm_subset n D M"
+  shows "dbm_subset n (conv_M D) (conv_M M)"
+  using assms unfolding dbm_subset_def pointwise_cmp_def check_diag_def
+  by (auto intro: conv_dbm_entry_mono dest!: conv_dbm_entry_mono_strict)
+
+lemma dbm_subset_conv_rev:
+  assumes "dbm_subset n (conv_M D) (conv_M M)"
+  shows "dbm_subset n D M"
+  using assms conv_dbm_entry_mono_strict_rev
+  unfolding dbm_subset_def pointwise_cmp_def check_diag_def
+  by (auto intro: conv_dbm_entry_mono_rev)
+
+(* Unused *)
+lemma dbm_subset_correct:
+  fixes D :: "real DBM'"
+  assumes "dbm_subset n D M"
+    and "canonical (curry D) n"
+    and "\<forall>i\<le>n. (curry D) i i \<le> 0"
+    and "\<forall>i\<le>n. (curry M) i i \<le> 0"
+    and "global_clock_numbering A v n"
+  shows "[curry D]v,n \<subseteq> [curry M]v,n"
+  using assms
+  apply (subst subset_eq_dbm_subset[where v= v and M' = M])
+       apply (auto; fail)
+      apply (auto; fail)
+     apply (auto; fail)
+  by blast+
+
+lemma dbm_subset_correct':
+  fixes D M :: "real DBM'"
+  assumes "canonical (curry D) n \<or> check_diag n D"
+    and "\<forall>i\<le>n. (curry D) i i \<le> 0"
+    and "\<forall>i\<le>n. (curry M) i i \<le> 0"
+    and "global_clock_numbering A v n"
+  shows "[curry D]v,n \<subseteq> [curry M]v,n \<longleftrightarrow> dbm_subset n D M"
+  using assms
+  apply -
+  apply (rule subset_eq_dbm_subset[OF assms(1)])
+     apply (auto; fail)
+    apply (auto; fail)
+  by blast+
+
+lemma step_z_dbm_diag_preservation:
+  assumes "step_z_dbm A l D v n a l' D'" "\<forall> i \<le> n. D i i \<le> 0"
+  shows "\<forall> i \<le> n. D' i i \<le> 0"
+  using assms
+  apply cases
+  using And_diag1 up_diag_preservation apply blast
+  by (metis And_diag1 order_mono_setup.refl reset'_diag_preservation)
+
+context AlphaClosure
+begin
+
+lemma step_z_dbm_mono:
+  "\<exists> D'. A \<turnstile> \<langle>l, D\<rangle> \<leadsto>v,n,a \<langle>l', D'\<rangle> \<and> [M']v,n \<subseteq> [D']v,n" if
+  gcn: "global_clock_numbering A v n" and that: "A \<turnstile> \<langle>l, M\<rangle> \<leadsto>v,n,a \<langle>l', M'\<rangle>" "[M]v,n \<subseteq> [D]v,n"
+  using step_z_mono[of A] step_z_dbm_sound[OF _ gcn] step_z_dbm_DBM[OF _ gcn] that by metis
+
+end
+
+lemma FW_canonical:
+  "canonical (FW M n) n \<or> (\<exists> i \<le> n. (FW M n) i i < 0)"
+  using FW_canonical' leI by blast
+
+lemma FW'_canonical:
+  "canonical (curry (FW' M n)) n \<or> (\<exists> i \<le> n. (FW' M n) (i, i) < 0)"
+  by (metis FW'_FW FW_canonical curry_def)
+
+lemma fw_upd_conv_M'':
+  "fw_upd (map_DBMEntry real_of_int \<circ>\<circ> M) k i j = map_DBMEntry real_of_int \<circ>\<circ> fw_upd M k i j"
+  unfolding fw_upd_def Floyd_Warshall.upd_def
+  by (rule ext; simp split: prod.split; cases "M i j"; cases "M i k"; cases "M k j")
+     (force simp: add min_def | force simp: min_inf_r | force simp: min_inf_l)+
+
+lemma fw_upd_conv_M':
+  "conv_M (uncurry (fw_upd M k i j)) = uncurry (fw_upd (map_DBMEntry real_of_int \<circ>\<circ> M) k i j)"
+  unfolding fw_upd_conv_M'' by auto
+
+lemma fw_upd_conv_M:
+  "uncurry (fw_upd (curry (conv_M (uncurry M))) h i j)
+  = (map_DBMEntry real_of_int \<circ>\<circ> uncurry) (fw_upd M h i j)"
+  unfolding fw_upd_conv_M' by (auto simp: curry_def comp_def)
+
+lemma fwi_conv_M'':
+  "map_DBMEntry real_of_int \<circ>\<circ> fwi M n k i j = fwi (map_DBMEntry real_of_int \<circ>\<circ> M) n k i j"
+  apply (induction _ "(i, j)" arbitrary: i j
+    rule: wf_induct[of "less_than <*lex*> less_than"]
+  )
+  apply (auto; fail)
+  subgoal for i j
+    by (cases i; cases j; auto simp add: fw_upd_conv_M''[symmetric])
+  done
+
+lemma fwi_conv_M':
+  "conv_M (uncurry (fwi M n k i j)) = uncurry (fwi (map_DBMEntry real_of_int \<circ>\<circ> M) n k i j)"
+  unfolding fwi_conv_M''[symmetric] by auto
+
+lemma fwi_conv_M:
+  "conv_M (uncurry (fwi (curry M) n k i j)) = uncurry (fwi (curry (conv_M M)) n k i j)"
+  unfolding fwi_conv_M' by (auto simp: curry_def comp_def)
+
+lemma fw_conv_M'':
+  "map_DBMEntry real_of_int \<circ>\<circ> fw M n k = fw (map_DBMEntry real_of_int \<circ>\<circ> M) n k"
+  by (induction k; simp only: fw.simps fwi_conv_M'')
+
+lemma fw_conv_M':
+  "conv_M (uncurry (fw M n k)) = uncurry (fw (map_DBMEntry real_of_int \<circ>\<circ> M) n k)"
+  unfolding fw_conv_M''[symmetric] by auto
+
+lemma fw_conv_M:
+  "conv_M (uncurry (fw (curry M) n k)) = uncurry (fw (curry (conv_M M)) n k)"
+  unfolding fw_conv_M' by (auto simp: curry_def comp_def)
+
+lemma FW_conv_M:
+  "uncurry (FW (curry (conv_M M)) n) = conv_M (uncurry (FW (curry M) n))"
+  using fw_conv_M by metis
+
+lemma FW'_conv_M:
+  "FW' (conv_M M) n = conv_M (FW' M n)"
+  using FW_conv_M unfolding FW'_def by simp
+
+lemma (in Regions) v_v':
+  "\<forall> c \<in> X. v' (v c) = c"
+  using clock_numbering unfolding v'_def by auto
+
+definition
+  "subsumes n
+  = (\<lambda> (l, M) (l', M'). check_diag n M \<or> l = l' \<and> pointwise_cmp (\<le>) n (curry M) (curry M'))"
+
+lemma subsumes_simp_1:
+  "subsumes n (l, M) (l', M') = dbm_subset n M M'" if "l = l'"
+  using that unfolding subsumes_def dbm_subset_def by simp
+
+lemma subsumes_simp_2:
+  "subsumes n (l, M) (l', M') = check_diag n M" if "l \<noteq> l'"
+  using that unfolding subsumes_def dbm_subset_def by simp
+
+
+lemma TA_clkp_set_unfold:
+  "Timed_Automata.clkp_set A = \<Union> (clkp_set A ` UNIV)"
+  unfolding clkp_set_def Timed_Automata.clkp_set_def
+  unfolding Timed_Automata.collect_clki_def Timed_Automata.collect_clkt_def
+  unfolding collect_clki_def collect_clkt_def
+  by blast
+
+locale TA_Start_No_Ceiling_Defs =
+  fixes A :: "('a, nat, int, 's) ta"
+    and l0 :: 's
+    and n :: nat
+begin
+
+definition "X \<equiv> {1..n}"
+
+(* Tailored towards the needs of the specific implementation semantics *)
+definition "v i \<equiv> if i > 0 \<and> i \<le> n then i else (Suc n)"
+definition "x \<equiv> SOME x. x \<notin> X"
+
+definition "a0 = (l0, init_dbm)"
+
+abbreviation
+  "canonical' D \<equiv> canonical (curry D) n"
+
+abbreviation
+  "canonical_diag' D \<equiv> canonical' D \<or> check_diag n D"
+
+abbreviation
+  "canonical_diag D \<equiv> canonical' (conv_M D) \<or> check_diag n D"
+
+abbreviation
+  "canonical_subs' I M \<equiv> canonical_subs n I (curry M)"
+
+definition "locations \<equiv> {l0} \<union> fst ` trans_of A \<union> (snd o snd o snd o snd) ` trans_of A"
+
+lemma X_alt_def:
+  "X = {1..<Suc n}"
+  unfolding X_def by auto
+
+lemma v_bij:
+  "bij_betw v {1..<Suc n} {1..n}"
+  unfolding v_def[abs_def] bij_betw_def inj_on_def by auto
+
+lemma triv_numbering:
+  "\<forall> i \<in> {1..n}. v i = i"
+  unfolding v_def by auto
+
+lemma canonical_diagI:
+  "canonical_diag D"  if "canonical_diag' D"
+  using that canonical_conv by auto
+
+(* This used (by used theorems) *)
+lemma v_id:
+  "v c = c" if "v c \<le> n"
+  by (metis Suc_n_not_le_n that v_def)
+
+end
+
+
+locale TA_Start_No_Ceiling = TA_Start_No_Ceiling_Defs +
+  assumes finite_trans[intro, simp]: "finite (trans_of A)"
+      and finite_inv[intro]: "finite (range (inv_of A))"
+      and clocks_n: "clk_set A \<subseteq> {1..n}"
+      and consts_nats: "\<forall>(_, d) \<in> Timed_Automata.clkp_set A. d \<in> \<nat>"
+      and n_gt0[intro, simp]: "n > 0"
+begin
+
+lemma clock_range:
+  "\<forall>c\<in>clk_set A. c \<le> n \<and> c \<noteq> 0"
+  using clocks_n by force
+
+lemma clk_set_X:
+  "clk_set A \<subseteq> X"
+  unfolding X_def using clocks_n .
+
+lemma finite_X:
+  "finite X"
+  unfolding X_def by (metis finite_atLeastAtMost)
+
+lemma finite_clkp_set_A[intro, simp]:
+  "finite (Timed_Automata.clkp_set A)"
+  unfolding Timed_Automata.clkp_set_def ta_collect_clkt_alt_def ta_collect_clki_alt_def by fast
+
+lemma finite_collect_clkvt[intro]:
+  "finite (collect_clkvt (trans_of A))"
+  unfolding collect_clkvt_def using [[simproc add: finite_Collect]] by auto
+
+lemma consts_ints:
+  "\<forall>(_, d) \<in> Timed_Automata.clkp_set A. d \<in> \<int>"
+  using consts_nats unfolding Nats_def by auto
+
+lemma n_bounded:
+  "\<forall> c \<in> X. c \<le> n"
+  unfolding X_def by auto
+
+lemma finite_locations:
+  "finite locations"
+  unfolding locations_def using finite_trans by auto
+
+lemma RI_I_conv_cc:
+  "RI_I n (conv_cc o snd A) (snd A)"
+  using clocks_n
+  unfolding RI_I_def rel_fun_def Timed_Automata.clkp_set_def ta_collect_clki_alt_def inv_of_def
+  by (force intro: ri_conv_cc)
+
+lemma triv_numbering':
+  "\<forall> c \<in> clk_set A. v c = c"
+  using triv_numbering clocks_n by auto
+
+lemma triv_numbering'':
+  "\<forall> c \<in> clk_set (conv_A A). v c = c"
+  using triv_numbering' unfolding clk_set_conv_A .
+
+lemma global_clock_numbering:
+  "global_clock_numbering A v n"
+  using clocks_n unfolding v_def by auto
+
+lemmas global_clock_numbering' = global_clock_numbering_conv[OF global_clock_numbering]
+
+lemma ri_conv_t:
+  assumes "t \<in> fst A"
+  shows "(RI_T n) (conv_t t) t"
+  unfolding RI_T_def apply (auto split: prod.split)
+   apply (rule ri_conv_cc)
+  using assms clocks_n unfolding Timed_Automata.clkp_set_def ta_collect_clkt_alt_def trans_of_def
+   apply fastforce
+  using assms using clocks_n unfolding clkp_set_def collect_clkvt_alt_def trans_of_def eq_onp_def
+  by (fastforce intro: list_all2I simp: zip_same)
+
+lemma RI_T_conv_t:
+  "rel_set (RI_T n) (conv_t ` fst A) (fst A)"
+  using ri_conv_t unfolding rel_set_def by (fastforce split: prod.split)
+
+lemma RI_A_conv_A:
+  "RI_A n (conv_A A) A"
+  using RI_T_conv_t RI_I_conv_cc unfolding RI_A_def conv_A_def by (auto split: prod.split)
+
+lemma step_impl_diag_preservation:
+  assumes step: "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+    and diag: "\<forall> i \<le> n. (curry M) i i \<le> 0"
+  shows
+    "\<forall> i \<le> n. (curry M') i i \<le> 0"
+proof -
+  have FW':
+    "(FW' M n) (i, i) \<le> 0" if "i \<le> n" "\<forall> i \<le> n. M (i, i) \<le> 0" for i and M :: "int DBM'"
+    using that FW'_diag_preservation[OF that(2)] diag by auto
+  have *:
+    "\<forall> c \<in> collect_clks (inv_of A l). c \<noteq> 0"
+    using clock_range collect_clks_inv_clk_set[of A l] by auto
+  from step diag * show ?thesis
+    apply cases
+
+    subgoal \<comment> \<open>delay step\<close>
+      apply simp
+      apply (rule FW'_diag_preservation)
+      apply (rule abstr_upd_diag_preservation')
+       apply (subst up_canonical_upd_diag_preservation)
+      by auto
+
+    subgoal \<comment> \<open>action step\<close>
+      apply simp
+      apply (rule FW'_diag_preservation)
+      apply (rule abstr_upd_diag_preservation')
+       apply (standard, standard)
+       apply (subst reset'_upd_diag_preservation)
+         defer
+         apply assumption
+        apply (erule FW')
+        apply (erule abstr_upd_diag_preservation')
+        apply (metis (no_types) clock_range collect_clocks_clk_set subsetCE)
+       apply (metis (no_types) clock_range collect_clks_inv_clk_set subsetCE)
+      apply (drule reset_clk_set)
+      using clock_range by blast
+    done
+qed
+
+lemma step_impl_neg_diag_preservation:
+  assumes step: "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+    and le: "i \<le> n"
+    and diag: "(curry M) i i < 0"
+  shows "(curry M') i i < 0"
+  using assms
+  apply cases
+
+  subgoal \<comment> \<open>delay step\<close>
+    apply simp
+    apply (rule FW'_neg_diag_preservation)
+     apply (subst abstr_upd_diag_preservation)
+       apply assumption
+      apply (metis (no_types) clock_range collect_clks_inv_clk_set subsetCE)
+     apply (subst up_canonical_upd_diag_preservation)
+    by auto
+
+  subgoal \<comment> \<open>action step\<close>
+    apply simp
+    apply (rule FW'_neg_diag_preservation)
+     apply (subst abstr_upd_diag_preservation)
+       apply assumption
+      apply (metis (no_types) clock_range collect_clks_inv_clk_set subsetCE)
+     apply (subst reset'_upd_diag_preservation)
+      defer
+      apply assumption
+     apply (rule FW'_neg_diag_preservation)
+      apply (subst abstr_upd_diag_preservation)
+        apply assumption
+       apply (metis (no_types) clock_range collect_clocks_clk_set subsetCE)
+      apply assumption+
+    apply (drule reset_clk_set)
+    using clock_range by blast
+  done
+
+lemma step_impl_canonical:
+  assumes step: "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+    and diag: "\<forall> i \<le> n. (curry M) i i \<le> 0"
+  shows
+    "canonical (curry (conv_M M')) n \<or> (\<exists> i \<le> n. M' (i, i) < 0)"
+proof -
+  from step_impl_diag_preservation[OF assms] have diag: "\<forall>i\<le>n. curry M' i i \<le> 0" .
+  from step obtain M'' where "M' = FW' M'' n" by cases auto
+  show ?thesis
+  proof (cases "\<exists> i \<le> n. M' (i, i) < 0")
+    case True
+    then show ?thesis by auto
+  next
+    case False
+    with diag have "\<forall>i\<le>n. curry M' i i \<ge> 0" by auto
+    then have
+      "\<forall>i\<le>n. curry (conv_M M') i i \<ge> 0"
+      unfolding neutral by (auto dest!: conv_dbm_entry_mono)
+    with FW_canonical'[of n "curry (conv_M M'')"] canonical_conv_M_FW' diag_conv_M show ?thesis
+      unfolding \<open>M' = _\<close> FW'_FW[symmetric] canonical_conv_M_FW' by auto
+  qed
+qed
+
+lemma step_impl_int_preservation:
+  assumes step: "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and  int: "dbm_int (curry D) n"
+  shows "dbm_int (curry D') n"
+  using assms
+  apply cases
+
+  subgoal premises prems
+    unfolding prems
+    apply (intro
+        FW'_int_preservation norm_upd_int_preservation abstr_upd_int_preservation[rotated]
+        up_canonical_upd_int_preservation
+        )
+    using consts_ints assms unfolding Timed_Automata.clkp_set_def Timed_Automata.collect_clki_def
+    by auto
+
+  subgoal premises prems
+    unfolding prems
+    apply (intro
+        FW'_int_preservation norm_upd_int_preservation abstr_upd_int_preservation[rotated]
+        reset'_upd_int_preservation
+        )
+    using assms prems(4) consts_ints collect_clock_pairs_inv_clkp_set[of A l']
+        apply (auto dest!: collect_clock_pairs_trans_clkp_set)
+    using prems(4) clock_range by (auto dest!: reset_clk_set)
+  done
+
+lemma check_diag_empty_spec:
+  assumes "check_diag n M"
+  shows "[curry M]v,n = {}"
+  apply (rule check_diag_empty)
+  using assms global_clock_numbering by fast+
+
+lemma init_dbm_semantics:
+  assumes "u \<in> [curry (conv_M init_dbm)]v,n" "0 < c" "c \<le> n"
+  shows "u c = 0"
+proof -
+  from assms(2,3) have "v c \<le> n" unfolding v_def by auto
+  with assms(1) show ?thesis unfolding DBM_zone_repr_def init_dbm_def DBM_val_bounded_def v_def
+    by auto
+qed
+
+lemma dbm_int_conv:
+  "dbm_int (curry (conv_M Z)) n"
+  apply safe
+  subgoal for i j
+    by (cases "Z (i, j)") auto
+  done
+
+lemma RI_init_dbm:
+  "RI n init_dbm init_dbm"
+  unfolding init_dbm_def rel_fun_def ri_def by auto
+
+lemma conv_M_init_dbm[simp]:
+  "conv_M init_dbm = init_dbm"
+  unfolding init_dbm_def by auto
+
+lemmas
+  step_z_dbm_equiv' = step_z_dbm_equiv[OF global_clock_numbering']
+
+lemma step_impl_complete':
+  assumes step: "step_z_dbm (conv_A A) l (curry (conv_M M)) v n a l' M'"
+    and canonical: "canonical (curry (conv_M M)) n"
+    and diag: "\<forall>i\<le>n. conv_M M (i, i) \<le> 0"
+  shows "\<exists> D'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle> \<and> [curry (conv_M D')]v,n = [M']v,n"
+proof -
+  let ?A = "conv_A A"
+  have A: "RI_A n ?A A" by (rule RI_A_conv_A)
+  have M_conv: "RI n (conv_M M) M" unfolding eq_onp_def by auto
+  let ?M' = "uncurry M'"
+  from
+    step_impl_complete[of ?A l "conv_M M" _ _ a _ ?M',
+      OF _ canonical global_clock_numbering' triv_numbering'' diag
+      ] step
+  obtain M'' where M'':
+    "?A \<turnstile>I \<langle>l, conv_M M\<rangle> \<leadsto>n,a \<langle>l', M''\<rangle>" "[curry M'']v,n = [curry ?M']v,n"
+    by auto
+  from RI_complete[OF M_conv A this(1)] obtain MM where MM:
+    "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', MM\<rangle>" "RI n M'' MM"
+    by auto
+  moreover from MM(2) M''(2) M''(2) have
+    "[M']v,n = [curry (conv_M MM)]v,n"
+    by (auto dest!: RI_zone_equiv[where v = v])
+  ultimately show ?thesis by auto
+qed
+
+lemma step_impl_check_diag:
+  assumes "check_diag n M" "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+  shows "check_diag n M'"
+  using step_impl_neg_diag_preservation assms unfolding check_diag_def neutral by auto
+
+lemmas
+  step_z_dbm_sound = step_z_dbm_sound[OF _ global_clock_numbering']
+
+lemmas
+  step_z_dbm_DBM = step_z_dbm_DBM[OF _ global_clock_numbering']
+
+lemmas dbm_subset_correct' = dbm_subset_correct'[OF _ _ _ global_clock_numbering]
+
+subsubsection \<open>Bisimilarity\<close>
+
+definition dbm_equiv (infixr "\<simeq>" 60) where
+  "dbm_equiv M M' \<equiv> [curry (conv_M M)]v,n = [curry (conv_M M')]v,n"
+
+definition state_equiv (infixr "\<sim>" 60) where
+  "state_equiv \<equiv> \<lambda> (l, M) (l', M'). l = l' \<and> M \<simeq> M'"
+
+lemma dbm_equiv_trans[intro]:
+  "a \<simeq> c" if "a \<simeq> b" "b \<simeq> c"
+  using that unfolding dbm_equiv_def by simp
+
+lemma state_equiv_trans[intro]:
+  "a \<sim> c" if "a \<sim> b" "b \<sim> c"
+  using that unfolding state_equiv_def by blast
+
+lemma dbm_equiv_refl[intro]:
+  "a \<simeq> a"
+  unfolding dbm_equiv_def by simp
+
+lemma state_equiv_refl[intro]:
+  "a \<sim> a"
+  unfolding state_equiv_def by blast
+
+lemma dbm_equiv_sym:
+  "a \<simeq> b" if "b \<simeq> a"
+  using that unfolding dbm_equiv_def by simp
+
+lemma state_equiv_sym:
+  "a \<sim> b" if "b \<sim> a"
+  using that unfolding state_equiv_def by (auto intro: dbm_equiv_sym)
+
+lemma state_equiv_D:
+  "M \<simeq> M'" if "(l, M) \<sim> (l', M')"
+  using that unfolding state_equiv_def by auto
+
+
+lemma finite_trans':
+    "finite (trans_of (conv_A A))"
+    using finite_trans unfolding trans_of_def conv_A_def by (cases A) auto
+
+lemma init_dbm_semantics':
+    assumes "u \<in> [(curry init_dbm :: real DBM)]v,n"
+    shows "\<forall> c \<in> {1..n}. u c = 0"
+  using assms init_dbm_semantics by auto
+
+  lemma init_dbm_semantics'':
+    assumes "\<forall> c \<in> {1..n}. u c = 0"
+    shows "u \<in> [(curry init_dbm :: real DBM)]v,n"
+  unfolding DBM_zone_repr_def DBM_val_bounded_def init_dbm_def using assms by (auto simp: v_def)
+
+end (* TA Start No Ceiling *)
+
+
+locale TA_Start_Defs =
+  TA_Start_No_Ceiling_Defs A l0 n for A :: "('a, nat, int, 's) ta" and l0 n +
+  fixes k :: "'s \<Rightarrow> nat \<Rightarrow> nat"
+begin
+
+definition "k' l \<equiv> map (int o k l) [0..<Suc n]"
+
+definition "E = (\<lambda> (l, M) (l'', M''').
+  \<exists> M' l' M'' a.
+  A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle>
+  \<and> A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>
+  \<and> M''' = FW' (norm_upd M'' (k' l'') n) n)"
+
+lemma E_alt_def: "E = (\<lambda> (l, M) (l', M''').
+  \<exists> g a r M' M''.
+  A \<turnstile> l \<longrightarrow>g,a,r l' \<and>
+  M' = FW' (abstr_upd (inv_of A l) (up_canonical_upd M n)) n \<and>
+  M'' = FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g M') n) n r 0)) n \<and>
+  M''' = FW' (norm_upd M'' (k' l') n) n
+  )"
+  unfolding E_def by (force elim!: step_impl.cases)
+
+  lemma length_k':
+    "length (k' l) = Suc n"
+  unfolding k'_def by auto
+
+end
+
+
+locale Reachability_Problem_No_Ceiling_Defs =
+  TA_Start_No_Ceiling_Defs A l0 n for A :: "('a, nat, int, 's) ta" and l0 n+
+  fixes F :: "'s \<Rightarrow> bool"
+begin
+
+definition "F_rel \<equiv> \<lambda> (l, M). F l \<and> \<not> check_diag n M"
+
+end
+
+(* locale Reachability_Problem_no_ceiling = TA_Start + Reachability_Problem_no_ceiling_Defs *)
+
+locale Reachability_Problem_Defs = TA_Start_Defs + Reachability_Problem_No_Ceiling_Defs
+begin
+
+end (* Reachability Problem Defs *)
+
+lemma check_diag_conv_M[intro]:
+  assumes "check_diag n M"
+  shows "check_diag n (conv_M M)"
+  using assms unfolding check_diag_def by (auto dest!: conv_dbm_entry_mono_strict)
+
+lemma canonical_variant:
+  "canonical (curry (conv_M D)) n \<or> check_diag n D
+  \<longleftrightarrow> canonical (curry (conv_M D)) n \<or> (\<exists>i\<le>n. D (i, i) < 0)"
+  unfolding check_diag_def neutral ..
+
+locale TA_Start =
+  TA_Start_No_Ceiling A l0 n +
+  TA_Start_Defs A l0 n k for A :: "('a, nat, int, 's) ta" and l0 n k +
+  assumes k_ceiling:
+    "\<forall> l. \<forall>(x,m) \<in> clkp_set A l. m \<le> k l x"
+    "\<forall> l g a r l' c. A \<turnstile> l \<longrightarrow>g,a,r l' \<and> c \<notin> set r \<longrightarrow> k l' c \<le> k l c"
+  and k_bound: "\<forall> l. \<forall> i > n. k l i = 0"
+  and k_0: "\<forall> l. k l 0 = 0"
+begin
+
+lemma E_closure:
+  "E** a0 (l', M') \<longleftrightarrow> A \<turnstile>I \<langle>l0, init_dbm\<rangle> \<leadsto>k',n* \<langle>l', M'\<rangle>"
+  unfolding a0_def E_def
+  apply safe
+   apply (drule rtranclp_induct[where P = "\<lambda> (l', M'). A \<turnstile>I \<langle>l0, init_dbm\<rangle> \<leadsto>k',n* \<langle>l', M'\<rangle>"];
+      auto dest: step intro: refl; fail
+      )
+  apply (induction rule: steps_impl.induct; simp add: rtranclp.intros(2))
+  by (rule rtranclp.intros(2)) auto
+
+lemma FW'_normalized_integral_dbms_finite':
+  "finite {FW' (norm_upd M (k' l) n) n | M. dbm_default (curry M) n}"
+  (is "finite ?S")
+proof -
+  let ?S' = "{norm_upd M (k' l) n | M. dbm_default (curry M) n}"
+  have "finite ?S'"
+    using normalized_integral_dbms_finite'[where k = "map (k l) [0..<Suc n]"] by (simp add: k'_def)
+  moreover have "?S = (\<lambda> M. FW' M n) ` ?S'" by auto
+  ultimately show ?thesis by auto
+qed
+
+lemma E_closure_finite:
+  "finite {x. E** a0 x}"
+proof -
+  have k': "map int (map (k l) [0..<Suc n]) = k' l" for l unfolding k'_def by auto
+  have *: "(l, M) = a0 \<or> (\<exists>M'. M = FW' (norm_upd M' (k' l) n) n \<and> dbm_default (curry M') n)"
+    if "E** a0 (l, M)" for l M
+    using that unfolding E_closure
+    apply -
+    apply (drule steps_impl_norm_dbm_default_dbm_int)
+          apply (auto simp: init_dbm_def neutral)[]
+         apply (auto simp: init_dbm_def)[]
+        defer
+        defer
+        apply (simp add: k'_def; fail)
+       apply (simp add: k'_def; fail)
+      apply (simp add: a0_def)
+    using clock_range consts_ints by (auto simp: X_def)
+  moreover have **: "l \<in> locations" if "E** a0 (l, M)" for l M
+    using that unfolding E_closure locations_def
+    apply (induction rule: steps_impl.induct)
+     apply (simp; fail)
+    by (force elim!: step_impl.cases)
+  have
+    "{x. E** a0 x} \<subseteq>
+        {a0} \<union> (locations \<times> {FW' (norm_upd M (k' l) n) n | l M. l \<in> locations \<and> dbm_default (curry M) n})"
+    (is "_ \<subseteq> _ \<union> ?S")
+    apply safe
+     apply (rule **, assumption)
+    apply (frule *)
+    by (auto intro: **)
+  moreover have "finite ?S"
+    using FW'_normalized_integral_dbms_finite' finite_locations
+    using [[simproc add: finite_Collect]]
+    by (auto simp: k'_def finite_project_snd)
+  ultimately show ?thesis by (auto intro: finite_subset)
+qed
+
+sublocale Regions "{1..<Suc n}" v n k "Suc n"
+  apply standard
+       apply (simp; fail)
+  using default_numbering(2)[OF finite_X] apply (subst (asm) X_def) apply (simp add: v_def; fail)
+  using default_numbering(2)[OF finite_X] apply (subst (asm) X_def) apply (simp add: v_def; fail)
+  by (auto simp: v_def)
+
+lemma k_simp_1:
+  "(\<lambda> l i. if i \<le> n then map (k l) [0..<Suc n] ! i else 0) = k"
+proof (rule ext)+
+  fix l i
+  show "(if i \<le> n then map (k l) [0..<Suc n] ! i else 0) = k l i"
+  proof (cases "i \<le> n")
+    case False
+    with k_bound show ?thesis by auto
+  next
+    case True
+    with map_nth[OF this] show ?thesis by auto
+  qed
+qed
+
+lemma k_simp_2:
+  "(\<lambda> l. k l \<circ> v') = k"
+proof (rule ext)+
+  fix l i
+  show "(\<lambda> l. k l o v') l i = k l i"
+  proof (cases "i > n")
+    case True
+    then show ?thesis unfolding v'_def by (auto simp: k_bound)
+  next
+    case False
+    show ?thesis
+    proof (cases "i = 0")
+      case True
+      with k_0 have "k l i = 0" by simp
+      moreover have "v' 0 = Suc n" unfolding v'_def by auto
+      ultimately show ?thesis using True by (auto simp: k_bound)
+    next
+      case False
+      with \<open>\<not> n < i\<close> have "v' (v i) = i" using v_v' by auto
+      moreover from False \<open>\<not> n < i\<close> triv_numbering have "v i = i" by auto
+      ultimately show ?thesis by auto
+    qed
+  qed
+qed
+
+lemma valid_abstraction:
+  "valid_abstraction A X k"
+  using k_ceiling consts_nats clk_set_X unfolding X_def
+  by (fastforce intro!: valid_abstraction.intros simp: TA_clkp_set_unfold)
+
+lemma norm_upd_diag_preservation:
+  assumes "i \<le> n" "M (i, i) \<le> 0"
+  shows "(norm_upd M (k' l) n) (i, i) \<le> 0"
+  apply (subst norm_upd_norm)
+  apply (subst norm_k_cong[where k' = "k l"])
+   apply (safe; simp only: k'_def map_nth; simp; fail)
+  using norm_diag_preservation_int assms unfolding neutral by auto
+
+lemma norm_upd_neg_diag_preservation:
+  assumes "i \<le> n" "M (i, i) < 0"
+  shows "(norm_upd M (k' l) n) (i, i) < 0"
+  apply (subst norm_upd_norm)
+  apply (subst norm_k_cong[where k' = "k l"])
+   apply (safe; simp only: k'_def map_nth; simp; fail)
+  using norm_empty_diag_preservation_int assms unfolding neutral by auto
+
+lemmas valid_abstraction' = valid_abstraction_conv[OF valid_abstraction, unfolded X_alt_def]
+
+lemma step_impl_V_preservation_canonical:
+  assumes step: "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and canonical: "canonical (curry (conv_M D)) n"
+    and diag: "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+    and valid: "valid_dbm (curry (conv_M D))"
+  shows "[curry (conv_M D')]v,n \<subseteq> V"
+proof -
+  let ?A = "conv_A A"
+  have A: "RI_A n ?A A" by (rule RI_A_conv_A)
+  have "RI n (conv_M D) D" unfolding eq_onp_def by auto
+  from IR_complete[OF this A step] obtain M' where M':
+    "?A \<turnstile>I \<langle>l, conv_M D\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>" "RI n M' D'"
+    by auto
+  from
+    step_impl_sound[of ?A l "conv_M D",
+      OF this(1) canonical global_clock_numbering' triv_numbering'' diag(1)
+      ]
+  obtain M'' where M'':
+    "?A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,a \<langle>l', M''\<rangle>" "[curry M']v,n = [M'']v,n"
+    by (auto simp add: k_simp_1)
+  from step_z_valid_dbm[OF M''(1) global_clock_numbering' valid_abstraction'] valid have
+    "valid_dbm M''"
+    by auto
+  then have "[M'']v,n \<subseteq> V" by cases
+  with M''(2) RI_zone_equiv[OF M'(2)] show ?thesis by auto
+qed
+
+lemma step_impl_V_preservation:
+  assumes step: "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and canonical: "canonical (curry (conv_M D)) n \<or> (\<exists>i\<le>n. D (i, i) < 0)"
+    and diag: "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+    and valid: "valid_dbm (curry (conv_M D))"
+  shows "[curry (conv_M D')]v,n \<subseteq> V"
+proof -
+  from canonical show ?thesis
+  proof (standard, goal_cases)
+    case 1
+    from step_impl_V_preservation_canonical[OF step this diag valid] show ?thesis .
+  next
+    case 2
+    with step_impl_neg_diag_preservation[OF step] have "\<exists>i\<le>n. D' (i, i) < 0" by auto
+    then have
+      "[curry (conv_M D')]v,n = {}"
+      by - (rule check_diag_empty_spec;
+          auto dest!: conv_dbm_entry_mono_strict simp: check_diag_def neutral
+          )
+    then show ?thesis by blast
+  qed
+qed
+
+lemma norm_step_diag_preservation:
+  "\<forall>i\<le>n. curry (FW' (norm_upd D (k' l) n) n) i i \<le> 0" if "\<forall>i\<le>n. (curry D) i i \<le> 0"
+  by (metis FW'_diag_preservation curry_conv norm_upd_diag_preservation that)
+
+lemma norm_step_check_diag_preservation:
+  "check_diag n (FW' (norm_upd D (k' l) n) n)" if "check_diag n D"
+  by (metis FW'_neg_diag_preservation check_diag_def neutral norm_upd_neg_diag_preservation that)
+
+lemma diag_reachable:
+  assumes "E** a0 (l, M)"
+  shows "\<forall> i \<le> n. (curry M) i i \<le> 0"
+  using assms unfolding E_closure
+  apply (induction Z \<equiv> "init_dbm :: int DBM'" _ _ _ _ rule: steps_impl_induct)
+   apply (auto simp: init_dbm_def neutral; fail)
+  apply (assumption | rule norm_step_diag_preservation step_impl_diag_preservation)+
+  done
+
+lemma canonical_norm_step:
+  "canonical (curry (conv_M (FW' (norm_upd M (k' l) n) n))) n
+   \<or> (\<exists> i \<le> n. (FW' (norm_upd M (k' l) n) n) (i, i) < 0)"
+  by (metis FW'_canonical canonical_conv)
+
+lemma canonical_reachable:
+  assumes "E** a0 (l, M)"
+  shows "canonical (curry (conv_M M)) n \<or> (\<exists>i \<le> n. M (i, i) < 0)"
+  using assms unfolding E_closure
+proof (induction l \<equiv> l0 Z \<equiv> "init_dbm :: int DBM'" _ _ _ _ rule: steps_impl_induct)
+  case refl
+  show ?case unfolding init_dbm_def by simp (simp add: neutral[symmetric])
+next
+  case step
+  show ?case by (rule canonical_norm_step)
+qed
+
+lemma diag_reachable':
+  assumes "E** a0 (l, M)"
+  shows "\<forall> i \<le> n. (curry (conv_M M)) i i \<le> 0"
+  using diag_reachable[OF assms] by (auto simp: neutral dest!: conv_dbm_entry_mono)
+
+context
+    fixes l' :: 's
+  begin
+
+  interpretation regions: Regions_global "{1..<Suc n}" v n "k l'" "Suc n"
+    by (standard; rule finite clock_numbering not_in_X non_empty)
+
+  lemma v'_v'[simp]:
+    "Regions.v' {1..<Suc n} v n (Suc n) = Beta_Regions'.v' {1..<Suc n} v n (Suc n)"
+    unfolding v'_def regions.beta_interp.v'_def ..
+
+  lemma k_simp_2': "(\<lambda> x. real ((k l' \<circ>\<circ>\<circ> Beta_Regions'.v' {1..<Suc n} v) n (Suc n) x)) = k l'"
+    using k_simp_2 v'_v' by metis
+
+  lemma norm_upd_V_preservation:
+    "[curry (conv_M (norm_upd M (k' l') n))]v,n \<subseteq> V"
+    if "[curry (conv_M M)]v,n \<subseteq> V" "canonical (curry (conv_M M)) n"
+  proof -
+    from regions.beta_interp.norm_V_preservation[OF that] have
+      "[norm (curry (conv_M M)) (\<lambda> x. real (k l' x)) n]v,n \<subseteq> V"
+      by (simp only: k_simp_2')
+    then have *: "[norm (curry (conv_M M)) ((!) (map real_of_int (k' l'))) n]v,n \<subseteq> V"
+      apply (subst norm_k_cong[of _ _ "(\<lambda> x. real (k l' x))"])
+       apply safe
+      (* s/h *)
+      proof -
+        fix i :: nat
+        assume "i \<le> n"
+        then have "map real_of_int (k' l') ! i = (real_of_int \<circ>\<circ>\<circ> (\<circ>)) int (k l') i"
+          by (metis (no_types) Normalized_Zone_Semantics_Impl.map_nth k'_def map_map)
+        then show "map real_of_int (k' l') ! i = real (k l' i)"
+          by simp
+      qed
+
+    note [transfer_rule] = norm_upd_transfer[folded ri_len_def]
+    have [transfer_rule]: "eq_onp (\<lambda>x. x = n) n n" by (simp add: eq_onp_def)
+    have [transfer_rule]: "(ri_len n) (map real_of_int (k' l')) (k' l')"
+      by (simp add: ri_len_def eq_onp_def length_k' list_all2_conv_all_nth ri_def)
+    have "RI n (norm_upd (conv_M M) (map real_of_int (k' l')) n) (norm_upd M (k' l') n)"
+      by transfer_prover
+    from RI_zone_equiv[OF this] * have
+      "[curry (conv_M (norm_upd M (k' l') n))]v,n \<subseteq> V"
+      by (auto simp add: norm_upd_norm'[symmetric])
+    then show ?thesis .
+  qed
+
+  lemma norm_step_correct:
+    assumes diag: "\<forall>i\<le>n. conv_M D (i, i) \<le> 0" "\<forall>i\<le>n. M i i \<le> 0"
+        and canonical: "canonical (curry (conv_M D)) n \<or> (\<exists>i\<le>n. D (i, i) < 0)"
+        and equiv: "[curry (conv_M D)]v,n = [M]v,n"
+        and valid: "valid_dbm M"
+    shows
+    "[curry (conv_M (FW' (norm_upd D (k' l') n) n))]v,n
+   = [norm (FW M n) (\<lambda>x. real (k l' x)) n]v,n"
+  proof (cases "check_diag n D")
+    case False
+    let ?k = "map real_of_int (k' l')"
+    have k_alt_def: "?k = map (k l') [0..<Suc n]" unfolding k'_def by auto
+    have k: "list_all2 ri ?k (k' l')" by (simp add: list_all2_conv_all_nth ri_def)
+    from canonical False have "canonical (curry (conv_M D)) n"
+      unfolding check_diag_def neutral by simp
+    from canonical_conv_rev[OF this] have "canonical (curry D) n"
+      by simp
+    from FW_canonical_id[OF this] have "FW (curry D) n = curry D" .
+    then have "FW' D n = D"
+      apply (subst (asm) FW'_FW[symmetric])
+      unfolding curry_def by (rule ext, safe, meson)
+    have "RI n (FW' (norm_upd (FW' (conv_M D) n) (map real_of_int (k' l')) n) n)
+       (FW' (norm_upd (FW' D n) (k' l') n) n)"
+      by (intro RI_norm_step RI_conv_M k; simp add: length_k')
+    moreover have "[curry (conv_M D)]v,n = [M]v,n" by (rule equiv)
+    moreover have "\<forall> i \<le> n. (\<lambda>x. real (map (k l') [0..<Suc n] ! x)) i = (\<lambda>x. real (k l' x)) i"
+      using map_nth by auto
+    ultimately show ?thesis
+      using diag
+      apply -
+      apply (subst \<open>_ = D\<close>[symmetric])
+      apply (subst RI_zone_equiv[symmetric, where M = "FW' (norm_upd (FW' (conv_M D) n) ?k n) n"])
+       apply assumption
+      unfolding k_alt_def
+      using norm_impl_correct[where D = "conv_M D", where M = M]
+      apply (subst norm_impl_correct[where M = M])
+            prefer 4
+      using global_clock_numbering apply satx
+      using global_clock_numbering apply satx
+          apply (simp; fail)+
+      apply (subst norm_k_cong[of _ "(\<lambda>x. real (k l' x))"])
+      by auto
+  next
+    case True
+    then have "check_diag n (conv_M D)" by auto
+    from check_diag_empty_spec[OF this] equiv have *:
+      "[curry (conv_M D)]v,n = {}" "[M]v,n = {}"
+      by auto
+    from regions.norm_FW_empty[OF valid *(2)] k_simp_2' have **:
+      "[norm (FW M n) (\<lambda>x. real (k l' x)) n]v,n = {}"
+      by auto
+    from norm_step_check_diag_preservation[OF True] have
+      "check_diag n (FW' (norm_upd D (k' l') n) n)" .
+    then have "check_diag n (conv_M (FW' (norm_upd D (k' l') n) n))" by auto
+    from check_diag_empty_spec[OF this] have
+      "[curry (conv_M (FW' (norm_upd D (k' l') n) n))]v,n = {}" .
+    with ** step' show ?thesis by auto
+  qed
+
+  lemma step_z_norm'_empty_preservation:
+    assumes "step_z_norm' (conv_A A) l D a l' D'" "valid_dbm D" "[D]v,n = {}"
+    shows "[D']v,n = {}"
+    using assms(1)
+    apply cases
+    unfolding v'_v'
+    apply simp
+    apply (rule regions.norm_FW_empty[simplified])
+    using assms(2) step_z_valid_dbm[OF _ global_clock_numbering' valid_abstraction'] apply (simp; fail)
+    apply (rule step_z_dbm_empty[OF global_clock_numbering'])
+    using assms(3) by auto
+
+  lemma canonical_empty_check_diag:
+    assumes "canonical (curry (conv_M D')) n \<or> (\<exists>i\<le>n. D' (i, i) < 0)" "[curry (conv_M D')]v,n = {}"
+    shows "check_diag n D'"
+  proof -
+    from assms(1)
+    show ?thesis
+    proof
+      assume "canonical (curry (conv_M D')) n"
+      from regions.beta_interp.canonical_empty_zone_spec[OF this] assms(2) have
+        "\<exists>i\<le>n. curry (conv_M D') i i < 0"
+      by fast
+      with conv_dbm_entry_mono_strict_rev show ?thesis unfolding check_diag_def neutral by force
+    next
+      assume "\<exists>i\<le>n. D' (i, i) < 0"
+      then show ?thesis unfolding check_diag_def neutral by auto
+    qed
+  qed
+
+  end (* End of context for global set of regions *)
+
+lemma FW'_valid_preservation:
+  assumes "valid_dbm (curry (conv_M M))"
+  shows "valid_dbm (curry (conv_M (FW' M n)))"
+proof -
+  from FW_valid_preservation[OF assms]
+  show ?thesis by (simp add: FW'_FW[symmetric] FW'_conv_M)
+qed
+
+lemma norm_step_valid_dbm:
+  "valid_dbm (curry (conv_M (FW' (norm_upd M (k' l') n) n)))" if
+  "[curry (conv_M M)]v,n \<subseteq> V" "canonical (curry (conv_M M)) n \<or> (\<exists> i \<le> n. M (i, i) < 0)"
+proof -
+  let ?M1 = "curry (conv_M (norm_upd M (k' l') n))"
+  have "dbm_int ?M1 n" by (rule dbm_int_conv)
+  from that(2) have "[?M1]v,n \<subseteq> V"
+  proof
+    assume "canonical (curry (conv_M M)) n"
+    from that(1) norm_upd_V_preservation[OF _ this] show "[?M1]v,n \<subseteq> V" by auto
+  next
+    assume "\<exists>i\<le>n. M (i, i) < 0"
+    then have "[?M1]v,n = {}"
+      by (intro check_diag_empty_spec, simp add: check_diag_def)
+         (metis DBMEntry.simps(15) conv_dbm_entry_mono_strict neutral norm_upd_neg_diag_preservation
+            of_int_0)
+    then show ?thesis by simp
+  qed
+  with \<open>dbm_int ?M1 n\<close> have "valid_dbm ?M1" by - rule
+  from FW'_valid_preservation[OF this] show ?thesis .
+qed
+
+lemma valid_dbm_V:
+  "[M]v,n \<subseteq> V" if "valid_dbm M"
+  using that by cases
+
+lemma step_impl_valid_dbm:
+  "valid_dbm (curry (conv_M M'))" if
+  "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>" "valid_dbm (curry (conv_M M))"
+  "canonical (curry (conv_M M)) n \<or> (\<exists>i\<le>n. M (i, i) < 0)" "\<forall>i\<le>n. conv_M M (i, i) \<le> 0"
+  apply (rule valid_dbm.intros)
+  subgoal
+    using that by - (erule step_impl_V_preservation; assumption)
+  subgoal
+    by (rule dbm_int_conv)
+  done
+
+lemma valid_dbm_reachable:
+  assumes "E** a0 (l, M)"
+  shows "valid_dbm (curry (conv_M M))"
+  using assms unfolding E_closure
+  apply (induction l \<equiv> "l0" Z \<equiv> "init_dbm :: int DBM'" _ _ _ _ rule: steps_impl_induct)
+  subgoal
+    apply (rule valid_dbm.intros)
+    using init_dbm_semantics unfolding V_def apply force
+    by (auto simp: init_dbm_def ; fail)
+
+  subgoal
+    apply (rule norm_step_valid_dbm)
+     apply (rule valid_dbm_V)
+
+     apply (rule step_impl_valid_dbm, assumption)
+       apply (rule step_impl_valid_dbm, assumption)
+         apply assumption
+        apply (rule canonical_reachable[unfolded E_closure], assumption)
+       apply (drule diag_conv[OF diag_reachable, unfolded E_closure]; simp; fail)
+      apply (rule step_impl_canonical step_impl_diag_preservation, assumption)
+      apply (drule diag_conv[OF diag_reachable, unfolded E_closure],
+        simp add: conv_dbm_entry_mono_rev neutral; fail)
+
+    subgoal
+      apply (drule step_impl_diag_preservation)
+       apply (drule diag_conv[OF diag_reachable, unfolded E_closure])
+       apply (auto simp: conv_dbm_entry_mono_rev neutral)[]
+      apply (auto simp: conv_dbm_entry_mono_rev neutral)[]
+      using conv_dbm_entry_mono by fastforce
+
+    apply (rule step_impl_canonical)
+     apply assumption
+    apply (rule step_impl_diag_preservation)
+     apply assumption
+    apply (rule diag_reachable)
+    unfolding E_closure .
+  done
+
+lemma step_impl_sound':
+  assumes step: "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>"
+    and canonical: "canonical (curry (conv_M D)) n \<or> check_diag n D"
+    and diag: "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+    and valid: "valid_dbm (curry (conv_M D))"
+  shows
+    "\<exists> M'. step_z_dbm (conv_A A) l (curry (conv_M D)) v n a l' M'
+          \<and> [curry (conv_M D')]v,n = [M']v,n"
+proof -
+  let ?A = "conv_A A"
+  have A: "RI_A n ?A A" by (rule RI_A_conv_A)
+  have "RI n (conv_M D) D" unfolding eq_onp_def by auto
+  from IR_complete[OF this A step] obtain M' where M':
+    "?A \<turnstile>I \<langle>l, conv_M D\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>" "RI n M' D'"
+    by auto
+  show ?thesis
+  proof (cases "check_diag n D")
+    case False
+    with
+      step_impl_sound[of ?A l "conv_M D",
+        OF M'(1) _ global_clock_numbering' triv_numbering'' diag
+        ] canonical(1)
+    obtain M'' where M'':
+      "?A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,a \<langle>l', M''\<rangle>" "[curry M']v,n = [M'']v,n"
+      by auto
+    with M'(2) M''(2) show ?thesis by (auto dest!: RI_zone_equiv[where v = v])
+  next
+    case True (* This part is duplicated very often *)
+    with step_impl_neg_diag_preservation[OF step] have
+      "check_diag n D'"
+      unfolding check_diag_def neutral by auto
+    moreover from step obtain M' where M': "conv_A A \<turnstile> \<langle>l,curry (conv_M D)\<rangle> \<leadsto>v,n,a \<langle>l',M'\<rangle>"
+    proof cases
+      case prems: step_t_impl
+      then show thesis by - (rule that; simp; rule step_t_z_dbm; rule HOL.refl)
+    next
+      case prems: (step_a_impl g a' r)
+      then show thesis
+        apply -
+        apply (rule that)
+        unfolding \<open>a = _\<close>
+        apply (rule step_a_z_dbm[where A = "conv_A A", of l "map conv_ac g" a' r l'])
+        by (fastforce split: prod.split simp: trans_of_def conv_A_def)
+    qed
+    moreover from step_z_dbm_empty[OF global_clock_numbering' this check_diag_empty_spec] True have
+      "[M']v,n = {}"
+      by auto
+    ultimately show ?thesis using check_diag_empty_spec[OF check_diag_conv_M] by auto
+  qed
+qed
+
+lemma step_impl_norm_sound:
+  assumes step: "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>"
+    and canonical: "canonical (curry (conv_M D)) n \<or> check_diag n D"
+    and diag: "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+    and valid: "valid_dbm (curry (conv_M D))"
+  shows
+    "\<exists> M'. step_z_norm' (conv_A A) l (curry (conv_M D)) a l' M'
+           \<and> [curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n = [M']v,n"
+proof -
+  from step_impl_sound'[OF step canonical diag valid] obtain M' where M':
+    "step_z_dbm (conv_A A) l (curry (conv_M D)) v n a l' M'" "[curry (conv_M D')]v,n = [M']v,n"
+    by auto
+  from step_z_norm[OF this(1), of k] have
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>k,v,n,a \<langle>l', norm (FW M' n) (k l') n\<rangle>" .
+  then have step':
+    "step_z_norm' (conv_A A) l (curry (conv_M D)) a l' (norm (FW M' n) (k l') n)"
+    using k_simp_2 by auto
+  from diag have "\<forall>i\<le>n. curry D i i \<le> 0"
+    by (simp add: conv_dbm_entry_mono_rev neutral)
+  from step_impl_canonical[OF step this] have
+    "canonical (curry (conv_M D')) n \<or> (\<exists>i\<le>n. D' (i, i) < 0)" .
+  moreover have "\<forall>i\<le>n. conv_M D' (i, i) \<le> 0"
+    using step_impl_diag_preservation[OF assms(1) \<open>\<forall>i\<le>n. curry D i i \<le> 0\<close>]
+    by (auto dest!: conv_dbm_entry_mono simp: neutral)
+  moreover have "\<forall> i \<le> n. M' i i \<le> 0"
+    using step_z_dbm_diag_preservation[OF M'(1)] diag by auto
+  moreover from step_z_valid_dbm[OF M'(1) global_clock_numbering' valid_abstraction' valid] have
+    "valid_dbm M'" .
+  ultimately have "[curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n
+         = [norm (FW M' n) (\<lambda>x. real (k l' x)) n]v,n"
+    using M'(2) by - (rule norm_step_correct; auto)
+  with step' show ?thesis
+    by auto
+qed
+
+lemmas steps_z_norm'_valid_dbm_preservation =
+  steps_z_norm'_valid_dbm_invariant[OF global_clock_numbering' valid_abstraction']
+
+lemma valid_init_dbm[intro]:
+  "valid_dbm (curry (init_dbm :: real DBM'))"
+  apply rule
+  subgoal
+    unfolding V_def
+    apply safe
+    subgoal for u c
+      by (auto dest: init_dbm_semantics[unfolded conv_M_init_dbm, where c = c])
+    done
+  unfolding init_dbm_def by auto
+
+lemmas
+  step_z_norm_equiv' = step_z_norm_equiv[OF _ global_clock_numbering' valid_abstraction']
+
+lemmas
+  step_z_valid_dbm' = step_z_valid_dbm[OF _ global_clock_numbering' valid_abstraction']
+
+lemma step_impl_norm_complete:
+  assumes step: "step_z_norm' (conv_A A) l (curry (conv_M M)) a l' M'"
+    and canonical: "canonical (curry (conv_M M)) n"
+    and diag: "\<forall>i\<le>n. conv_M M (i, i) \<le> 0"
+    and valid: "valid_dbm (curry (conv_M M))"
+  shows
+    "\<exists> D'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>
+    \<and> [curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n = [M']v,n"
+proof -
+  let ?k = "map real_of_int (k' l')"
+  have k_alt_def: "?k = map (k l') [0..<Suc n]" unfolding k'_def by auto
+  have k: "list_all2 ri ?k (k' l')" by (simp add: list_all2_conv_all_nth ri_def)
+  have "length ?k = Suc n" using length_k' by auto
+  let ?A = "conv_A A"
+  have A: "RI_A n ?A A" by (rule RI_A_conv_A)
+  have M_conv: "RI n (conv_M M) M" unfolding eq_onp_def by auto
+  let ?M' = "uncurry M'"
+  from step obtain D' where D':
+    "M' = norm (FW D' n) (k l') n"
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M M)\<rangle> \<leadsto>v,n,a \<langle>l', D'\<rangle>"
+    apply cases
+    apply (subst (asm) v'_v')
+    apply (subst (asm) k_simp_2')
+    by (auto simp: k_simp_2')
+  from step_impl_complete'[OF D'(2) canonical diag] obtain D'' where D'':
+    "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', D''\<rangle>" "[curry (conv_M D'')]v,n = [D']v,n"
+    by auto
+  from diag have "\<forall>i\<le>n. curry M i i \<le> 0"
+    by (simp add: conv_dbm_entry_mono_rev neutral)
+  have "\<forall>i\<le>n. conv_M D'' (i, i) \<le> 0"
+    using step_impl_diag_preservation[OF D''(1) \<open>\<forall>i\<le>n. curry M i i \<le> 0\<close>]
+    by (auto dest!: conv_dbm_entry_mono simp: neutral)
+  moreover have "\<forall>i\<le>n. D' i i \<le> 0"
+    using step_z_dbm_diag_preservation[OF D'(2)] diag by auto
+  moreover have "canonical (curry (conv_M D'')) n \<or> (\<exists>i\<le>n. D'' (i, i) < 0)"
+    using step_impl_canonical[OF D''(1) \<open>\<forall>i\<le>n. curry M i i \<le> 0\<close>] .
+  moreover have "[curry (conv_M D'')]v,n = [D']v,n" by fact
+  moreover have "valid_dbm D'"
+    using step_z_valid_dbm[OF D'(2) global_clock_numbering' valid_abstraction' valid] .
+  ultimately have
+    "[curry ((map_DBMEntry real_of_int \<circ>\<circ>\<circ> FW') (norm_upd D'' (k' l') n) n)]v,n
+     = [norm (FW D' n) (\<lambda>x. real (k l' x)) n]v,n"
+    by (rule norm_step_correct)
+  with D''(1) show ?thesis by (auto simp add: D'(1))
+qed
+
+definition wf_dbm where
+  "wf_dbm D \<equiv>
+    (canonical (curry (conv_M D)) n \<or> check_diag n D)
+    \<and> (\<forall>i\<le>n. conv_M D (i, i) \<le> 0)
+    \<and> valid_dbm (curry (conv_M D))"
+
+lemma wf_dbm_D:
+  "canonical (curry (conv_M D)) n \<or> check_diag n D"
+  "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+  "valid_dbm (curry (conv_M D))"
+  if "wf_dbm D"
+  using that unfolding wf_dbm_def by auto
+
+lemma wf_dbm_I:
+  "wf_dbm D" if
+  "canonical (curry (conv_M D)) n \<or> check_diag n D"
+  "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+  "valid_dbm (curry (conv_M D))"
+  using that unfolding wf_dbm_def by auto
+
+lemma reachable_wf_dbm:
+  "wf_dbm M" if "E** a0 (l, M)"
+  using canonical_reachable[OF that] diag_reachable[OF that] valid_dbm_reachable[OF that]
+  apply (intro wf_dbm_I)
+  unfolding check_diag_def neutral[symmetric] using diag_conv by auto
+
+lemma canonical_check_diag_empty_iff:
+  "[curry (conv_M D)]v,n = {} \<longleftrightarrow> check_diag n D" if "canonical_diag' D"
+  apply standard
+  subgoal
+    apply (rule canonical_empty_check_diag)
+    using canonical_diagI[OF that] unfolding check_diag_def neutral by auto
+  by (intro check_diag_empty_spec check_diag_conv_M)
+
+lemma step_impl_norm_complete'':
+  assumes step: "step_z_norm' (conv_A A) l (curry (conv_M M)) a l' D"
+    and valid: "valid_dbm (curry (conv_M M))"
+    and canonical: "canonical (curry (conv_M M)) n \<or> check_diag n M"
+    and diag: "\<forall>i\<le>n. conv_M M (i, i) \<le> 0"
+  shows
+    "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>
+           \<and> [curry (conv_M (FW' (norm_upd M' (k' l') n) n))]v,n = [D]v,n"
+proof (cases "check_diag n M")
+  case True
+  then have "[curry (conv_M M)]v,n = {}" by (intro check_diag_empty_spec check_diag_conv_M)
+  with step valid have
+    "[D]v,n = {}"
+    by (rule step_z_norm'_empty_preservation)
+  from step obtain M' where M': "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+    apply cases
+    apply (cases rule: step_z_dbm.cases)
+      apply assumption
+  proof goal_cases
+    case 1
+    then show ?thesis by - (rule that; simp; rule step_t_impl)
+  next
+    case prems: (2 _ g a' r)
+    obtain g' where "A \<turnstile> l \<longrightarrow>g',a',r l'"
+    proof -
+      obtain T I where "A = (T, I)" by force
+      from prems(6) show ?thesis by (fastforce simp: \<open>A = _\<close> trans_of_def conv_A_def intro: that)
+    qed
+    then show ?thesis
+      apply -
+      apply (rule that)
+      unfolding \<open>a = _\<close> by (rule step_a_impl)
+  qed
+  from step_impl_check_diag[OF \<open>check_diag n M\<close> M'] have "check_diag n M'" .
+  from norm_step_check_diag_preservation[OF this] have
+    "check_diag n (FW' (norm_upd M' (k' l') n) n)"
+    by auto
+  with M' \<open>[D]v,n = {}\<close> canonical_check_diag_empty_iff show ?thesis
+    by blast
+next
+  case False
+  with canonical have
+    "canonical (curry (conv_M M)) n"
+    unfolding check_diag_def neutral by auto
+  then show ?thesis using diag valid step_impl_norm_complete[OF step] by auto
+qed
+
+lemmas
+  step_z_dbm_mono = step_z_dbm_mono[OF global_clock_numbering']
+
+lemmas
+  step_z_norm_mono' = step_z_norm_mono[OF _ global_clock_numbering' valid_abstraction']
+
+lemma dbm_subset_correct'':
+  assumes "wf_dbm D" and "wf_dbm M"
+  shows "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n \<longleftrightarrow> dbm_subset n (conv_M D) (conv_M M)"
+  apply (rule dbm_subset_correct')
+  using wf_dbm_D[OF assms(1)] wf_dbm_D[OF assms(2)] by auto
+
+lemma step_impl_wf_dbm:
+  assumes step: "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,a \<langle>l', D'\<rangle>"
+    and wf_dbm: "wf_dbm D"
+  shows "wf_dbm D'"
+  apply (rule wf_dbm_I)
+  using
+    step_impl_diag_preservation[OF step] step_impl_canonical[OF step] step_impl_valid_dbm[OF step]
+    wf_dbm_D[OF wf_dbm]
+  unfolding canonical_variant
+  using diag_conv_rev diag_conv by simp+
+
+lemma norm_step_wf_dbm:
+  "wf_dbm (FW' (norm_upd D (k' l) n) n)" if "wf_dbm D"
+  apply (rule wf_dbm_I)
+  subgoal
+    using canonical_norm_step
+    unfolding canonical_variant .
+  subgoal
+    using wf_dbm_D[OF that] norm_step_diag_preservation
+    using diag_conv[of n "FW' (norm_upd D (k' l) n) n"] diag_conv_rev by simp
+  subgoal
+    using wf_dbm_D[OF that]
+    unfolding canonical_variant
+    by (force intro!: valid_dbm_V norm_step_valid_dbm)
+  done
+
+lemma step_impl_complete''_improved:
+  assumes step: "conv_A A \<turnstile> \<langle>l, curry (conv_M M)\<rangle> \<leadsto>v,n,a \<langle>l', D\<rangle>"
+    and wf_dbm: "wf_dbm M"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle> \<and> [curry (conv_M M')]v,n = [D]v,n"
+proof -
+  note prems = wf_dbm_D[OF wf_dbm]
+  show ?thesis
+  proof (cases "check_diag n M")
+    case True
+    then have "[curry (conv_M M)]v,n = {}" by (intro check_diag_empty_spec check_diag_conv_M)
+    with step prems have
+      "[D]v,n = {}"
+      by - (rule step_z_dbm_empty[OF global_clock_numbering'])
+    moreover from step obtain M' where M': "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>"
+    apply cases
+    proof goal_cases
+      case 1
+      then show ?thesis by - (rule that; simp; rule step_t_impl)
+    next
+      case prems: (2 g a' r)
+      obtain g' where "A \<turnstile> l \<longrightarrow>g',a',r l'"
+      proof -
+        obtain T I where "A = (T, I)" by force
+        from prems(4) show ?thesis by (fastforce simp: \<open>A = _\<close> trans_of_def conv_A_def intro: that)
+      qed
+      then show ?thesis
+        apply -
+        apply (rule that)
+        unfolding \<open>a = _\<close> by (rule step_a_impl)
+    qed
+    ultimately show ?thesis
+      using step_impl_check_diag[OF True M', THEN check_diag_conv_M, THEN check_diag_empty_spec]
+      by auto
+  next
+    case False
+    with prems have
+      "canonical (curry (conv_M M)) n"
+    unfolding check_diag_def neutral by fast
+    moreover from prems have
+      "\<forall>i\<le>n. conv_M M (i, i) \<le> 0" by fast
+    ultimately show ?thesis using step_impl_complete'[OF step] by fast
+  qed
+qed
+
+
+subsubsection \<open>Bisimilarity\<close>
+
+lemma step_impl_equiv:
+  assumes "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and     "wf_dbm D" "wf_dbm M"
+    and "D \<simeq> M"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle> \<and> D' \<simeq> M'"
+proof -
+  note prems_D = wf_dbm_D[OF assms(2)]
+  from step_impl_sound'[OF assms(1) prems_D] diag_conv obtain M' where
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,a \<langle>l', M'\<rangle>"
+    "[curry (conv_M D')]v,n = [M']v,n"
+    unfolding check_diag_def neutral by auto
+  from step_z_dbm_equiv'[OF this(1) assms(4)[unfolded dbm_equiv_def]] this(2)
+  show ?thesis
+    by (auto simp: dbm_equiv_def dest!: step_impl_complete''_improved[OF _ assms(3)])
+qed
+
+lemma norm_step_correct':
+  assumes wf_dbm: "wf_dbm D" "wf_dbm M"
+    and equiv:  "D \<simeq> M"
+  shows
+    "[curry (conv_M (FW' (norm_upd D (k' l') n) n))]v,n
+     = [norm (FW (curry (conv_M M)) n) (\<lambda>x. real (k l' x)) n]v,n"
+  apply (rule norm_step_correct)
+  using wf_dbm_D[OF wf_dbm(1)] wf_dbm_D[OF wf_dbm(2)] equiv
+  unfolding dbm_equiv_def[symmetric] canonical_variant
+  by auto
+
+lemma step_impl_norm_equiv:
+  assumes step:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', D'\<rangle>"
+    "A \<turnstile>I \<langle>l', D'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D''\<rangle>"
+    "D''' = FW' (norm_upd D'' (k' l'') n) n"
+    and wf_dbm: "wf_dbm D" "wf_dbm M"
+    and equiv: "D \<simeq> M"
+  shows "\<exists> M' M'' M'''. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle> \<and> A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>
+       \<and> M''' = FW' (norm_upd M'' (k' l'') n) n \<and> D''' \<simeq> M'''"
+proof -
+  from step_impl_equiv[OF step(1) wf_dbm equiv] obtain M' where M':
+    "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle>" "D' \<simeq> M'"
+    by auto
+  from wf_dbm step(1) M'(1) have wf_dbm':
+    "wf_dbm D'" "wf_dbm M'"
+    by (auto intro: step_impl_wf_dbm)
+  from step_impl_equiv[OF step(2) this M'(2)] obtain M'' where M'':
+    "A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>" "D'' \<simeq> M''"
+    by auto
+  with wf_dbm' step(2) have wf_dbm'':
+    "wf_dbm D''" "wf_dbm M''"
+    by (auto intro: step_impl_wf_dbm)
+  let ?M''' = "FW' (norm_upd M'' (k' l'') n) n"
+  have "D''' \<simeq> ?M'''"
+    unfolding step(3) dbm_equiv_def
+    apply (subst norm_step_correct'[of M'' D''])
+       prefer 4
+       apply (subst norm_step_correct'[of D'' D''])
+    using wf_dbm'' \<open>D'' \<simeq> M''\<close>
+    by (auto intro: dbm_equiv_sym)
+  with M'(1) M''(1) show ?thesis by auto
+qed
+
+definition
+  "wf_state \<equiv> \<lambda> (l, M). wf_dbm M"
+
+lemma E_equiv:
+  "\<exists> b'. E b b' \<and> a' \<sim> b'" if "E a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that
+  unfolding wf_state_def E_def
+  apply safe
+  by (drule step_impl_norm_equiv; force simp: state_equiv_def dest: state_equiv_D)
+
+lemma E_wf_dbm[intro]:
+  "wf_dbm D'" if "E (l, D) (l', D')" "wf_dbm D"
+  using that unfolding wf_state_def E_def by (auto simp: norm_step_wf_dbm step_impl_wf_dbm)
+
+lemma E_wf_state[intro]:
+  "wf_state b" if "E a b" "wf_state a"
+  using that unfolding wf_state_def by auto
+
+lemma E_steps_wf_state[intro]:
+  "wf_state b" if "E** a b" "wf_state a"
+  using that by (induction rule: rtranclp_induct) auto
+
+lemma wf_dbm_init_dbm[intro, simp]:
+  "wf_dbm init_dbm"
+  apply (rule wf_dbm_I)
+  using valid_init_dbm
+  unfolding conv_M_init_dbm
+  unfolding init_dbm_def neutral[symmetric]
+  by simp+
+
+lemma wf_state_init[intro, simp]:
+  "wf_state a0"
+  unfolding wf_state_def a0_def by simp
+
+context
+  fixes E1 :: "'s \<times> _ \<Rightarrow> 's \<times> _ \<Rightarrow> bool" and P
+  assumes E_E1_step: "E a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E1 a c \<and> b \<sim> c)"
+  assumes E1_E_step: "E1 a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E a c \<and> b \<sim> c)"
+  assumes E1_equiv: "E1 a b \<Longrightarrow> a \<sim> a' \<Longrightarrow> P a \<Longrightarrow> P a' \<Longrightarrow> \<exists> b'. E1 a' b' \<and> b \<sim> b'"
+  assumes E1_P[intro]: "P a \<Longrightarrow> E1 a b \<Longrightarrow> P b"
+  assumes wf_state_P[intro]: "wf_state a \<Longrightarrow> P a"
+begin
+
+private lemma E1_steps_P:
+  "P b" if "E1** a b" "P a"
+  using that by (induction rule: rtranclp_induct) auto
+
+context
+  notes [intro] = E1_steps_P
+begin
+
+private lemma E_E1_step':
+  "(\<exists> b'. E1 b b' \<and> a' \<sim> b')" if "E a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that E_equiv[OF that] by (blast dest: E_E1_step)
+
+private lemma E1_E_step':
+  "(\<exists> b'. E b b' \<and> a' \<sim> b')" if "E1 a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that
+  apply -
+  apply (drule E1_E_step, assumption)
+  apply safe
+  by (drule E_equiv; blast)
+
+private lemma E_E1_steps:
+  "\<exists> b'. E1** a b' \<and> a' \<sim> b'" if "E** a a'" "wf_state a"
+  using that
+proof (induction rule: rtranclp_induct)
+  case base
+  then show ?case by blast
+next
+  case (step y z)
+  then obtain b where b: "E1** a b" "y \<sim> b" by auto
+  from step obtain z' where "E1 y z'" "z \<sim> z'"
+    by atomize_elim (blast intro: E_E1_step)
+  from step b \<open>E1 y z'\<close> obtain b' where "E1 b b'" "z' \<sim> b'"
+    by atomize_elim (blast intro: E1_equiv)
+  with b \<open>z \<sim> z'\<close> show ?case
+    by (blast intro: rtranclp.intros(2))
+qed
+
+private lemma E1_E_steps:
+  "\<exists> b'. E** a b' \<and> a' \<sim> b'" if "E1** a a'" "wf_state a"
+  using that
+proof (induction rule: rtranclp_induct)
+  case base
+  then show ?case by blast
+next
+  case (step y z)
+  then obtain y' where y': "E** a y'" "y \<sim> y'"
+    by auto
+  with step obtain z' where "E1 y' z'" "z \<sim> z'"
+    by atomize_elim (blast intro: E1_equiv)
+  with y' \<open>wf_state _\<close> obtain z'' where "E y' z''" "z' \<sim> z''"
+    by atomize_elim (blast intro: E1_E_step)
+  with y' \<open>z \<sim> z'\<close> show ?case
+    by (blast intro: rtranclp.intros(2))
+qed
+
+lemma E_E1_steps_equiv:
+  "(\<exists> M'. E** a0 (l', M') \<and> [curry (conv_M M')]v,n \<noteq> {}) \<longleftrightarrow>
+   (\<exists> M'. E1** a0 (l', M') \<and> [curry (conv_M M')]v,n \<noteq> {})"
+  by (auto 4 4 simp: state_equiv_def dbm_equiv_def dest: E_E1_steps E1_E_steps)
+
+end (* End of anonymous context for intro declaration *)
+
+end (* End of anonymous context *)
+
+definition "E_step \<equiv>
+  \<lambda> (l, D) (l'', D'').
+    \<exists> l' M' a M''. conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l', M'\<rangle>
+           \<and> step_z_norm' (conv_A A) l' M' (\<upharpoonleft>a) l'' M''
+           \<and> [curry (conv_M D'')]v,n = [M'']v,n"
+
+lemma step_impl_sound_wf:
+  assumes step: "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and     wf: "wf_dbm D"
+  shows
+    "\<exists> M'. step_z_norm' (conv_A A) l (curry (conv_M D)) a l' M'
+    \<and> [curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n = [M']v,n"
+  using step_impl_norm_sound[OF step wf_dbm_D[OF wf]] .
+
+lemmas valid_dbm_convI = valid_dbm.intros[OF _ dbm_int_conv]
+lemmas step_z_norm_valid_dbm'_spec =
+  step_z_norm_valid_dbm'[OF global_clock_numbering' valid_abstraction']
+
+lemma E_E1_step_aux:
+  "\<exists> l' M' a M''.
+    conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l', M'\<rangle> \<and>
+    step_z_norm' (conv_A A) l' M' (\<upharpoonleft>a) l'' M'' \<and> [curry (conv_M D'')]v,n = [M'']v,n"
+  if "E (l, D) (l'', D'')" "wf_dbm D"
+  using that unfolding E_def
+proof (safe, goal_cases)
+  case prems: (1 D1 l' D2 a)
+  from step_impl_sound'[OF prems(2)] prems(1) obtain M' where M':
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l', M'\<rangle>" "[curry (conv_M D1)]v,n = [M']v,n"
+    by (blast dest: wf_dbm_D)
+  from prems have "wf_dbm D1"
+    by (blast intro: step_impl_wf_dbm)
+  from step_impl_sound_wf[OF prems(3) this] obtain M'' where M'':
+    "step_z_norm'(conv_A A) l' (curry (conv_M D1)) (\<upharpoonleft>a) l'' M''"
+    "[curry (conv_M (FW' (norm_upd D2 (k' l'') n) n))]v,n = [M'']v,n"
+    by blast
+  from \<open>wf_dbm D1\<close> \<open>wf_dbm D\<close> M'(1) have "valid_dbm (curry (conv_M D1))" "valid_dbm M'"
+    by (blast intro: wf_dbm_D step_z_valid_dbm')+
+  with step_z_norm_equiv'[OF M''(1) _ _ M'(2)] obtain M''' where
+    "step_z_norm'(conv_A A) l' M' (\<upharpoonleft>a) l'' M'''" "[M'']v,n = [M''']v,n"
+    by blast
+  with M'(1) M''(2) show ?case by auto
+qed
+
+lemma E_E1_step: "\<exists> c. E_step a c \<and> b \<sim> c" if "E a b" "wf_state a"
+  apply (cases a)
+  apply (cases b)
+  using that
+  apply simp
+  unfolding wf_state_def
+  apply (drule E_E1_step_aux)
+  unfolding E_step_def
+  unfolding state_equiv_def dbm_equiv_def
+  by blast+
+
+lemma E1_E_step: "\<exists> c. E a c \<and> b \<sim> c" if "E_step a b" "wf_state a"
+  using that
+  unfolding wf_state_def
+  unfolding E_step_def
+proof (safe del: step_z_norm'_elims, goal_cases)
+  case prems: (1 l D l'' D'' l' M' a M'')
+  from step_impl_complete''_improved[OF prems(4,1)] obtain M1 where M1:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', M1\<rangle>" "[curry (conv_M M1)]v,n = [M']v,n"
+    by blast
+  with prems have "valid_dbm M'" "wf_dbm M1" "valid_dbm (curry (conv_M M1))"
+    by (blast intro: wf_dbm_D step_z_valid_dbm' step_impl_wf_dbm)+
+  with step_z_norm_equiv'[OF prems(5) _ _ M1(2)[symmetric]] obtain M2 where M2:
+    "step_z_norm' (conv_A A) l' (curry (conv_M M1)) (\<upharpoonleft>a) l'' M2" "[M'']v,n = [M2]v,n"
+    by blast
+  from wf_dbm_D(1)[OF \<open>wf_dbm M1\<close>] have "canonical (curry (conv_M M1)) n \<or> check_diag n M1" .
+  then show ?case
+  proof standard
+    assume "canonical (curry (conv_M M1)) n"
+    with \<open>wf_dbm M1\<close> step_impl_norm_complete[OF M2(1)] obtain D3 where
+      "A \<turnstile>I \<langle>l', M1\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D3\<rangle>"
+      "[curry (conv_M (FW' (norm_upd D3 (k' l'') n) n))]v,n = [M2]v,n"
+      by (blast dest: wf_dbm_D)
+    with M1(1) M2(2) prems(6) show ?case
+      unfolding E_def state_equiv_def dbm_equiv_def by blast
+  next
+    assume "check_diag n M1"
+    with M1(2) have "[M']v,n = {}" by (auto dest: check_diag_conv_M simp: check_diag_empty_spec)
+    with prems(5) \<open>valid_dbm M'\<close> have "[M'']v,n = {}"
+      by - (drule step_z_norm'_empty_preservation; assumption)
+    from step_impl_norm_complete''[OF M2(1)] \<open>wf_dbm M1\<close> obtain D3 where D3:
+      "A \<turnstile>I \<langle>l', M1\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D3\<rangle>"
+      "[M2]v,n \<subseteq> [curry (conv_M (FW' (norm_upd D3 (k' l'') n) n))]v,n"
+      by (blast dest: wf_dbm_D)
+    from step_impl_check_diag[OF \<open>check_diag n M1\<close> D3(1)] have "check_diag n D3" .
+    then have "[curry (conv_M (FW' (norm_upd D3 (k' l'') n) n))]v,n = {}"
+      using check_diag_empty_spec norm_step_check_diag_preservation by blast
+    with D3 \<open>[M'']v,n = {}\<close> M1(1) prems(6) show ?case
+      unfolding E_def state_equiv_def dbm_equiv_def by blast
+  qed
+qed
+
+lemma E1_equiv:
+  defines "P \<equiv> valid_dbm o curry o conv_M o snd"
+  assumes that: "E_step a b" "a \<sim> a'" "P a" "P a'"
+  shows "\<exists> b'. E_step a' b' \<and> b \<sim> b'"
+  using that
+  unfolding E_step_def state_equiv_def dbm_equiv_def
+  apply (safe del: step_z_norm'_elims)
+  apply (frule step_z_dbm_equiv', assumption)
+  apply (erule exE conjE)+
+  unfolding P_def
+  by (drule step_z_norm_equiv'; fastforce dest: step_z_valid_dbm')
+
+lemma E1_P:
+  defines "P \<equiv> valid_dbm o curry o conv_M o snd"
+  assumes that: "E_step a b" "P a"
+  shows "P b"
+  using that
+  unfolding E_step_def P_def
+  apply (safe del: step_z_norm'_elims)
+  apply (drule step_z_valid_dbm', (simp; fail))
+  apply (drule step_z_norm_valid_dbm'_spec, (simp; fail))
+  using valid_dbm_convI by (fastforce dest: valid_dbm_V)
+
+lemma E_steps_equiv:
+  "(\<exists> M'. E** a0 (l', M') \<and> [curry (conv_M M')]v,n \<noteq> {}) \<longleftrightarrow>
+   (\<exists> M'. E_step** a0 (l', M') \<and> [curry (conv_M M')]v,n \<noteq> {})"
+proof -
+  define P :: "'s \<times> _ \<Rightarrow> bool" where "P \<equiv> valid_dbm o curry o conv_M o snd"
+  show ?thesis
+    apply (rule E_E1_steps_equiv[where P = P])
+        apply (rule E_E1_step; assumption)
+       apply (rule E1_E_step; assumption)
+      apply (rule E1_equiv; simp add: P_def)
+     apply (drule E1_P; simp add: P_def)
+    unfolding P_def wf_state_def by (auto dest: wf_dbm_D)
+qed
+
+
+subsubsection \<open>First Layer of Simulation\<close>
+
+lemma E_step_valid_dbm[intro]:
+  assumes "valid_dbm (curry (conv_M D))" "E_step (l, D) (l', D')"
+  shows "valid_dbm (curry (conv_M D'))"
+  using assms unfolding E_step_def
+  proof (safe del: step_z_norm'_elims, goal_cases)
+    case prems: (1 l' M' a M'')
+    then have "valid_dbm M''"
+      by (elim step_z_norm_valid_dbm'_spec step_z_valid_dbm')
+    with prems(4) show ?case
+      by (blast elim!: valid_dbm.cases intro: valid_dbm_convI)
+  qed
+
+lemma E_steps_valid_dbm[intro]:
+  assumes "E_step** a0 (l, D)"
+  shows "valid_dbm (curry (conv_M D))"
+  using assms
+  apply (induction "(l, D)" arbitrary: l D rule: rtranclp_induct)
+  subgoal
+    using valid_init_dbm by (auto simp only: conv_M_init_dbm a0_def)
+  subgoal for x l D
+    by (cases x) blast
+  done
+
+lemma E_steps_steps_z_norm':
+  "\<exists> M'. steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M' \<and> [M']v,n = [curry (conv_M D')]v,n"
+  if "E_step** a0 (l', D')"
+using that proof (induction "(l', D')" arbitrary: l' D' rule: rtranclp_induct)
+  case base
+  then show ?case by (auto simp: a0_def; fail)
+next
+  case (step y)
+  obtain l D where [simp]: "y = (l, D)"
+    by force
+  from step(3-)[OF this] obtain M' where M':
+    "steps_z_norm' (conv_A A) l0 (curry init_dbm) l M'" "[M']v,n = [curry (conv_M D)]v,n"
+    by atomize_elim
+  from step(2) obtain Mi li a M'' where step':
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>li, Mi\<rangle>"
+    "step_z_norm' (conv_A A) li Mi (\<upharpoonleft>a) l' M''"
+    "[curry (conv_M D')]v,n = [M'']v,n"
+    unfolding E_step_def by auto
+  from step_z_dbm_equiv'[OF this(1) M'(2)[symmetric]] obtain Mi' where Mi':
+    "conv_A A \<turnstile> \<langle>l, M'\<rangle> \<leadsto>v,n,\<tau> \<langle>li, Mi'\<rangle>" "[Mi]v,n = [Mi']v,n"
+    by atomize_elim
+  with step'(1) M'(1) step(1) have "valid_dbm Mi" "valid_dbm Mi'"
+    unfolding \<open>y = _\<close> by (blast intro: steps_z_norm'_valid_dbm_preservation step_z_valid_dbm')+
+  from step_z_norm_equiv'[OF step'(2) this Mi'(2)] obtain M''' where
+    "step_z_norm' (conv_A A) li Mi' (\<upharpoonleft>a) l' M'''" "[M'']v,n = [M''']v,n"
+    by atomize_elim
+  with Mi'(1) M'(1) step'(3) show ?case
+    unfolding step_z_norm''_def by (auto 4 5 elim: rtranclp.intros(2))
+qed
+
+lemma dbm_int_conv_M_equiv:
+  "\<exists> D. [M]v,n = [curry (conv_M D)]v,n" if "dbm_int M n"
+proof -
+  define D where "D \<equiv> \<lambda> (i, j). map_DBMEntry floor (M i j)"
+  have "rel_DBMEntry ri (M i j) (D (i, j))" if "i \<le> n" "j \<le> n" for i j
+  proof -
+    from that \<open>dbm_int M n\<close> have "M i j \<noteq> \<infinity> \<longrightarrow> get_const (M i j) \<in> \<int>" by blast
+    then show ?thesis
+      unfolding D_def ri_def by (cases "M i j") (auto elim: Ints_cases)
+  qed
+  then have "RI n (uncurry M) D"
+    unfolding eq_onp_def by auto
+  from RI_zone_equiv[OF this] show ?thesis
+    by auto
+qed
+
+definition "A' = conv_A A"
+
+interpretation Bisimulation_Invariant
+  "\<lambda> (l, Z) (l', Z'). \<exists> a. step_z_norm'' (conv_A A) l Z a l' Z'"
+  E_step
+  "\<lambda> (l, M) (l', D). l' = l \<and> [M]v,n = [curry (conv_M D)]v,n"
+  "\<lambda> (l, M). valid_dbm M" "\<lambda> (l, D). valid_dbm (curry (conv_M D))"
+  apply standard
+  subgoal premises prems for a b a'
+  proof -
+    from prems(2) obtain l M D l' M' where [simp]: "a = (l, M)" "a' = (l, D)" "b = (l', M')"
+      by force
+    from prems[unfolded step_z_norm''_def, simplified, folded A'_def] obtain l1 M1 a where steps:
+      "A' \<turnstile> \<langle>l, M\<rangle> \<leadsto>v,n,\<tau> \<langle>l1, M1\<rangle>"
+      "A' \<turnstile> \<langle>l1, M1\<rangle> \<leadsto>(\<lambda>l. (k l \<circ>\<circ>\<circ> Regions.v' {Suc 0..<Suc n} v) n (Suc n)),v,n,\<upharpoonleft>a \<langle>l', M'\<rangle>"
+      by auto
+    from step_z_dbm_equiv'[folded A'_def, OF this(1), of "curry (conv_M D)"] prems(2) obtain M2
+      where M2: "A' \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l1, M2\<rangle>" "[M1]v,n = [M2]v,n"
+      by auto
+    from steps(2) obtain M3
+      where M3:
+        "A' \<turnstile> \<langle>l1, M2\<rangle> \<leadsto>(\<lambda>l. (k l \<circ>\<circ>\<circ> Regions.v' {Suc 0..<Suc n} v) n (Suc n)),v,n,\<upharpoonleft>a \<langle>l', M3\<rangle>"
+        "[M']v,n = [M3]v,n"
+      using prems(3,4) steps(1) M2(1)
+      by atomize_elim (auto
+          intro!: step_z_norm_equiv'[folded A'_def, simplified, OF steps(2) _ _ M2(2)]
+          dest: step_z_valid_dbm'[folded A'_def]
+          )
+    from prems(4) M2(1) M3(1) have "valid_dbm M3"
+      by - (rule step_z_norm_valid_dbm'_spec[folded A'_def], fastforce,
+          fastforce dest: step_z_valid_dbm'[folded A'_def]
+          )
+    then have "dbm_int M3 n"
+      by - (erule valid_dbm_cases)
+    from dbm_int_conv_M_equiv[OF this] obtain D' where "[M3]v,n = [curry (conv_M D')]v,n"
+      by auto
+    with \<open>[M']v,n = _\<close> show ?thesis
+      unfolding E_step_def A'_def[symmetric] by (fastforce intro: M2(1) M3(1))
+  qed
+
+  subgoal for a a' b'
+    unfolding step_z_norm''_def E_step_def A'_def[symmetric]
+    apply clarsimp
+    apply (frule step_z_dbm_equiv'[folded A'_def], rule sym, assumption)
+    apply clarify
+    by (frule step_z_norm_equiv'[folded A'_def, simplified];
+        auto elim: step_z_valid_dbm'[folded A'_def, simplified]
+        ) auto
+
+  subgoal for a b
+    by (blast intro: steps_z_norm'_valid_dbm_preservation)
+
+  subgoal for a b
+    by blast
+  done
+
+lemma steps_z_norm'_E_steps:
+  "\<exists> D'. E_step** a0 (l', D') \<and> [M']v,n = [curry (conv_M D')]v,n" if
+  "steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M'"
+  using bisim.A_B_reaches[OF that, of a0] valid_init_dbm unfolding a0_def equiv'_def by auto
+
+lemma E_steps_steps_z_norm'_iff:
+  "(\<exists> D'. E_step** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+  \<longleftrightarrow> (\<exists> M'. steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M' \<and> [M']v,n \<noteq> {})"
+  using steps_z_norm'_E_steps E_steps_steps_z_norm' by fast
+
+subsubsection \<open>Monotonicity\<close>
+
+lemma E_step_mono_reachable:
+  assumes "E_step (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+  shows "\<exists> M'. E_step (l,M) (l',M') \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+  using assms
+  unfolding E_step_def
+  apply (safe del: step_z_norm'_elims)
+  apply (frule step_z_dbm_mono, assumption)
+  apply (safe del: step_z_norm'_elims)
+  apply (drule step_z_norm_mono')
+     prefer 3
+     apply assumption
+    apply (blast intro: step_z_valid_dbm' wf_dbm_D)
+   apply (blast intro: step_z_valid_dbm' wf_dbm_D)
+  apply (safe del: step_z_norm'_elims)
+  subgoal premises prems for l'' M' a M'' D'' M'''
+  proof -
+    from prems have "valid_dbm M'''"
+      by (fast intro: step_z_valid_dbm' wf_dbm_D step_z_norm_valid_dbm'_spec)
+    with dbm_int_conv_M_equiv[of M'''] obtain D''' where D''':
+      "[M''']v,n = [curry (conv_M D''')]v,n"
+      by (blast elim!: valid_dbm.cases)
+    with prems show ?thesis
+      by fastforce
+  qed
+  done
+
+lemma E_mono:
+  assumes "E (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+  shows "\<exists> M'. E (l,M) (l',M') \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+  using assms
+  apply -
+  apply (drule E_E1_step)
+   apply (simp add: wf_state_def; fail)
+  apply safe
+  apply (drule E_step_mono_reachable[where M = M], assumption+)
+  by (force simp: wf_state_def state_equiv_def dbm_equiv_def dest: E1_E_step)
+
+lemma E_mono':
+  assumes "E (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "dbm_subset n D M"
+  shows "\<exists> M'. E (l,M) (l',M') \<and> dbm_subset n D' M'"
+  using assms
+  apply -
+  apply (frule E_mono[where M = M], assumption+)
+   apply (subst dbm_subset_correct'', assumption+)
+   apply (rule dbm_subset_conv, assumption)
+  apply safe
+  apply (subst (asm) dbm_subset_correct'')
+    apply (blast intro: dbm_subset_conv_rev)+
+  done
+
+lemma reachable_steps_z_norm':
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+    \<longleftrightarrow> (\<exists> M'. steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M' \<and> [M']v,n \<noteq> {})"
+  unfolding E_steps_steps_z_norm'_iff E_steps_equiv ..
+
+subsubsection \<open>Instantiating the Reachability Problem\<close>
+
+theorem reachable_decides_emptiness:
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+    \<longleftrightarrow> (\<exists> u \<in> [curry init_dbm]v,n. \<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>)"
+  by (subst reachable_steps_z_norm')
+    (subst
+      steps_z_norm_decides_emptiness
+      [OF global_clock_numbering' valid_abstraction'[unfolded X_alt_def]
+        finite_trans_of_finite_state_set[OF finite_trans'] valid_init_dbm
+        ],
+      rule HOL.refl
+      )
+
+lemma reachable_decides_emptiness':
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+    \<longleftrightarrow> (\<exists> u u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0))"
+  using reachable_decides_emptiness init_dbm_semantics' init_dbm_semantics'' by blast
+
+lemma reachable_empty_check_diag:
+  assumes "E** a0 (l', D')" "[curry (conv_M D')]v,n = {}"
+  shows "check_diag n D'"
+  using canonical_reachable[OF assms(1)] canonical_empty_check_diag assms(2) by simp
+
+lemma check_diag_E_preservation:
+  "check_diag n M'" if "check_diag n M" "E (l, M) (l', M')"
+  using that unfolding E_def check_diag_def neutral[symmetric]
+  by (fastforce
+      simp: curry_def dest: step_impl_neg_diag_preservation
+      intro: FW'_neg_diag_preservation norm_upd_neg_diag_preservation
+      )
+
+lemma cn_weak:
+  "\<forall>c. 0 < v c"
+  using clock_numbering(1) by auto
+
+end (* TA Start *)
+
+
+locale Reachability_Problem =
+  TA_Start A l0 n k +
+  Reachability_Problem_Defs A l0 n k F for A :: "('a, nat, int, 's) ta" and l0 n k F
+begin
+
+theorem reachability_check:
+  "(\<exists> D'. E** a0 (l', D') \<and> F_rel (l', D'))
+    \<longleftrightarrow> (\<exists> u u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0) \<and> F l')"
+  using reachable_decides_emptiness'[of l'] check_diag_empty_spec reachable_empty_check_diag
+  unfolding F_rel_def by auto
+
+sublocale Standard_Search_Space: Search_Space E a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M"
+  apply standard
+  subgoal for a
+    apply (rule prod.exhaust[of a])
+    by (auto simp add: subsumes_simp_1 dbm_subset_refl)
+  subgoal for a b c
+    apply (rule prod.exhaust[of a], rule prod.exhaust[of b], rule prod.exhaust[of c])
+    subgoal for l1 M1 l2 M2 l3 M3
+      apply simp
+      apply (cases "check_diag n M1")
+       apply (simp add: subsumes_def; fail)
+      apply (simp add: subsumes_def)
+      by (meson check_diag_subset dbm_subset_def dbm_subset_trans)
+    done
+  subgoal for a b a'
+    apply (rule prod.exhaust[of a], rule prod.exhaust[of b], rule prod.exhaust[of a'])
+    apply safe
+    by (drule E_mono';
+        fastforce simp: E_def subsumes_def dbm_subset_def Graph_Start_Defs.reachable_def
+        intro!: reachable_wf_dbm)
+  subgoal
+    unfolding F_rel_def subsumes_def by auto
+  subgoal
+    using check_diag_subset unfolding subsumes_def dbm_subset_def by auto
+  subgoal
+    using check_diag_E_preservation by auto
+  unfolding F_rel_def subsumes_def
+  unfolding check_diag_def pointwise_cmp_def
+  by fastforce
+
+sublocale Standard_Search_Space_finite_strict:
+  Search_Space_finite_strict E a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M"
+  apply standard
+  using E_closure_finite unfolding Graph_Start_Defs.reachable_def by (auto; fail)
+
+
+paragraph \<open>Nice to have\<close>
+
+lemma V_I:
+  assumes "\<forall> i \<in> {1..<Suc n}. M 0 i \<le> 0"
+  shows "[M]v,n \<subseteq> V"
+  unfolding V_def DBM_zone_repr_def
+proof (safe, goal_cases)
+  case prems: (1 u i)
+  then have "v i = i"
+    using X_alt_def X_def triv_numbering by blast
+  with prems have "v i > 0" "v i \<le> n" by auto
+  with prems have "dbm_entry_val u None (Some i) (M 0 (v i))"
+    unfolding DBM_val_bounded_def by auto
+  moreover from assms \<open>v i > 0\<close> \<open>v i \<le> n\<close> have "M 0 (v i) \<le> 0" by auto
+  ultimately
+  show ?case
+    apply (cases "M 0 (v i)")
+    unfolding neutral less_eq dbm_le_def
+    by (auto elim!: dbm_lt.cases simp: \<open>v i = i\<close>)
+qed
+
+lemma DBM_val_bounded_cong':
+  "\<forall>c\<in>{Suc 0..n}. u c = u' c \<Longrightarrow> u \<turnstile>v,n M \<Longrightarrow> u' \<turnstile>v,n M"
+  unfolding DBM_val_bounded_def
+proof (safe, goal_cases)
+  case prems: (1 c)
+  from prems have v: "v c = c" "v c > 0" by (auto simp: cn_weak intro: v_id)
+  show ?case
+  proof (cases "M 0 (v c)")
+    case (Le d)
+    with prems have "- u c \<le> d" by auto
+    with Le v prems show ?thesis by auto
+  next
+    case (Lt d)
+    with prems have "- u c < d" by auto
+    with Lt v prems show ?thesis by auto
+  qed auto
+next
+  case prems: (2 c)
+  from prems have v: "v c = c" "v c > 0" by (auto simp: cn_weak intro: v_id)
+  show ?case
+  proof (cases "M (v c) 0")
+    case (Le d)
+    with prems have "u c \<le> d" by auto
+    with Le v prems show ?thesis by auto
+  next
+    case (Lt d)
+    with prems have "u c < d" by auto
+    with Lt v prems show ?thesis by auto
+  qed auto
+next
+  case prems: (3 c1 c2)
+  from prems have v: "v c1 = c1" "v c1 > 0" "v c2 = c2" "v c2 > 0"
+    by (auto simp: cn_weak intro: v_id)
+  show ?case
+  proof (cases "M (v c1) (v c2)")
+    case (Le d)
+    with prems have "u c1 - u c2 \<le> d" by (auto 4 3)
+    with Le v prems show ?thesis by auto
+  next
+    case (Lt d)
+    with prems have "u c1 - u c2 < d" by (auto 4 3)
+    with Lt v prems show ?thesis by auto
+  qed auto
+qed
+
+lemma DBM_val_bounded_cong:
+  "\<forall>c\<in>{Suc 0..n}. u c = d \<Longrightarrow> \<forall>c\<in>{Suc 0..n}. u' c = d \<Longrightarrow> u \<turnstile>v,n M \<Longrightarrow> u' \<turnstile>v,n M"
+  using DBM_val_bounded_cong' by force
+
+end (* End of locale context *)
+
+lemma abstr_upd_correct:
+  assumes gcn: "\<forall>c. 0 < v c \<and> (\<forall>x y. v x \<le> n \<and> v y \<le> n \<and> v x = v y \<longrightarrow> x = y)"
+      and surj: "\<forall> k \<le> n. k > 0 \<longrightarrow> (\<exists> c. v c = k)"
+      and clocks: "\<forall>c\<in>collect_clks (inv_of A l). v c = c \<and> 0 < c \<and> v c \<le> n"
+  shows
+  "[curry (abstr_upd (inv_of A l) D)]v,n = [And (curry D) (abstr (inv_of A l) (\<lambda>i j. \<infinity>) v)]v,n"
+  apply (subst abstr_upd_abstr')
+   defer
+   apply (subst abstr_abstr'[symmetric])
+    defer
+    apply (subst And_abstr[symmetric])
+  using assms by fastforce+
+
+locale Reachability_Problem_start =
+  Reachability_Problem +
+  assumes start_inv: "[(curry init_dbm :: real DBM)]v,n \<subseteq> {u. u \<turnstile> conv_cc (inv_of A l0)}"
+begin
+
+lemma start_inv':
+  "[(curry init_dbm :: real DBM)]v,n \<subseteq> {u. u \<turnstile> inv_of (conv_A A) l0}"
+  using start_inv unfolding conv_A_def
+  by (smt case_prod_conv comp_apply inv_of_def prod.collapse snd_conv subset_Collect_conv)
+
+lemma start_vals:
+  "u \<turnstile> inv_of (conv_A A) l0" if "\<forall> c \<in> {1..n}. u c = 0"
+  using that start_inv' init_dbm_semantics'' by auto
+
+lemma steps_steps'_equiv':
+  "(\<exists> u u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0))
+  \<longleftrightarrow> (\<exists> u u'. conv_A A \<turnstile> \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0))"
+  using steps_steps'_equiv[of _ \<open>conv_A A\<close> l0] start_vals by fast
+
+corollary reachability_check_start:
+  "(\<exists> D'. E** a0 (l', D') \<and> F_rel (l', D'))
+  \<longleftrightarrow> (\<exists> u u'. conv_A A \<turnstile> \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0) \<and> F l')"
+  using reachability_check steps_steps'_equiv' by fast
+
+end (* End of locale context for ensuring that the invariant of the start state is satisfied *)
+
+lemma check_diag_conv_M_rev:
+  assumes "check_diag n (conv_M M)"
+  shows "check_diag n M"
+using assms unfolding check_diag_def by (auto intro!: conv_dbm_entry_mono_strict_rev)
+
+context TA_Start_No_Ceiling
+begin
+
+lemma surj_numbering:
+  "\<forall>k\<le>n. 0 < k \<longrightarrow> (\<exists>c. v c = k)"
+  using global_clock_numbering[THEN conjunct2, THEN conjunct2] .
+
+lemma collect_clks_numbering:
+  "\<forall>c\<in>collect_clks (inv_of A l0). v c = c \<and> 0 < c \<and> v c \<le> n"
+  by (metis collect_clks_inv_clk_set global_clock_numbering subsetCE v_id)
+
+lemma collect_clks_numbering':
+  "\<forall>c\<in>collect_clks (inv_of (conv_A A) l0). v c = c \<and> 0 < c \<and> v c \<le> n"
+  by (metis (no_types, lifting) collect_clks_inv_clk_set global_clock_numbering' subsetCE triv_numbering'')
+
+lemmas dbm_abstr_zone_eq' = dbm_abstr_zone_eq[OF global_clock_numbering[THEN conjunct1]]
+
+end
+
+
+context TA_Start
+begin
+
+lemmas abstr_upd_correct' =
+  abstr_upd_correct[OF clock_numbering(1) surj_numbering collect_clks_numbering']
+
+lemma FW'_zone_equiv:
+  "[curry (conv_M (FW' M n))]v,n = [curry (conv_M M)]v,n"
+  by (metis FW'_FW FW'_conv_M FW_zone_equiv_spec)
+
+lemma unbounded_dbm_semantics:
+  "[curry (unbounded_dbm n)]v,n = UNIV"
+  unfolding unbounded_dbm_def
+  unfolding DBM_zone_repr_def apply auto
+  unfolding DBM_val_bounded_def
+  apply auto
+  (* s/h *)
+  using cn_weak gr_implies_not_zero apply blast
+  using cn_weak gr_implies_not_zero apply blast
+  (* s/h *)
+  by (metis dbm_entry_val.intros(5) diff_self order_mono_setup.refl v_id)
+
+lemma And_unbounded_dbm:
+  "[And (curry (unbounded_dbm n)) M]v,n = [M]v,n"
+  by (subst And_correct[symmetric]) (simp add: unbounded_dbm_semantics)
+
+definition
+  "start_inv_check \<equiv> dbm_subset n init_dbm (FW' (abstr_upd (inv_of A l0) (unbounded_dbm n)) n)"
+
+lemma conv_inv:
+  "conv_cc (inv_of A l0) = inv_of (conv_A A) l0"
+  unfolding inv_of_def conv_A_def by (simp split!: prod.split)
+
+lemma start_inv_check:
+  "start_inv_check \<longleftrightarrow> [(curry init_dbm :: real DBM)]v,n \<subseteq> {u. u \<turnstile> conv_cc (inv_of A l0)}"
+proof -
+  let ?M = "FW' (abstr_upd (inv_of A l0) (unbounded_dbm n)) n"
+  have canonical: "canonical (curry init_dbm) n \<or> check_diag n init_dbm"
+    unfolding init_dbm_def by (auto simp: neutral[symmetric])
+  have diag: "\<forall>i\<le>n. curry (conv_M ?M) i i \<le> 0"
+    apply (rule diag_conv)
+    unfolding curry_def unbounded_dbm_def
+    apply (rule FW'_diag_preservation)
+    apply (rule abstr_upd_diag_preservation')
+    using collect_clks_numbering unfolding neutral by auto
+  have "\<forall>i\<le>n. curry init_dbm i i \<le> 0" unfolding init_dbm_def neutral by auto
+  from dbm_subset_correct'[OF canonical this diag] canonical have
+    "dbm_subset n init_dbm (conv_M ?M)
+    \<longleftrightarrow> [(curry init_dbm :: real DBM)]v,n \<subseteq> [curry (conv_M ?M)]v,n"
+    by auto
+  moreover have "[curry (conv_M ?M)]v,n = {u. u \<turnstile> inv_of (conv_A A) l0}"
+    apply (subst FW'_zone_equiv)
+    apply (subst RI_zone_equiv[symmetric,
+           where M = "abstr_upd (inv_of (conv_A A) l0) (unbounded_dbm n)"])
+     defer
+     apply (subst abstr_upd_correct')
+     apply (subst And_unbounded_dbm)
+     apply (subst dbm_abstr_zone_eq')
+    using collect_clks_numbering' apply force
+     apply simp
+    subgoal
+    proof -
+      have [transfer_rule]: "rel_DBMEntry ri 0 0" "rel_DBMEntry ri \<infinity> \<infinity>"
+        unfolding ri_def neutral by auto
+      have [transfer_rule]: "eq_onp (\<lambda>x. x < Suc n) n n"
+        by (simp add: eq_onp_def)
+      have [transfer_rule]:
+        "rel_fun (eq_onp (\<lambda>x. x < Suc n)) (rel_fun (eq_onp (\<lambda>x. x < Suc n)) (=)) (<) (<)"
+        unfolding rel_fun_def eq_onp_def by simp
+      have [transfer_rule]: "(RI n) (unbounded_dbm n) (unbounded_dbm n)"
+        unfolding unbounded_dbm_def by transfer_prover
+      have [transfer_rule]: "RI_A n (conv_A A) A" by (rule RI_A_conv_A)
+      note [transfer_rule] = inv_of_transfer[unfolded RI_I_def]
+      show ?thesis by transfer_prover
+    qed
+    done
+  ultimately show ?thesis
+    unfolding start_inv_check_def conv_inv using dbm_subset_conv_rev dbm_subset_conv conv_M_init_dbm
+    by metis
+qed
+
+lemma start_inv_check_correct:
+  "start_inv_check \<longleftrightarrow> (\<forall> u. (\<forall> c \<in> {1..n}. u c = 0) \<longrightarrow> u \<turnstile> inv_of (conv_A A) l0)"
+  unfolding start_inv_check conv_inv using  init_dbm_semantics'' init_dbm_semantics' by fast
+
+end
+
+context Reachability_Problem
+begin
+
+lemma F_reachable_equiv':
+  "(\<exists> l' u u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0))
+    \<longleftrightarrow> (\<exists> l' u u'. conv_A A \<turnstile> \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0))"
+  if start_inv_check
+proof -
+
+  interpret start: Reachability_Problem_start A l0 n k
+    apply standard
+    using that unfolding start_inv_check .
+  from start.steps_steps'_equiv' show ?thesis by fast
+qed
+
+lemma F_reachable_equiv:
+  "(\<exists> l' u u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0) \<and> F l')
+    \<longleftrightarrow> (\<exists> l' u u'. conv_A A \<turnstile> \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..n}. u c = 0) \<and> F l')"
+  if start_inv_check
+proof -
+  interpret start: Reachability_Problem_start A l0 n k
+    apply standard
+    using that unfolding start_inv_check .
+  from start.steps_steps'_equiv' show ?thesis by fast
+qed
+
+end
+
+(* This is obsolete now *)
+context TA_Start_No_Ceiling
+begin
+
+  context
+  begin
+
+  private abbreviation "k \<equiv> default_ceiling A"
+
+  lemma k_bound:
+    assumes "i > n"
+    shows "k i = 0"
+  proof -
+    from \<open>i > n\<close> have "i \<notin> {1..n}" by auto
+    then have
+      "i \<notin> clk_set A"
+      using clocks_n by fast
+    then have "{m. (i, m) \<in> Timed_Automata.clkp_set A} = {}" by auto
+    then show ?thesis unfolding default_ceiling_def by auto
+  qed
+
+  lemma k_0:
+   "k 0 = 0"
+    using clock_range unfolding default_ceiling_def by auto
+
+  lemma valid_abstraction:
+     "Timed_Automata.valid_abstraction A X k"
+    by (rule standard_abstraction_int;
+        force intro!: consts_nats clk_set_X[unfolded X_def] clk_set_X finite_X)
+
+  (* This is a bit clumsy *)
+  lemma k_ge:
+    "\<forall>(x,m) \<in> Timed_Automata.clkp_set A. m \<le> k x"
+    using valid_abstraction by (auto elim!: Timed_Automata.valid_abstraction.cases)
+
+  end
+
+end
+
+lemma mem_nth:
+  "x \<in> set xs \<Longrightarrow> \<exists> i < length xs. xs ! i = x"
+by (metis index_less_size_conv nth_index)
+
+lemma collect_clock_pairs_empty[simp]:
+  "collect_clock_pairs [] = {}"
+unfolding collect_clock_pairs_def by auto
+
+subsubsection \<open>Pre-compiled automata with states and clocks as natural numbers\<close>
+locale Reachability_Problem_precompiled_defs =
+  fixes n :: nat \<comment> \<open>Number of states. States are 0 through n - 1\<close>
+    and m :: nat \<comment> \<open>Number of clocks\<close>
+    and k :: "nat list list"
+    \<comment> \<open>Clock ceiling. Maximal constant appearing in automaton for each clock for each state\<close>
+    and inv :: "(nat, int) cconstraint list" \<comment> \<open>Clock invariants on states\<close>
+    and trans :: "((nat, int) cconstraint * nat list * nat) list list"
+        \<comment> \<open>Transitions between states\<close>
+    and final :: "nat list" \<comment> \<open>Final states. Initial location is 0\<close>
+begin
+
+  definition "clkp_set' l \<equiv>
+    collect_clock_pairs (inv ! l) \<union> \<Union> ((\<lambda> (g, _). collect_clock_pairs g) ` set (trans ! l))"
+  definition clk_set'_def: "clk_set' =
+    (\<Union> ((\<lambda> l. fst ` clkp_set' l) ` {0..<n}) \<union> \<Union> ((\<lambda> (g, r, _). set r) ` set (concat trans)))"
+
+  text \<open>Definition of the corresponding automaton\<close>
+  definition "label a \<equiv> \<lambda> (g, r, l'). (g, a, r, l')"
+  definition "I l \<equiv> if l < n then inv ! l else []"
+  definition "T \<equiv> {(l, label i (trans ! l ! i)) | l i. l < n \<and> i < length (trans ! l)}"
+  definition "A \<equiv> (T, I)"
+  definition "F \<equiv> \<lambda> x. x \<in> set final"
+
+end
+
+
+locale Reachability_Problem_precompiled = Reachability_Problem_precompiled_defs +
+  assumes inv_length: "length inv = n"
+    and trans_length: "length trans = n" (* "\<forall> xs \<in> set trans. length xs \<ge> n" *)
+    and state_set: "\<forall> xs \<in> set trans. \<forall> (_, _, l) \<in> set xs. l < n"
+    and k_length: "length k = n" "\<forall> l \<in> set k. length l = m + 1"
+    \<comment> \<open>Zero entry is just a dummy for the zero clock\<close>
+  assumes k_ceiling:
+    "\<forall> l < n. \<forall> (c, d) \<in> clkp_set' l. k ! l ! c \<ge> nat d" "\<forall> l < n. \<forall> c \<in> {1..m}. k ! l ! c \<ge> 0"
+    "\<forall> l < n. k ! l ! 0 = 0"
+    "\<forall> l < n. \<forall> (_, r, l') \<in> set (trans ! l). \<forall> c \<le> m. c \<notin> set r
+        \<longrightarrow> k ! l ! c \<ge> k ! l' ! c"
+  assumes consts_nats: "\<forall> l < n. snd ` clkp_set' l \<subseteq> \<nat>"
+  assumes clock_set: "clk_set' = {1..m}"
+    and m_gt_0: "m > 0"
+    and n_gt_0: "n > 0" and start_has_trans: "trans ! 0 \<noteq> []" \<comment> \<open>Necessary for refinement\<close>
+begin
+
+lemma consts_nats':
+  "\<forall> cc \<in> set inv. \<forall> (c, d) \<in> collect_clock_pairs cc. d \<in> \<nat>"
+  "\<forall> xs \<in> set trans. \<forall> (g, _) \<in> set xs. \<forall> (c, d) \<in> collect_clock_pairs g. d \<in> \<nat>"
+  using consts_nats unfolding clkp_set'_def
+   apply safe
+  using inv_length apply (auto dest!: mem_nth; fail)
+  using trans_length by - (drule mem_nth, auto)
+
+lemma clkp_set_simp_1:
+  "collect_clock_pairs (inv ! l) = collect_clki (inv_of A) l" if "l < n"
+  unfolding A_def inv_of_def collect_clki_def I_def[abs_def] using inv_length that by auto
+
+lemma clkp_set_simp_1':
+  "\<Union> (collect_clock_pairs ` set inv) = Timed_Automata.collect_clki (inv_of A)"
+  unfolding A_def inv_of_def Timed_Automata.collect_clki_def I_def[abs_def] using inv_length
+  by (auto simp add: collect_clks_id image_Union dest: nth_mem dest!: mem_nth)
+
+lemma clk_set_simp_2:
+  "\<Union> ((\<lambda> (g, r, _). set r) ` set (concat trans)) = collect_clkvt (trans_of A)"
+  unfolding A_def trans_of_def collect_clkvt_def T_def[abs_def] label_def using trans_length
+  apply (clarsimp simp add: image_Union dest: nth_mem)
+  apply safe
+   apply ((drule mem_nth)+, force)
+  apply (drule nth_mem, simp)
+  subgoal for _ a _ _ _ _ i
+    apply (rule bexI[where x = "trans ! a"])
+    apply (rule bexI[where x = "trans ! a ! i"])
+    by auto
+  done
+
+lemma clkp_set_simp_3:
+  "\<Union> ((\<lambda> (g, _). collect_clock_pairs g) ` set (trans ! l))
+      = collect_clkt (trans_of A) l" if "l < n"
+  unfolding A_def trans_of_def collect_clkt_def T_def[abs_def] label_def
+  using trans_length that
+  apply (auto simp add: collect_clks_id image_Union dest: nth_mem)
+  by (force dest!: mem_nth)
+
+lemma clkp_set_simp_3':
+  "\<Union> ((\<lambda> (g, _). Timed_Automata.collect_clock_pairs g) ` set (concat trans))
+      = Timed_Automata.collect_clkt (trans_of A)"
+  unfolding A_def trans_of_def Timed_Automata.collect_clkt_def T_def[abs_def] label_def
+  using trans_length
+  apply (auto simp add: collect_clks_id image_Union dest: nth_mem)
+   apply (auto dest!: mem_nth)[]
+   apply force
+  apply (auto dest!: nth_mem)
+  apply (rule_tac x = "trans ! aa" in bexI)
+   apply (rule_tac x = "trans ! aa ! i" in bexI)
+  by auto
+
+lemma clkp_set'_eq:
+  "clkp_set A l = clkp_set' l" if "l < n"
+  using that
+  by (fastforce simp: clkp_set'_def clkp_set_def image_Un image_Union
+      clkp_set_simp_1[symmetric] clkp_set_simp_3[symmetric]
+      )
+
+lemma clkp_set_out_of_bounds:
+  "clkp_set A l = {}" if "l \<ge> n" for l
+  using that
+  unfolding clkp_set_def collect_clki_def collect_clkt_def
+  unfolding inv_of_def A_def I_def T_def trans_of_def
+  by simp
+
+lemma clkp_setD:
+  "(x, d) \<in> clkp_set' l" if "(x, d) \<in> Closure.clkp_set A l"
+  using that by (cases "l < n") (auto simp: clkp_set'_eq clkp_set_out_of_bounds)
+
+lemma clkp_set_boundD:
+  "l < n" if "pair \<in> clkp_set A l" for pair
+  using that by - (rule ccontr, auto simp: clkp_set_out_of_bounds)
+
+lemma clkp_set'_eq':
+  "Timed_Automata.clkp_set A = \<Union> (clkp_set' ` {0..<n})"
+proof -
+  have "\<Union> (clkp_set A ` UNIV) = \<Union> (clkp_set A ` {0..<n})"
+    apply safe
+    subgoal for _ _ x
+      by (cases "x < n") (auto simp: clkp_set_out_of_bounds)
+    by auto
+  then show ?thesis by (simp add: TA_clkp_set_unfold clkp_set'_eq)
+qed
+
+lemma clk_set'_eq[simp]:
+  "clk_set A = clk_set'"
+  unfolding clkp_set'_eq' clk_set'_def clk_set_simp_2 by auto
+
+lemma Collect_fold_pair:
+  "{f a b | a b. P a b} = (\<lambda> (a, b). f a b) ` {(a, b). P a b}"
+  by auto
+
+lemma finite_T[intro, simp]:
+  "finite (trans_of A)"
+proof -
+  have
+    "{(l, i). l < n \<and> i < length (trans ! l)}
+      = \<Union> ((\<lambda> l. {(l, i) | i. i < length (trans ! l)}) ` {l. l < n})"
+    by auto
+  then show ?thesis unfolding trans_of_def A_def T_def by (auto simp: Collect_fold_pair)
+qed
+
+lemma has_clock[intro]:
+  "clk_set A \<noteq> {}"
+  using clock_set m_gt_0 by simp
+
+lemma clk_set:
+  "clk_set A = {1..m}"
+  using clock_set m_gt_0 by auto
+
+lemma clk_set_eq:
+  "clk_set A = {1..card (clk_set A)}"
+  using clk_set by auto
+
+lemma
+  "\<forall>c\<in>clk_set A. c \<le> card (clk_set A) \<and> c \<noteq> 0"
+  using clock_set by auto
+
+lemma clkp_set_consts_nat:
+  "\<forall>(_, d)\<in>Timed_Automata.clkp_set A. d \<in> \<nat>"
+  unfolding
+    Timed_Automata.clkp_set_def ta_collect_clki_alt_def ta_collect_clkt_alt_def
+    A_def trans_of_def inv_of_def
+    T_def I_def[abs_def] label_def
+  using consts_nats' inv_length trans_length by (force dest: nth_mem)
+
+lemma finite_range_inv_of_A[intro, simp]:
+  "finite (range (inv_of A))"
+  unfolding inv_of_def A_def I_def[abs_def] by (auto intro: finite_subset[where B = "{[]}"])
+
+lemma transD:
+  "(g, r, l') \<in> set (trans ! l) \<and> l < n" if "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  using that
+  unfolding trans_of_def A_def T_def label_def
+  by (auto dest!: nth_mem; simp split: prod.split_asm)
+
+lemma clkp_set_clk_bound:
+  "a \<le> m" if "(a, b) \<in> clkp_set' l" "l < n"
+  using that clock_set clk_set'_def
+  by fastforce
+
+sublocale TA_Start_No_Ceiling A 0 m
+  using has_clock clkp_set_consts_nat clk_set m_gt_0 by - (standard, auto)
+
+sublocale Reachability_Problem A 0 m "\<lambda> l i. if l < n \<and> i \<le> m then k ! l ! i else 0" "PR_CONST F"
+  apply standard
+  subgoal
+    apply safe
+    apply (frule clkp_set_boundD)
+    unfolding clkp_set'_eq
+    using k_ceiling
+    by (auto 4 10 dest: clkp_set_clk_bound)
+  subgoal
+    using k_ceiling(4) by (fastforce dest!: transD)
+  subgoal
+    using k_ceiling(2,3) by simp
+  subgoal
+    using k_ceiling(3) by simp
+  done
+
+end (* End of locale *)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Extra.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Extra.thy
@@ -0,0 +1,280 @@
+theory Normalized_Zone_Semantics_Impl_Extra
+  imports Normalized_Zone_Semantics_Impl
+begin
+
+context Regions
+begin
+
+lemma steps_z_norm'_induct[consumes 1, case_names refl step]:
+  assumes "A \<turnstile> \<langle>l0, D0\<rangle> \<leadsto>\<N>* \<langle>l', D'\<rangle>"
+  assumes "P l0 D0"
+  assumes "\<And>l D l' D' l'' D'' a. A \<turnstile> \<langle>l0, D0\<rangle> \<leadsto>\<N>* \<langle>l, D\<rangle> \<Longrightarrow> P l D
+  \<Longrightarrow> A \<turnstile> \<langle>l, D\<rangle> \<leadsto>v,n,\<tau> \<langle>l', D'\<rangle> \<Longrightarrow> A \<turnstile> \<langle>l', D'\<rangle> \<leadsto>\<N>(\<upharpoonleft>a) \<langle>l'', D''\<rangle> \<Longrightarrow> P l'' D''"
+  shows "P l' D'"
+  using assms
+  apply (induction "(l', D')" arbitrary: l' D' rule: rtranclp_induct)
+  apply clarsimp+
+  apply (fastforce simp: step_z_norm''_def)
+  done
+
+end
+
+context Reachability_Problem
+begin
+
+lemma step_z_norm''_cong[cong]:
+  "\<And>a b c d e f. step_z_norm'' a b c d e f \<equiv> step_z_norm'' a b c d e f" .
+
+lemma step_z_norm'_cong[cong]:
+  "\<And>a b c d e f. step_z_norm' a b c d e f \<equiv> step_z_norm' a b c d e f" .
+
+no_notation comp3  (infixl "\<circ>\<circ>\<circ>" 55)
+
+lemma step_impl_sound'':
+  assumes step: "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and     reachable: "E** a0 (l, D)"
+  shows
+    "\<exists> M'. step_z_norm' (conv_A A) l (curry (conv_M D)) a l' M'
+    \<and> [curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n = [M']v,n"
+  apply -
+  apply (rule step_impl_norm_sound[OF step])
+  using canonical_reachable reachable apply (simp only: check_diag_def neutral; blast)
+  using valid_dbm_reachable reachable diag_reachable' by auto
+
+lemma reachable_sound:
+  assumes "E** a0 (l', D')"
+  shows
+    "\<exists> M'. steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M' \<and> [curry (conv_M D')]v,n = [M']v,n"
+  using assms unfolding E_closure
+proof (induction l \<equiv> "l0" Z \<equiv> "init_dbm :: int DBM'" _ _ _ _ rule: steps_impl_induct)
+  case refl
+  show ?case
+    apply (rule exI[where x = "curry init_dbm"])
+    apply standard
+     apply blast
+    using RI_zone_equiv[symmetric] RI_init_dbm by fastforce
+next
+  case (step l' M' l'' M'' a l''' M''')
+  from step have reachable: "E** a0 (l', M')" unfolding E_closure by auto
+  note valid = valid_dbm_reachable[OF reachable]
+  from step.hyps(2) obtain D' where D':
+    "steps_z_norm' (conv_A A) l0 (curry init_dbm) l' D'" "[curry (conv_M M')]v,n = [D']v,n"
+    by auto
+  from step_impl_sound'[OF step(3)] canonical_reachable diag_reachable' valid reachable
+  obtain D'' where D'':
+    "(conv_A A) \<turnstile> \<langle>l', curry (conv_M M')\<rangle> \<leadsto>v,n,\<tau> \<langle>l'', D''\<rangle>" "[curry (conv_M M'')]v,n = [D'']v,n"
+    unfolding check_diag_def neutral by auto
+  from valid diag_reachable' canonical_reachable reachable have valid_M'':
+    "valid_dbm (curry (conv_M M''))"
+    apply -
+    apply (rule step_impl_valid_dbm)
+       apply (rule step(3))
+      apply assumption+
+    by auto
+  with canonical_reachable diag_reachable' valid reachable
+  obtain D''' where D''':
+    "step_z_norm' (conv_A A) l'' (curry (conv_M M'')) (\<upharpoonleft>a) l''' D'''"
+    "[curry (conv_M (FW' (norm_upd M''' (k' l''') n) n))]v,n = [D''']v,n"
+    using reachable_wf_dbm step.hyps(3) step_impl_wf_dbm wf_dbm_D
+    by (atomize_elim, intro step_impl_norm_sound[OF step(4)]) auto
+  from step_z_dbm_equiv'[OF D''(1) D'(2)] obtain M2 where M2:
+    "conv_A A \<turnstile> \<langle>l', D'\<rangle> \<leadsto>v,n,\<tau> \<langle>l'', M2\<rangle>" "[D'']v,n = [M2]v,n"
+    by blast
+  have "valid_dbm D'" using D'(1) steps_z_norm'_valid_dbm_preservation by blast
+  then have valid_M2: "valid_dbm M2" by (rule step_z_valid_dbm'[OF M2(1)])
+  from M2 D''(2) have "[curry (conv_M M'')]v,n = [M2]v,n" by simp
+  from step_z_norm_equiv'[OF D'''(1) valid_M'' valid_M2 this] obtain M3 where
+    "step_z_norm' (conv_A A) l'' M2 (\<upharpoonleft>a) l''' M3" "[D''']v,n = [M3]v,n"
+    by blast
+  with M2(1) D'(1) D'''(2) show ?case
+    by (fastforce simp: step_z_dbm_delay_loc step_z_norm''_def elim: rtranclp.rtrancl_into_rtrancl)
+qed
+
+lemma step_impl_complete2:
+  assumes step: "conv_A A \<turnstile> \<langle>l, curry (conv_M M)\<rangle> \<leadsto>v,n,a \<langle>l', D\<rangle>"
+    and wf_dbm: "wf_dbm M"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle> \<and> [curry (conv_M M')]v,n \<supseteq> [D]v,n"
+  using assms(1) wf_dbm_D[OF assms(2)] by (metis step_impl_complete''_improved subsetI wf_dbm)
+
+lemma reachable_complete:
+  assumes "steps_z_norm' (conv_A A) l0 (curry init_dbm) l' M'"
+  shows
+    "\<exists> D'. E** a0 (l', D') \<and> [M']v,n \<subseteq> [curry (conv_M D')]v,n"
+  using assms unfolding E_closure
+proof (induction rule: steps_z_norm'_induct)
+  case refl
+  show ?case
+    by (intro exI conjI) (rule steps_impl.refl, auto)
+next
+  case (step l' M' l'' M'' l''' M''' a)
+  then obtain D' where D':
+    "A \<turnstile>I \<langle>l0, init_dbm\<rangle> \<leadsto>k',n* \<langle>l', D'\<rangle>" "[M']v,n \<subseteq> [curry (conv_M D')]v,n"
+    by fast
+  from step_z_dbm_mono[OF step(2) D'(2)] obtain D'' where D'':
+    "conv_A A \<turnstile> \<langle>l', curry (conv_M D')\<rangle> \<leadsto>v,n,\<tau> \<langle>l'', D''\<rangle>"
+    "[M'']v,n \<subseteq> [D'']v,n"
+    by auto
+  from step have reachable: "E** a0 (l', D')" using D'(1) E_closure by blast (* s/h *)
+  from step_impl_complete2[OF D''(1) reachable_wf_dbm[OF this]] step obtain D2 where D2:
+    "A \<turnstile>I \<langle>l', D'\<rangle> \<leadsto>n,\<tau> \<langle>l'', D2\<rangle>" "[D'']v,n \<subseteq> [curry (conv_M D2)]v,n"
+    by auto
+  have "valid_dbm M''" (* s/h *)
+    using step(1,2) step_z_valid_dbm' steps_z_norm'_valid_dbm_preservation valid_init_dbm by blast
+  have valid3:"valid_dbm (curry (conv_M D2))"
+    apply (rule step_impl_valid_dbm[OF D2(1)])
+    using valid_dbm_reachable canonical_reachable reachable
+    using diag_reachable[OF reachable] diag_conv by (auto intro!: canonical_reachable)
+  from step_z_norm_mono'[OF step(3) \<open>valid_dbm M''\<close> valid3] D''(2) D2(2) obtain M3 where M3:
+    "step_z_norm' (conv_A A) l'' (curry (conv_M D2)) (\<upharpoonleft>a) l''' M3" "[M''']v,n \<subseteq> [M3]v,n"
+    by auto
+  have canonical: "canonical (curry (conv_M D2)) n \<or> check_diag n D2"
+    using step_impl_canonical[OF D2(1) diag_reachable[OF reachable]]
+    unfolding check_diag_def neutral by auto
+  have diag: "\<forall>i\<le>n. conv_M D2 (i, i) \<le> 0"
+    using step_impl_diag_preservation[OF D2(1) diag_reachable[OF reachable]] diag_conv by auto
+  from step_impl_norm_complete''[OF M3(1) valid3 canonical diag] obtain D3 where
+    "A \<turnstile>I \<langle>l'', D2\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l''', D3\<rangle> "
+    "[M3]v,n \<subseteq> [curry (conv_M (FW' (norm_upd D3 (k' l''') n) n))]v,n"
+    by auto
+  with D'(1) D2 M3(2) show ?case
+     by (blast intro: steps_impl.step)
+qed
+
+lemma step_impl_mono_reachable:
+    assumes "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+    and     "wf_dbm D" "wf_dbm M"
+    and "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+  shows "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle> \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+  proof -
+    note prems_D = wf_dbm_D[OF assms(2)]
+    note prems_M = wf_dbm_D[OF assms(3)]
+    from step_impl_sound'[OF assms(1) prems_D] diag_conv obtain M' where
+      "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,a \<langle>l', M'\<rangle>"
+      "[curry (conv_M D')]v,n = [M']v,n"
+      unfolding check_diag_def neutral by auto
+    with step_impl_complete2[OF _ assms(3)] step_z_dbm_mono[OF _ assms(4)] show ?thesis by force
+  qed
+
+lemma step_impl_norm_mono_reachable:
+    assumes "A \<turnstile>I \<langle>l,D\<rangle> \<leadsto>n,a \<langle>l',D'\<rangle>"
+        and prems_D: "wf_dbm D"
+        and prems_M: "wf_dbm M"
+    and subs: "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+    shows
+      "\<exists> M'. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,a \<langle>l', M'\<rangle>
+      \<and> [curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n
+      \<subseteq> [curry (conv_M (FW' (norm_upd M' (k' l') n) n))]v,n"
+  proof -
+    note prems_D = wf_dbm_D[OF prems_D]
+    note prems_M = wf_dbm_D[OF prems_M]
+    from step_impl_norm_sound[OF assms(1) prems_D] obtain M' where M':
+      "step_z_norm' (conv_A A) l (curry (conv_M D)) a l' M'"
+      "[curry (conv_M (FW' (norm_upd D' (k' l') n) n))]v,n = [M']v,n"
+      by auto
+    from
+      step_z_norm_mono'[OF this(1) prems_D(3) prems_M(3) subs]
+      step_impl_norm_complete''[OF _ prems_M(3,1,2)] M'(2)
+    show ?thesis by fast
+  qed
+
+lemma step_impl_mono_reachable':
+    assumes step:
+      "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', D'\<rangle>"
+      "A \<turnstile>I \<langle>l', D'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D''\<rangle>"
+      "D''' = FW' (norm_upd D'' (k' l'') n) n"
+    and wf_dbm: "wf_dbm D" "wf_dbm M"
+    and subset: "dbm_subset n D M"
+    shows
+      "\<exists> M' M'' M'''. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle> \<and> A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>
+       \<and> M''' = FW' (norm_upd M'' (k' l'') n) n \<and> dbm_subset n D''' M'''"
+  proof -
+    from dbm_subset_correct''[OF wf_dbm] subset[THEN dbm_subset_conv] have
+      "([curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n)"
+      by simp
+    from step_impl_mono_reachable[OF step(1) wf_dbm this] obtain M' where M':
+      "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle>" "[curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+      by auto
+    from step_impl_wf_dbm[OF step(1) wf_dbm(1)] step_impl_wf_dbm[OF M'(1) wf_dbm(2)] have wf_dbm':
+      "wf_dbm D'" "wf_dbm M'" by auto
+    from step_impl_norm_mono_reachable[OF step(2) wf_dbm' M'(2)] obtain M'' where M'':
+      "A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>"
+      "[curry (conv_M (FW' (norm_upd D'' (k' l'') n) n))]v,n
+      \<subseteq> [curry (conv_M (FW' (norm_upd M'' (k' l'') n) n))]v,n"
+      using diag_conv by auto
+    from step_impl_wf_dbm[OF step(2)] step_impl_wf_dbm[OF M''(1)] wf_dbm' have
+      "wf_dbm D''" "wf_dbm M''"
+      by auto
+    from norm_step_wf_dbm[OF this(1)] norm_step_wf_dbm[OF this(2)] have
+      "wf_dbm (FW' (norm_upd D'' (k' l'') n) n)" "wf_dbm (FW' (norm_upd M'' (k' l'') n) n)"
+      by auto
+    from M'(1) dbm_subset_correct''[OF this] M''(2) dbm_subset_conv_rev have
+      "dbm_subset n (FW' (norm_upd D'' (k' l'') n) n) (FW' (norm_upd M'' (k' l'') n) n)"
+      by auto
+    with M'(1) M''(1) show ?thesis unfolding \<open>D''' = _\<close> by auto
+  qed
+
+lemma step_impl_mono_reachable'':
+    assumes step:
+      "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', D'\<rangle>"
+      "A \<turnstile>I \<langle>l', D'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D''\<rangle>"
+      "D''' = FW' (norm_upd D'' (k' l'') n) n"
+    and wf_dbm: "wf_dbm D" "wf_dbm M"
+    and subset: "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+    shows
+      "\<exists> M' M'' M'''. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle> \<and> A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>
+       \<and> M''' = FW' (norm_upd M'' (k' l'') n) n
+       \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n
+       \<and> [curry (conv_M D''')]v,n \<subseteq> [curry (conv_M M''')]v,n"
+  proof -
+    from step_impl_mono_reachable[OF step(1) wf_dbm subset] obtain M' where M':
+      "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle>" "[curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+      by auto
+    from step_impl_wf_dbm[OF step(1) wf_dbm(1)] step_impl_wf_dbm[OF M'(1) wf_dbm(2)] have wf_dbm':
+      "wf_dbm D'" "wf_dbm M'" by auto
+    from step_impl_norm_mono_reachable[OF step(2) wf_dbm' M'(2)] obtain M'' where M'':
+      "A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>"
+      "[curry (conv_M (FW' (norm_upd D'' (k' l'') n) n))]v,n
+      \<subseteq> [curry (conv_M (FW' (norm_upd M'' (k' l'') n) n))]v,n"
+      using diag_conv by auto
+    with M' M''(1) show ?thesis unfolding \<open>D''' = _\<close> by auto
+  qed
+
+lemma step_impl_mono_reachable__:
+  assumes step:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', D'\<rangle>"
+    "A \<turnstile>I \<langle>l', D'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', D''\<rangle>"
+    "D''' = FW' (norm_upd D'' (k' l'') n) n"
+    and wf_dbm: "wf_dbm D" "wf_dbm M"
+    and subset: "dbm_subset n D M"
+  shows
+    "\<exists> M' M'' M'''. A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle> \<and> A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>
+     \<and> M''' = FW' (norm_upd M'' (k' l'') n) n \<and> dbm_subset n D''' M'''"
+proof -
+  from dbm_subset_correct''[OF wf_dbm] subset[THEN dbm_subset_conv] have
+    "([curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n)"
+    by simp
+  from step_impl_mono_reachable[OF step(1) wf_dbm this] obtain M' where M':
+    "A \<turnstile>I \<langle>l, M\<rangle> \<leadsto>n,\<tau> \<langle>l', M'\<rangle>" "[curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+    by auto
+  from step_impl_wf_dbm[OF step(1) wf_dbm(1)] step_impl_wf_dbm[OF M'(1) wf_dbm(2)] have wf_dbm':
+    "wf_dbm D'" "wf_dbm M'" by auto
+  from step_impl_norm_mono_reachable[OF step(2) wf_dbm' M'(2)] obtain M'' where M'':
+    "A \<turnstile>I \<langle>l', M'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', M''\<rangle>"
+    "[curry (conv_M (FW' (norm_upd D'' (k' l'') n) n))]v,n
+    \<subseteq> [curry (conv_M (FW' (norm_upd M'' (k' l'') n) n))]v,n"
+    using diag_conv by auto
+  from step_impl_wf_dbm[OF step(2)] step_impl_wf_dbm[OF M''(1)] wf_dbm' have
+    "wf_dbm D''" "wf_dbm M''"
+    by auto
+  from norm_step_wf_dbm[OF this(1)] norm_step_wf_dbm[OF this(2)] have
+    "wf_dbm (FW' (norm_upd D'' (k' l'') n) n)" "wf_dbm (FW' (norm_upd M'' (k' l'') n) n)"
+    by auto
+  from M'(1) dbm_subset_correct''[OF this] M''(2) dbm_subset_conv_rev have
+    "dbm_subset n (FW' (norm_upd D'' (k' l'') n) n) (FW' (norm_upd M'' (k' l'') n) n)"
+    by auto
+  with M'(1) M''(1) show ?thesis unfolding \<open>D''' = _\<close> by auto
+qed
+
+end
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Refine.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Refine.thy
@@ -0,0 +1,1797 @@
+theory Normalized_Zone_Semantics_Impl_Refine
+  imports
+    Normalized_Zone_Semantics_Impl Difference_Bound_Matrices.DBM_Operations_Impl_Refine
+    "HOL-Library.IArray"
+    HOL.Code_Numeral
+    Worklist_Algorithms.Worklist_Subsumption_Impl1 Worklist_Algorithms.Unified_PW_Impl
+    Worklist_Algorithms.Liveness_Subsumption_Impl Worklist_Algorithms.Leadsto_Impl
+    Normalized_Zone_Semantics_Impl_Semantic_Refinement
+    Munta_Base.Printing Show.Show_Instances
+    Munta_Base.Abstract_Term
+begin
+
+  unbundle no_library_syntax
+
+  notation fun_rel_syn (infixr "\<rightarrow>" 60)
+
+  section \<open>Imperative Implementation of Reachability Checking\<close>
+
+  subsection \<open>Misc\<close>
+
+  lemma (in -) rtranclp_equiv:
+    "R** x y \<longleftrightarrow> S** x y" if "\<And> x y. P x \<Longrightarrow> R x y \<longleftrightarrow> S x y" "\<And> x y. P x \<Longrightarrow> R x y \<Longrightarrow> P y" "P x"
+  proof
+    assume A: "R** x y"
+    note that(1)[iff] that(2)[intro]
+    from A \<open>P x\<close> have "P y \<and> S** x y"
+      by induction auto
+    then show "S** x y" ..
+  next
+    assume A: "S** x y"
+    note that(1)[iff] that(2)[intro] rtranclp.intros(2)[intro]
+    from A \<open>P x\<close> have "P y \<and> R** x y"
+      by (induction; blast)
+    then show "R** x y" ..
+  qed
+
+  lemma (in -) tranclp_equiv:
+    "R++ x y \<longleftrightarrow> S++ x y" if "\<And> x y. P x \<Longrightarrow> R x y \<longleftrightarrow> S x y" "\<And> x y. P x \<Longrightarrow> R x y \<Longrightarrow> P y" "P x"
+  proof
+    assume A: "R++ x y"
+    note that(1)[iff] that(2)[intro]
+    from A \<open>P x\<close> have "P y \<and> S++ x y"
+      by induction auto
+    then show "S++ x y" ..
+  next
+    assume A: "S++ x y"
+    note that(1)[iff] that(2)[intro] tranclp.intros(2)[intro]
+    from A \<open>P x\<close> have "P y \<and> R++ x y"
+      by (induction; blast)
+    then show "R++ x y" ..
+  qed
+
+  lemma (in -) rtranclp_tranclp_equiv:
+    "R** x y \<and> R++ y z \<longleftrightarrow> S** x y \<and> S++ y z" if
+    "\<And> x y. P x \<Longrightarrow> R x y \<longleftrightarrow> S x y" "\<And> x y. P x \<Longrightarrow> R x y \<Longrightarrow> P y" "P x"
+  proof
+    assume A: "R** x y \<and> R++ y z"
+    note that(1)[iff] that(2)[intro]
+    from A[THEN conjunct1] \<open>P x\<close> have "P y"
+      by induction auto
+    then show "S** x y \<and> S++ y z"
+      using rtranclp_equiv[of P R S x y, OF that] tranclp_equiv[of P R S y z, OF that(1,2)] A
+      by fastforce
+  next
+    assume A: "S** x y \<and> S++ y z"
+    note that(1)[iff] that(2)[intro]
+    from A[THEN conjunct1] \<open>P x\<close> have "P y"
+      by induction auto
+    then show "R** x y \<and> R++ y z"
+      using rtranclp_equiv[of P R S x y, OF that] tranclp_equiv[of P R S y z, OF that(1,2)] A
+      by force
+  qed
+
+
+  subsection \<open>Mapping Transitions and Invariants\<close>
+
+  type_synonym
+    ('a, 'c, 'time, 's) transition_fun = "'s \<Rightarrow> (('c, 'time) cconstraint * 'a * 'c list * 's) list"
+
+  definition transition_rel where
+    "transition_rel X = {(f, r). \<forall> l \<in> X. \<forall> x. (l, x) \<in> r \<longleftrightarrow> x \<in> set (f l)}"
+
+  definition inv_rel where \<comment> \<open>Invariant assignments for a restricted state set\<close>
+    (* Should it be "inv_rel X = Id_on X \<rightarrow> Id" ? *)
+    "inv_rel R X = b_rel R (\<lambda> x. x \<in> X) \<rightarrow> Id"
+
+  lemma transition_rel_anti_mono:
+  "transition_rel S \<subseteq> transition_rel R" if "R \<subseteq> S"
+  using that unfolding transition_rel_def by auto
+
+  lemma inv_rel_anti_mono:
+    "inv_rel A S \<subseteq> inv_rel A R" if "R \<subseteq> S"
+    using that unfolding inv_rel_def fun_rel_def b_rel_def by auto
+
+  definition state_set :: "('a, 'c, 'time, 's) transition set \<Rightarrow> 's set" where
+    "state_set T = fst ` T \<union> (snd o snd o snd o snd) ` T"
+
+definition
+  "trace_level (i :: int) (f :: unit \<Rightarrow> String.literal Heap) = ()"
+
+locale Show_State_Defs =
+  fixes n :: nat and show_state :: "'si \<Rightarrow> string" and show_clock :: "nat \<Rightarrow> string"
+begin
+
+definition
+  "tracei type \<equiv> \<lambda> (l, M).
+   let _ = trace_level 5 (
+    \<lambda>_. do {
+      let st = show_state l;
+      m \<leftarrow> show_dbm_impl n show_clock show M;
+      let s = type @ '': (''  @ st @ '', <'' @ m @ ''>)''; 
+      let s = String.implode s;
+      return s
+    })
+  in return ()
+"
+
+end
+
+locale TA_Impl_No_Ceiling_Defs =
+  Show_State_Defs n show_state +
+  TA_Start_No_Ceiling_Defs A l0 n
+  for show_state :: "'si :: {hashable, heap} \<Rightarrow> string"
+  and A :: "('a, nat, int, 's) ta" and l0 :: 's and n :: nat +
+
+  fixes trans_fun :: "('a, nat, int, 's) transition_fun"
+    and inv_fun :: "(nat, int, 'si :: {hashable, heap}) invassn"
+    and trans_impl :: "('a, nat, int, 'si) transition_fun"
+    and l0i :: 'si
+begin
+
+abbreviation "states \<equiv> {l0} \<union> (state_set (trans_of A))"
+
+end
+
+locale TA_Impl_Defs =
+  TA_Start_Defs A l0 n k +
+  TA_Impl_No_Ceiling_Defs show_clock show_state A l0 n trans_fun inv_fun trans_impl l0i
+  for show_clock and show_state :: "'si :: {hashable, heap} \<Rightarrow> string"
+  and A :: "('a, nat, int, 's) ta" and l0 :: 's and n :: nat and k
+  and trans_fun inv_fun trans_impl l0i
+  +
+  fixes ceiling :: "'si \<Rightarrow> int iarray"
+begin
+
+definition "inv_of_A = inv_of A"
+
+end
+
+locale Reachability_Problem_Impl_Defs =
+  Reachability_Problem_Defs A l0 n k F +
+  TA_Impl_No_Ceiling_Defs show_clock show_state A l0 n trans_fun inv_fun trans_impl l0i
+  for show_clock and show_state :: "'si :: {hashable, heap} \<Rightarrow> string"
+  and A :: "('a, nat, int, 's) ta" and l0 :: 's and  n :: nat and k and F :: "'s \<Rightarrow> bool"
+  and trans_fun inv_fun trans_impl and l0i
+  +
+  fixes F_fun :: "'si \<Rightarrow> bool"
+begin
+
+end (* Reachability Problem Impl Defs *)
+
+definition "FWI'' n M = FWI' M n"
+
+lemma fwi_impl'_refine_FWI'':
+  "(fwi_impl' n, RETURN oo PR_CONST (\<lambda> M. FWI'' n M)) \<in> Id \<rightarrow> Id \<rightarrow> \<langle>Id\<rangle> nres_rel"
+  unfolding FWI''_def by (rule fwi_impl'_refine_FWI')
+
+lemmas fwi_impl_refine_FWI'' = fwi_impl.refine[FCOMP fwi_impl'_refine_FWI'']
+
+context
+  fixes n :: nat
+begin
+
+  context
+    notes [map_type_eqs] = map_type_eqI[of "TYPE(nat * nat \<Rightarrow> 'e)" "TYPE('e i_mtx)"]
+  begin
+
+  sepref_register "PR_CONST (FWI'' n)"
+
+  end
+
+  lemmas [sepref_fr_rules] = fwi_impl_refine_FWI''
+
+  lemma [def_pat_rules]: "FWI'' $ n \<equiv> UNPROTECT (FWI'' n)" by simp
+
+  sepref_definition repair_pair_impl is
+    "uncurry2 (RETURN ooo PR_CONST (repair_pair n))" ::
+    "[\<lambda> ((_,a),b). a \<le> n \<and> b \<le> n]a (mtx_assn n)d *a nat_assnk *a nat_assnk \<rightarrow> mtx_assn n"
+    unfolding repair_pair_def FWI''_def[symmetric] PR_CONST_def
+    by sepref
+
+  sepref_register repair_pair
+
+  lemmas [sepref_fr_rules] = repair_pair_impl.refine
+
+end
+
+locale TA_Impl_No_Ceiling =
+  TA_Impl_No_Ceiling_Defs _ _ A l0 n trans_fun inv_fun trans_impl l0i +
+  TA_Start_No_Ceiling A l0 n
+  for A :: "('a, nat, int, 's) ta"
+  and l0 :: 's
+  and n :: nat
+  and trans_fun inv_fun
+  and trans_impl :: "('a, nat, int, 'si :: {hashable, heap}) transition_fun"
+  and l0i
+  +
+  fixes states' and loc_rel :: "('si \<times> 's) set"
+  assumes trans_fun: "(trans_fun, trans_of A) \<in> transition_rel states'"
+      and trans_impl:
+        "(trans_impl, trans_fun) \<in> fun_rel_syn loc_rel (list_rel (Id \<times>r Id \<times>r Id \<times>r loc_rel))"
+      and inv_fun: "(inv_fun, inv_of A) \<in> inv_rel loc_rel states'"
+      and states'_states: "states \<subseteq> states'"
+      and init_impl: "(l0i, l0) \<in> loc_rel"
+      and loc_rel_left_unique:
+        "\<And>l li li'. l \<in> states' \<Longrightarrow> (li, l) \<in> loc_rel \<Longrightarrow> (li', l) \<in> loc_rel \<Longrightarrow> li' = li"
+      and loc_rel_right_unique:
+        "\<And>l l' li. l \<in> states' \<Longrightarrow> l' \<in> states' \<Longrightarrow> (li,l) \<in> loc_rel \<Longrightarrow> (li,l') \<in> loc_rel
+          \<Longrightarrow> l' = l"
+begin
+
+lemma trans_fun':
+  "(trans_fun, trans_of A) \<in> transition_rel states"
+  using transition_rel_anti_mono[OF states'_states] trans_fun by blast
+
+lemma inv_fun': "(inv_fun, inv_of A) \<in> inv_rel loc_rel states"
+  using inv_fun inv_rel_anti_mono[OF states'_states] by blast
+
+abbreviation "location_rel \<equiv> b_rel loc_rel (\<lambda> x. x \<in> states')"
+
+abbreviation "location_assn \<equiv> pure location_rel"
+
+\<comment> \<open>Could probably be remove altogether.\<close>
+abbreviation "state_assn \<equiv> prod_assn id_assn (mtx_assn n)"
+
+abbreviation "state_assn' \<equiv> prod_assn location_assn (mtx_assn n)"
+
+interpretation DBM_Impl n .
+
+abbreviation
+  "op_impl_assn \<equiv>
+  location_assnk *a (list_assn clock_assn)k *a
+  (list_assn (acconstraint_assn clock_assn id_assn))k *a location_assnk *a mtx_assnd \<rightarrow>a mtx_assn"
+
+lemma tracei_refine:
+  "(uncurry tracei, uncurry (\<lambda>_ _. RETURN ())) \<in> id_assnk *a state_assn'k \<rightarrow>a unit_assn"
+    unfolding tracei_def
+    using show_dbm_impl.refine[to_hnr, unfolded hn_refine_def, of n]
+    by sepref_to_hoare sep_auto
+
+end
+
+locale TA_Impl =
+  TA_Impl_Defs _ _ A l0 n k trans_fun inv_fun trans_impl l0i ceiling +
+  TA_Impl_No_Ceiling _ _ A l0 n trans_fun inv_fun trans_impl l0i states' loc_rel +
+  TA_Start A l0 n k
+  for A :: "('a, nat, int, 's) ta"
+  and l0 :: 's
+  and n :: nat
+  and k
+  and trans_fun inv_fun
+  and trans_impl :: "('a, nat, int, 'si :: {hashable, heap}) transition_fun"
+  and l0i
+  and ceiling
+  and states' and loc_rel :: "('si \<times> 's) set"
+  +
+  assumes ceiling: "(ceiling, IArray o k') \<in> inv_rel loc_rel states'"
+begin
+
+lemma ceiling': "(ceiling, IArray o k') \<in> inv_rel loc_rel states"
+  using ceiling inv_rel_anti_mono[OF states'_states] by blast
+
+end
+
+locale Reachability_Problem_Impl =
+  Reachability_Problem_Impl_Defs _ _ A l0 n k F trans_fun inv_fun trans_impl l0i +
+  TA_Impl _ _ A l0 n k trans_fun inv_fun trans_impl l0i ceiling states' loc_rel +
+  Reachability_Problem A l0 n k F
+  for A :: "('a, nat, int, 's) ta"
+  and l0 :: 's
+  and n :: nat
+  and k
+  and F
+  and trans_fun inv_fun
+  and trans_impl :: "('a, nat, int, 'si :: {hashable, heap}) transition_fun"
+  and l0i
+  and ceiling
+  and states' and loc_rel :: "('si \<times> 's) set" +
+  assumes F_fun: "(F_fun, F) \<in> inv_rel loc_rel states'"
+begin
+
+lemma F_fun':  "(F_fun, F) \<in> inv_rel loc_rel states"
+  using F_fun inv_rel_anti_mono[OF states'_states] by blast
+
+end
+
+context Search_Space_finite
+begin
+
+sublocale liveness_search_space_pre:
+  Liveness_Search_Space_pre "\<lambda> x y. E x y \<and> F x \<and> F y \<and> \<not> empty y" a0 F "(\<preceq>)"
+  "\<lambda> v. reachable v \<and> \<not> empty v \<and> F v"
+  using finite_reachable apply -
+  apply (standard)
+      apply (blast intro: finite_subset[rotated] F_mono trans)
+     apply (blast intro: finite_subset[rotated] F_mono trans)
+  subgoal
+    by safe (blast dest: mono F_mono empty_mono)
+   apply (blast intro: finite_subset[rotated] F_mono trans)
+  apply (blast intro: finite_subset[rotated] F_mono trans)
+  done
+
+end
+
+locale TA_Impl_Op =
+  TA_Impl _ _ A l0 n k trans_fun inv_fun trans_impl l0i ceiling states' loc_rel
+  + op: E_From_Op_Bisim A l0 n k f
+  for A and l0 :: 's and n k
+  and f trans_fun inv_fun trans_impl and l0i :: "'si:: {hashable,heap}"
+  and ceiling  states' loc_rel
+  +
+  fixes op_impl
+  assumes op_impl: "(uncurry4 op_impl, uncurry4 (\<lambda> l r. RETURN ooo f l r)) \<in> op_impl_assn"
+
+(*
+locale Reachability_Problem_Impl_Op =
+  TA_Impl_Op _ _ A l0 n k f trans_fun inv_fun trans_impl l0i ceiling states' loc_rel
++ Reachability_Problem_Impl _ _ _ A l0 n k _ trans_fun inv_fun trans_impl l0i ceiling states' loc_rel
+  for A and l0 :: 's and n k
+  and f trans_fun inv_fun trans_impl and l0i :: "'si:: {hashable,heap}"
+  and ceiling  states' loc_rel
+*)
+
+locale Reachability_Problem_Impl_Op =
+  TA_Impl_Op _ _ A l0 n k f trans_fun inv_fun trans_impl l0i ceiling states' loc_rel
++ Reachability_Problem_Impl _ _ _ A l0 n k _ trans_fun inv_fun trans_impl l0i ceiling states' loc_rel
++ op: E_From_Op_Bisim_Finite_Reachability A l0 n k f
+  for A and l0 :: 's and n k
+  and f trans_fun inv_fun trans_impl and l0i :: "'si:: {hashable,heap}"
+  and ceiling  states' loc_rel
+
+
+subsection \<open>Implementing of the Successor Function\<close>
+
+paragraph \<open>Shared Setup\<close>
+
+context TA_Impl
+begin
+
+  interpretation DBM_Impl n .
+
+  context
+    notes [map_type_eqs] = map_type_eqI[of "TYPE(nat * nat \<Rightarrow> 'e)" "TYPE('e i_mtx)"]
+  begin
+
+  sepref_register trans_impl
+
+  sepref_register abstr_upd up_canonical_upd norm_upd reset'_upd reset_canonical_upd
+
+  sepref_register "PR_CONST (FW'' n)"
+
+  sepref_register "PR_CONST (inv_of_A)"
+
+  end
+
+  lemmas [sepref_fr_rules] = fw_impl.refine[FCOMP fw_impl'_refine_FW'']
+
+  lemma [def_pat_rules]: "FW'' $ n \<equiv> UNPROTECT (FW'' n)" by simp
+
+  lemma trans_impl_clock_bounds3:
+    "\<forall> c \<in> collect_clks (inv_of A l). c \<le> n"
+  using collect_clks_inv_clk_set[of A l] clocks_n by force
+
+  lemma inv_fun_rel:
+    assumes "l \<in> states'" "(l1, l) \<in> loc_rel"
+    shows "(inv_fun l1, inv_of A l) \<in> \<langle>\<langle>nbn_rel (Suc n), int_rel\<rangle>acconstraint_rel\<rangle>list_rel"
+  proof -
+    have "(xs, xs) \<in> \<langle>\<langle>nbn_rel (Suc n), int_rel\<rangle>acconstraint_rel\<rangle>list_rel"
+      if "\<forall> c \<in> collect_clks xs. c \<le> n" for xs
+    using that
+      apply (induction xs)
+      apply simp
+      apply simp
+      subgoal for a
+        apply (cases a)
+      by (auto simp: acconstraint_rel_def p2rel_def rel2p_def)
+    done
+    moreover have
+      "inv_fun l1 = inv_of A l"
+    using inv_fun assms unfolding inv_rel_def b_rel_def fun_rel_def by auto
+    ultimately show ?thesis using trans_impl_clock_bounds3 by auto
+  qed
+
+  lemma [sepref_fr_rules]:
+    "(return o inv_fun, RETURN o PR_CONST inv_of_A)
+    \<in> location_assnk \<rightarrow>a list_assn (acconstraint_assn clock_assn int_assn)"
+    using inv_fun_rel
+   apply (simp add: b_assn_pure_conv aconstraint_assn_pure_conv list_assn_pure_conv inv_of_A_def)
+  by sepref_to_hoare (sep_auto simp: pure_def)
+
+  lemma [sepref_fr_rules]:
+    "(return o ceiling, RETURN o PR_CONST k') \<in> location_assnk \<rightarrow>a iarray_assn"
+    using ceiling by sepref_to_hoare (sep_auto simp: inv_rel_def fun_rel_def br_def)
+
+  sepref_register "PR_CONST k'"
+
+  lemma [def_pat_rules]: "(TA_Start_Defs.k' $ n $ k) \<equiv> PR_CONST k'" by simp
+
+  lemma [simp]:
+    "length (k' l) > n"
+  by (simp add: k'_def)
+
+  lemma [def_pat_rules]: "(TA_Impl_Defs.inv_of_A $ A) \<equiv> PR_CONST inv_of_A" by simp
+
+  lemma reset_canonical_upd_impl_refine:
+    "(uncurry3 (reset_canonical_upd_impl n), uncurry3 (\<lambda>x. RETURN \<circ>\<circ>\<circ> reset_canonical_upd x))
+      \<in> [\<lambda>(((_, i), j), _).
+             i \<le> n \<and>
+             j \<le> n]a mtx_assnd *a nat_assnk *a clock_assnk *a id_assnk \<rightarrow> mtx_assn"
+    apply sepref_to_hnr
+    apply (rule hn_refine_preI)
+    apply (rule hn_refine_cons[OF _ reset_canonical_upd_impl.refine[to_hnr], simplified])
+    by (sep_auto simp: hn_ctxt_def b_assn_def entailst_def)+
+
+  lemmas [sepref_fr_rules] =
+    abstr_upd_impl.refine up_canonical_upd_impl.refine norm_upd_impl.refine
+    reset_canonical_upd_impl_refine
+
+  lemma constraint_clk_acconstraint_rel_simps:
+    assumes "(ai, a) \<in> \<langle>nbn_rel (Suc n), id_rel\<rangle>acconstraint_rel"
+    shows "constraint_clk a < Suc n" "constraint_clk ai = constraint_clk a"
+    using assms by (cases ai; cases a; auto simp: acconstraint_rel_def p2rel_def rel2p_def)+
+
+  lemma [sepref_fr_rules]:
+    "(return o constraint_clk, RETURN o constraint_clk)
+     \<in> (acconstraint_assn clock_assn id_assn)k \<rightarrow>a clock_assn"
+    apply (simp add: b_assn_pure_conv aconstraint_assn_pure_conv)
+    by sepref_to_hoare (sep_auto simp: pure_def constraint_clk_acconstraint_rel_simps)
+
+  sepref_register constraint_clk
+
+  sepref_register
+    "repair_pair n :: ((nat \<times> nat \<Rightarrow> ('b :: {linordered_cancel_ab_monoid_add}) DBMEntry)) \<Rightarrow> _" ::
+    "(('ee DBMEntry) i_mtx) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ((nat \<times> nat \<Rightarrow> 'ee DBMEntry))"
+
+  sepref_register abstra_upd :: "(nat, 'e) acconstraint \<Rightarrow> 'e DBMEntry i_mtx \<Rightarrow> 'e DBMEntry i_mtx"
+
+  sepref_register n
+
+  lemmas [sepref_import_param] = IdI[of n]
+
+  sepref_definition abstra_repair_impl is
+    "uncurry (RETURN oo PR_CONST abstra_repair)" ::
+    "(acconstraint_assn clock_assn id_assn)k *a mtx_assnd \<rightarrow>a mtx_assn"
+    unfolding abstra_repair_def PR_CONST_def by sepref
+
+  sepref_register abstra_repair :: "(nat, 'e) acconstraint \<Rightarrow> 'e DBMEntry i_mtx \<Rightarrow> 'e DBMEntry i_mtx"
+
+  lemmas [sepref_fr_rules] = abstra_repair_impl.refine
+
+  sepref_definition abstr_repair_impl is
+    "uncurry (RETURN oo PR_CONST abstr_repair)" ::
+    "(list_assn (acconstraint_assn clock_assn id_assn))k *a mtx_assnd \<rightarrow>a mtx_assn"
+    unfolding abstr_repair_def PR_CONST_def by sepref
+
+  sepref_register abstr_repair ::
+    "(nat, 'e) acconstraint list \<Rightarrow> 'e DBMEntry i_mtx \<Rightarrow> 'e DBMEntry i_mtx"
+
+  lemmas [sepref_fr_rules] = abstr_repair_impl.refine
+
+end (* TA Impl *)
+
+paragraph \<open>Implementation for an Arbitrary DBM Successor Operation\<close>
+
+context Reachability_Problem_Impl
+begin
+
+lemma F_rel_alt_def:
+  "F_rel = (\<lambda> (l, M). F l  \<and> \<not> check_diag n M)"
+unfolding F_rel_def by auto
+
+sepref_register F
+
+lemma [sepref_fr_rules]:
+  "(return o F_fun, RETURN o F) \<in> location_assnk \<rightarrow>a id_assn"
+using F_fun by sepref_to_hoare (sep_auto simp: inv_rel_def b_rel_def fun_rel_def)
+
+sepref_definition F_impl is
+  "RETURN o (\<lambda> (l, M). F l)" :: "state_assn'k \<rightarrow>a bool_assn"
+  by sepref
+
+end
+
+context TA_Impl_Op
+begin
+
+  definition succs where
+    "succs \<equiv> \<lambda> (l, D). rev (map (\<lambda> (g,a,r,l'). (l', f l r g l' D)) (trans_fun l))"
+
+  definition succs_P where
+    "succs_P P \<equiv> \<lambda> (l, D).
+      rev (map (\<lambda> (g,a,r,l'). (l', f l r g l' D)) (filter (\<lambda> (g,a,r,l'). P l') (trans_fun l)))"
+
+  definition succs' where
+    "succs' \<equiv> \<lambda> (l, D). do
+      {
+        xs \<leftarrow> nfoldli (trans_fun l) (\<lambda> _. True) (\<lambda> (g,a,r,l') xs.
+          RETURN ((l', f l r g l' (op_mtx_copy D)) # xs)
+        ) [];
+        RETURN xs
+      }"
+
+  definition succs_P' where
+    "succs_P' P \<equiv> \<lambda> (l, D). do
+      {
+        xs \<leftarrow> nfoldli (trans_fun l) (\<lambda> _. True) (\<lambda> (g,a,r,l') xs.
+          RETURN (if P l' then (l', f l r g l' (op_mtx_copy D)) # xs else xs)
+        ) [];
+        RETURN xs
+      }"
+
+  (* Unused *)
+  lemma RETURN_split:
+    "RETURN (f a b) = do
+      {
+        a \<leftarrow> RETURN a;
+        b \<leftarrow> RETURN b;
+        RETURN (f a b)
+      }"
+   by simp
+
+  lemma succs'_succs:
+    "succs' lD = RETURN (succs lD)"
+  unfolding succs'_def succs_def rev_map_fold
+    apply (cases lD)
+    apply simp
+    apply (rewrite fold_eq_nfoldli)
+    apply (fo_rule arg_cong fun_cong)+
+    apply auto
+  done
+
+  lemma succs_P'_succs_P:
+    "succs_P' P lD = RETURN (succs_P P lD)"
+    unfolding succs_P'_def succs_P_def rev_map_fold fold_filter
+    apply (cases lD)
+    apply simp
+    apply (rewrite fold_eq_nfoldli)
+    apply (fo_rule arg_cong fun_cong)+
+    by auto
+
+  lemmas [sepref_fr_rules] = dbm_subset'_impl.refine
+
+  sepref_register "PR_CONST (dbm_subset' n)" :: "'e DBMEntry i_mtx \<Rightarrow> 'e DBMEntry i_mtx \<Rightarrow> bool"
+
+  lemma [def_pat_rules]:
+    "dbm_subset' $ n \<equiv> PR_CONST (dbm_subset' n)"
+    by simp
+
+  lemmas [sepref_fr_rules] = check_diag_impl.refine
+
+  sepref_register "PR_CONST (check_diag n)" :: "'e DBMEntry i_mtx \<Rightarrow> bool"
+
+  lemma [def_pat_rules]:
+    "check_diag $ n \<equiv> PR_CONST (check_diag n)"
+  by simp
+
+  definition
+  "subsumes' =
+    (\<lambda> (l, M :: ('tt :: {zero,linorder}) DBM') (l', M').
+      l = l' \<and> pointwise_cmp (\<le>) n (curry M) (curry M'))"
+
+  context
+  begin
+
+  notation fun_rel_syn (infixr "\<rightarrow>" 60)
+
+  lemma left_unique_location_rel:
+    "IS_LEFT_UNIQUE location_rel"
+    unfolding IS_LEFT_UNIQUE_def
+    by (rule single_valuedI) (auto intro: loc_rel_left_unique single_valuedI)
+
+  lemma right_unique_location_rel:
+    "IS_RIGHT_UNIQUE location_rel"
+    by (rule single_valuedI) (auto intro: loc_rel_right_unique)
+
+  lemma [sepref_import_param]:
+    "((=), (=)) \<in> location_rel \<rightarrow> location_rel \<rightarrow> Id"
+    using left_unique_location_rel right_unique_location_rel
+    by (blast dest: IS_LEFT_UNIQUED IS_RIGHT_UNIQUED)
+
+  sepref_definition subsumes_impl is
+    "uncurry (RETURN oo subsumes')" :: "state_assn'k *a state_assn'k \<rightarrow>a bool_assn"
+    unfolding subsumes'_def short_circuit_conv dbm_subset'_def[symmetric]
+    by sepref
+
+  end
+
+  (* Somewhat peculiar use of the zero instance for DBM entries *)
+  lemma init_dbm_alt_def:
+    "init_dbm = op_amtx_dfltNxM (Suc n) (Suc n) (Le 0)"
+  unfolding init_dbm_def op_amtx_dfltNxM_def neutral by auto
+
+  lemma [sepref_import_param]: "(Le 0, PR_CONST (Le 0)) \<in> Id" by simp
+
+  lemma [def_pat_rules]: "Le $ 0 \<equiv> PR_CONST (Le 0)" by simp
+
+  sepref_register "PR_CONST (Le 0)"
+
+  sepref_definition init_dbm_impl is
+    "uncurry0 (RETURN (init_dbm :: nat \<times> nat \<Rightarrow> _))" :: "unit_assnk \<rightarrow>a (mtx_assn n)"
+  unfolding init_dbm_alt_def by sepref
+
+  lemmas [sepref_fr_rules] = init_dbm_impl.refine
+
+  sepref_register l0
+
+  lemma [sepref_import_param]:
+    "(l0i, l0) \<in> location_rel"
+    using init_impl states'_states by auto
+
+  sepref_definition a0_impl is
+    "uncurry0 (RETURN a0)" :: "unit_assnk \<rightarrow>a state_assn'"
+    unfolding a0_def by sepref
+
+  lemma trans_impl_trans_of[intro]:
+    "(g, a, r, l') \<in> set (trans_fun l) \<Longrightarrow> l \<in> states' \<Longrightarrow> A \<turnstile> l \<longrightarrow>g,a,r l'"
+    using trans_fun unfolding transition_rel_def by auto
+
+  lemma trans_of_trans_impl[intro]:
+    "A \<turnstile> l \<longrightarrow>g,a,r l' \<Longrightarrow> l \<in> states' \<Longrightarrow> (g, a, r, l') \<in> set (trans_fun l)"
+    using trans_fun unfolding transition_rel_def by auto
+
+  lemma trans_impl_clock_bounds1:
+    "(g, a, r, l') \<in> set (trans_fun l) \<Longrightarrow> l \<in> states' \<Longrightarrow> \<forall> c \<in> set r. c \<le> n"
+    using clocks_n reset_clk_set[OF trans_impl_trans_of] by fastforce
+
+  lemma trans_impl_clock_bounds2:
+    "(g, a, r, l') \<in> set (trans_fun l) \<Longrightarrow> l \<in> states' \<Longrightarrow> \<forall> c \<in> collect_clks g. c \<le> n"
+    using clocks_n collect_clocks_clk_set[OF trans_impl_trans_of] by fastforce
+
+  lemma trans_impl_states:
+    "(g, a, r, l') \<in> set (trans_fun l) \<Longrightarrow> l \<in> states' \<Longrightarrow> l' \<in> state_set (trans_of A)"
+     apply (drule trans_impl_trans_of)
+     apply assumption
+     unfolding state_set_def
+     apply (rule UnI2)
+     by force
+
+  lemma trans_impl_states':
+    "(g, a, r, l') \<in> set (trans_fun l) \<Longrightarrow> l \<in> states' \<Longrightarrow> l' \<in> states'"
+    using trans_impl_states states'_states by auto
+
+  lemma trans_of_states[intro]:
+    "l' \<in> states" if "A \<turnstile> l \<longrightarrow>g,a,r l'" "l \<in> states'"
+    using that by (auto intro: trans_impl_states)
+
+  lemma trans_of_states'[intro]:
+    "l' \<in> states'" if "A \<turnstile> l \<longrightarrow>g,a,r l'" "l \<in> states'"
+    using that states'_states by (auto intro: trans_impl_states)
+
+  abbreviation "clock_rel \<equiv> nbn_rel (Suc n)"
+
+  lemma [sepref_import_param]:
+    "(trans_impl, trans_fun)
+      \<in> location_rel \<rightarrow>
+        \<langle>\<langle>\<langle>clock_rel, int_rel\<rangle> acconstraint_rel\<rangle> list_rel \<times>r Id \<times>r \<langle>clock_rel\<rangle> list_rel \<times>r location_rel\<rangle>
+        list_rel"
+  proof (rule fun_relI)
+    fix li l
+    assume "(li, l) \<in> location_rel"
+    { fix xs  :: "((nat, int) acconstraint list \<times> 'a \<times> nat list \<times> 's) list"
+        and xs' :: "((nat, int) acconstraint list \<times> 'a \<times> nat list \<times> 'si) list"
+      assume A:
+        "\<forall> g a r l'. (g, a, r, l') \<in> set xs
+          \<longrightarrow> (\<forall> c \<in> collect_clks g. c \<le> n) \<and> (\<forall> c \<in> set r. c \<le> n) \<and> l' \<in> states'"
+      assume "(xs', xs) \<in> \<langle>Id \<times>r Id \<times>r Id \<times>r loc_rel\<rangle>list_rel"
+      then have
+        "(xs', xs) \<in>
+          \<langle>\<langle>\<langle>clock_rel, int_rel\<rangle>acconstraint_rel\<rangle>list_rel
+              \<times>r Id
+              \<times>r \<langle>clock_rel\<rangle>list_rel
+              \<times>r location_rel
+          \<rangle>list_rel"
+        using A
+      proof (induction xs' xs)
+        case 1 then show ?case
+          by simp
+      next
+        case (2 x' x xs' xs)
+        obtain g a r l' where [simp]: "x = (g, a, r, l')"
+          by (cases x)
+        obtain gi ai ri l'i where [simp]: "x' = (gi, ai, ri, l'i)"
+          by (cases x')
+        from 2 have r_bounds: "\<forall> c \<in> set r. c \<le> n"
+          by auto
+        from 2 have "\<forall> c \<in> collect_clks g. c \<le> n"
+          by auto
+        then have "(g, g) \<in> \<langle>\<langle>nbn_rel (Suc n), int_rel\<rangle>acconstraint_rel\<rangle>list_rel"
+          apply (induction g; simp)
+          subgoal for a
+            by (cases a)
+               (auto simp: acconstraint_rel_def p2rel_def rel2p_def split: acconstraint.split)
+          done
+        moreover from r_bounds have "(r, r) \<in> \<langle>nbn_rel (Suc n)\<rangle>list_rel"
+          by (induction r) auto
+        moreover from 2(1,4) have "(l'i, l') \<in> location_rel"
+          by auto
+        ultimately have
+          "(x', x) \<in>
+            \<langle>\<langle>clock_rel, int_rel\<rangle>acconstraint_rel\<rangle>list_rel \<times>r Id \<times>r \<langle>clock_rel\<rangle>list_rel \<times>r location_rel"
+          using 2(1) by simp
+        moreover from 2 have
+          "(xs', xs)
+            \<in> \<langle>\<langle>\<langle>clock_rel, int_rel\<rangle>acconstraint_rel\<rangle>list_rel \<times>r Id
+                \<times>r \<langle>clock_rel\<rangle>list_rel \<times>r location_rel\<rangle>list_rel"
+          by force
+        ultimately show ?case by simp
+      qed
+    } note * = this
+    from
+      \<open>(li, l) \<in> _\<close> trans_impl_clock_bounds1 trans_impl_clock_bounds2 trans_impl_states' trans_impl
+    show "(trans_impl li, trans_fun l)
+        \<in> \<langle>\<langle>\<langle>clock_rel, int_rel\<rangle>acconstraint_rel\<rangle>list_rel \<times>r
+           Id \<times>r \<langle>clock_rel\<rangle>list_rel \<times>r location_rel\<rangle>list_rel"
+      by - (rule *, auto, (drule fun_relD, assumption, simp add: relAPP_def)+)
+  qed
+
+  lemma b_rel_subtype[sepref_frame_match_rules]:
+    "hn_val (b_rel R P) a b \<Longrightarrow>t hn_val R a b"
+  by (rule enttI) (sep_auto simp: hn_ctxt_def pure_def)
+
+  lemma b_rel_subtype_merge[sepref_frame_merge_rules]:
+    "hn_val (b_rel R p) a b \<or>A hn_val R a b \<Longrightarrow>t hn_val R a b"
+    "hn_val R a b \<or>A hn_val (b_rel R p) a b \<Longrightarrow>t hn_val R a b"
+  by (simp add: b_rel_subtype entt_disjE)+
+
+  lemmas [safe_constraint_rules] = CN_FALSEI[of is_pure "asmtx_assn n a" for a]
+
+  sepref_register f ::
+    "'s \<Rightarrow> nat list \<Rightarrow> (nat, int) acconstraint list \<Rightarrow> 's \<Rightarrow> int DBMEntry i_mtx \<Rightarrow> int DBMEntry i_mtx"
+
+  lemmas [sepref_fr_rules] = op_impl
+
+  context
+    notes [id_rules] = itypeI[of "PR_CONST n" "TYPE (nat)"]
+      and [sepref_import_param] = IdI[of n]
+  begin
+
+  sepref_register trans_fun
+
+  sepref_definition succs_impl is
+    "RETURN o PR_CONST succs" :: "state_assn'k \<rightarrow>a list_assn state_assn'"
+  unfolding PR_CONST_def
+  unfolding
+    comp_def succs'_succs[symmetric] succs'_def
+    FW''_def[symmetric] rev_map_fold inv_of_A_def[symmetric]
+  apply (rewrite "HOL_list.fold_custom_empty")
+  by sepref
+
+  sepref_register succs :: "'s \<times> (int DBMEntry i_mtx) \<Rightarrow> ('s \<times> (int DBMEntry i_mtx)) list"
+
+  lemmas [sepref_fr_rules] = succs_impl.refine
+
+  sepref_definition succs_impl' is
+    "RETURN o (filter (\<lambda> (l, M). \<not>check_diag n M) o succs)"  :: "state_assn'k \<rightarrow>a list_assn state_assn'"
+    apply (rewrite in succs PR_CONST_def[symmetric])
+    unfolding list_filter_foldli (* This might be slightly inefficient *)
+    apply (rewrite "HOL_list.fold_custom_empty")
+    by sepref
+
+  end (* End sepref setup *)
+
+  context
+    fixes P :: "'s \<Rightarrow> bool" and P_fun
+    assumes P_fun: "(P_fun, P) \<in> inv_rel loc_rel states'"
+  begin
+
+  context
+    notes [id_rules] = itypeI[of "PR_CONST n" "TYPE (nat)"]
+      and [sepref_import_param] = IdI[of n]
+  begin
+
+  sepref_register "PR_CONST P"
+
+  lemma [sepref_fr_rules]:
+    "(return o P_fun, RETURN o PR_CONST P) \<in> location_assnk \<rightarrow>a id_assn"
+    using P_fun by sepref_to_hoare (sep_auto simp: inv_rel_def b_rel_def fun_rel_def)
+
+  sepref_definition succs_P_impl is
+    "RETURN o PR_CONST (succs_P P)" :: "state_assn'k \<rightarrow>a list_assn state_assn'"
+    unfolding PR_CONST_def
+    apply (rewrite in P PR_CONST_def[symmetric])
+    unfolding
+      comp_def succs_P'_succs_P[symmetric] succs_P'_def
+      FW''_def[symmetric] rev_map_fold inv_of_A_def[symmetric]
+    apply (rewrite "HOL_list.fold_custom_empty")
+    by sepref
+
+  sepref_register "succs_P P" :: "'s \<times> (int DBMEntry i_mtx) \<Rightarrow> ('s \<times> (int DBMEntry i_mtx)) list"
+
+  lemmas [sepref_fr_rules] = succs_P_impl.refine
+
+  sepref_definition succs_P_impl' is
+    "RETURN o (filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P P)"  :: "state_assn'k \<rightarrow>a list_assn state_assn'"
+    apply (rewrite in "succs_P P" PR_CONST_def[symmetric])
+    unfolding list_filter_foldli (* This might be slightly inefficient *)
+    apply (rewrite "HOL_list.fold_custom_empty")
+    by sepref
+
+  end (* End sepref setup *)
+
+  end (* Fixed predicate *)
+
+  \<comment> \<open>Could be moved to different context.\<close>
+  lemma reachable_states:
+    "l \<in> states" if "op.E_from_op** a0 (l, M)"
+    using that
+    by (induction "(l, M)" arbitrary: l M; force simp: a0_def state_set_def op.E_from_op_def)
+
+  definition "emptiness_check \<equiv> \<lambda> (l, M). check_diag n M"
+
+  sepref_definition emptiness_check_impl  is
+    "RETURN o emptiness_check" :: "state_assn'k \<rightarrow>a bool_assn"
+    unfolding emptiness_check_def
+    by sepref
+
+  lemma state_copy:
+    "RETURN \<circ> COPY = (\<lambda> (l, M). do {M' \<leftarrow> mop_mtx_copy M; RETURN (l, M')})"
+    by auto
+
+  sepref_definition state_copy_impl is
+    "RETURN \<circ> COPY" :: "state_assn'k \<rightarrow>a state_assn'"
+    unfolding state_copy
+    by sepref
+
+  lemma location_assn_constraints:
+    "is_pure location_assn"
+    "IS_LEFT_UNIQUE (the_pure location_assn)"
+    "IS_RIGHT_UNIQUE (the_pure location_assn)"
+    using left_unique_location_rel right_unique_location_rel
+    by (auto elim: single_valued_subset[rotated] simp: b_rel_def IS_LEFT_UNIQUE_def)
+
+  lemma not_check_diag_init_dbm[intro, simp]:
+      "\<not> check_diag n init_dbm"
+      unfolding check_diag_def init_dbm_def by auto
+
+end (* TA Impl Op *)
+
+  subsection \<open>Instantiation of Worklist Algorithms\<close>
+
+context Reachability_Problem_Impl_Op
+begin
+
+  sublocale Worklist0 op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M"
+    apply standard
+    apply (clarsimp simp: op.reachable_def split: prod.split dest!: reachable_states)
+    unfolding succs_def op.E_from_op_def using states'_states by force
+
+  sublocale Worklist1 op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" ..
+
+  sublocale Worklist2 op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes'
+    apply standard
+    unfolding subsumes_def subsumes'_def by auto
+
+  sublocale
+    Worklist3
+      op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes' "\<lambda> (l, M). F l"
+    apply standard
+    unfolding F_rel_def by auto
+
+  sublocale
+    Worklist4
+      op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes' "\<lambda> (l, M). F l"
+    apply standard
+    unfolding F_rel_def by auto
+
+  sublocale
+    Worklist_Map a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M" subsumes' op.E_from_op fst succs
+    apply standard
+    unfolding subsumes'_def by auto
+
+  sublocale
+    Worklist_Map2
+      a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M" subsumes' op.E_from_op
+      fst succs "\<lambda> (l, M). F l"
+    ..
+
+  sublocale Worklist4_Impl
+    op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes' "\<lambda> (l, M). F l"
+    state_assn' succs_impl a0_impl F_impl subsumes_impl emptiness_check_impl
+    apply standard
+    apply (unfold PR_CONST_def)
+    by (rule
+        a0_impl.refine F_impl.refine subsumes_impl.refine succs_impl.refine[unfolded PR_CONST_def]
+        emptiness_check_impl.refine[unfolded emptiness_check_def]
+        )+
+
+  sublocale Worklist_Map2_Impl
+    op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes'
+    "\<lambda> (l, M). F l" state_assn'
+    succs_impl a0_impl F_impl subsumes_impl emptiness_check_impl fst "return o fst" state_copy_impl
+    tracei location_assn
+    apply standard
+    subgoal
+    unfolding PR_CONST_def
+     apply sepref_to_hoare
+    apply sep_auto
+    done
+  subgoal
+    by (rule state_copy_impl.refine)
+  subgoal
+    unfolding trace_def by (rule tracei_refine)
+  by (rule location_assn_constraints)+
+
+  sublocale Worklist_Map2_Hashable
+    op.E_from_op a0 F_rel "subsumes n" succs "\<lambda> (l, M). check_diag n M" subsumes' "\<lambda> (l, M). F l"
+    state_assn' succs_impl a0_impl F_impl subsumes_impl emptiness_check_impl
+    fst "return o fst" state_copy_impl tracei location_assn by standard
+
+  sublocale liveness: Liveness_Search_Space_Key
+    "\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M'" a0
+    "\<lambda> _. False"
+    "subsumes n" "\<lambda> (l, M). op.reachable (l, M) \<and> \<not> check_diag n M \<and> F l"
+    "filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P F" fst subsumes'
+    apply standard
+           apply blast
+          apply (blast intro: op.liveness_search_space_pre.trans)
+    subgoal for a b a'
+      apply (drule op.liveness_search_space_pre.mono[where a'=a'])
+      unfolding op.liveness_search_space_pre.G.E'_def by (auto simp: F_rel_def)
+        apply blast
+    subgoal
+      using op.liveness_search_space_pre.finite_V
+      by (auto elim: finite_subset[rotated] simp: F_rel_def)
+    subgoal premises prems for a
+    proof -
+      have
+        "set ((filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P F) a)
+        = {x. (\<lambda>(l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l' \<and> \<not> check_diag n M') a x}"
+        unfolding op.E_from_op_def succs_P_def using prems states'_states
+        by (safe; force dest!: reachable_states simp: op.reachable_def)
+      also have "\<dots> =
+        {x. Subgraph_Node_Defs.E'
+         (\<lambda>(l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M')
+         (\<lambda>(l, M). op.reachable (l, M) \<and> \<not> check_diag n M \<and> F l) a x}"
+        unfolding Subgraph_Node_Defs.E'_def using prems by auto
+      finally show ?thesis .
+    qed
+    by (blast intro!: empty_subsumes')+
+
+  sublocale liveness: Liveness_Search_Space_Key_Impl
+    a0 "\<lambda> _. False" "subsumes n" "\<lambda> (l, M). op.reachable (l, M) \<and> \<not> check_diag n M \<and> F l"
+    "filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P F"
+    "\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M'"
+    subsumes' fst
+    state_assn' "succs_P_impl' F_fun" a0_impl subsumes_impl "return o fst" state_copy_impl
+    location_assn
+    apply standard
+    apply (rule location_assn_constraints)+
+        apply (unfold PR_CONST_def, rule a0_impl.refine; fail)
+       apply (unfold PR_CONST_def, rule subsumes_impl.refine; fail)
+      apply (unfold PR_CONST_def, rule succs_P_impl'.refine[OF F_fun])
+    subgoal
+      by sepref_to_hoare sep_auto
+    by (rule state_copy_impl.refine)
+
+  lemma precond_a0:
+    "case a0 of (l, M) \<Rightarrow> op.reachable (l, M) \<and> \<not> check_diag n M"
+    unfolding op.reachable_def unfolding a0_def by auto
+
+  lemma liveness_step_equiv:
+    fixes x y
+    assumes "(\<lambda> (l, M). op.reachable (l, M) \<and> \<not> check_diag n M \<and> F l) x"
+    shows "liveness.G.E' x y \<longleftrightarrow>
+      (\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M') x y"
+    using assms unfolding liveness.G.E'_def by auto
+
+  lemma liveness_check_equiv:
+    "(\<exists>x. liveness.G.G'.reaches a0 x \<and> liveness.G.G'.reaches1 x x) \<longleftrightarrow>
+       (\<exists> x. op.liveness_pre.reaches a0 x \<and> op.liveness_pre.reaches1 x x)"
+    if "F l0"
+    apply (subst rtranclp_tranclp_equiv[OF liveness_step_equiv])
+       apply assumption
+    subgoal
+      unfolding liveness.G.E'_def by auto
+    subgoal
+      using that precond_a0 by (auto simp: a0_def)
+    ..
+
+  lemma liveness_spec_refine:
+    "SPEC (\<lambda>r. r =
+       (\<exists>x. liveness.G.G'.reaches a0 x \<and> liveness.G.G'.reaches1 x x)) \<le>
+     (SPEC (\<lambda> r. r =
+     (\<exists> x.
+       (\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M')** a0 x \<and>
+       (\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M')++ x x)
+      )
+     )
+    " if "F l0"
+    using liveness_check_equiv[OF that] by auto
+
+  lemma liveness_hnr:
+    "(uncurry0
+      (dfs_map_impl' (succs_P_impl' F_fun) a0_impl subsumes_impl
+        (return \<circ> fst) state_copy_impl),
+     uncurry0 (SPEC (\<lambda>r. r = (\<exists>x. op.liveness_pre.reaches a0 x \<and> op.liveness_pre.reaches1 x x))))
+      \<in> unit_assnk \<rightarrow>a bool_assn"
+    if "F l0"
+    apply (rule liveness.dfs_map_impl'_hnr[
+        FCOMP liveness_spec_refine[THEN Id_SPEC_refine, THEN nres_relI]
+        ])
+    using that precond_a0 by (auto simp: a0_def)
+
+  context
+    fixes Q :: "'s \<Rightarrow> bool" and Q_fun
+    assumes Q_fun: "(Q_fun, Q) \<in> inv_rel loc_rel states'"
+  begin
+
+  interpretation leadsto: Leadsto_Search_Space_Key
+    a0 "\<lambda> _. False" "subsumes n" "\<lambda> (l, M). check_diag n M" subsumes'
+    "\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> \<not> check_diag n M'"
+    fst "\<lambda> _. False" "\<lambda> (l, M). F l" "\<lambda> (l, M). Q l"
+    "filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P Q" "filter (\<lambda> (l, M). \<not>check_diag n M) o succs"
+  proof (goal_cases)
+    case 1
+    interpret Subgraph_Start
+      op.E_from_op a0 "\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> \<not> check_diag n M'"
+      by standard auto
+    show ?case
+      apply standard
+                    apply blast
+                   apply (blast intro: liveness.trans)
+      subgoal for a b a'
+        by (drule op.mono[where a' = a'], auto dest: op.empty_mono[rotated] intro: reachable)
+                 apply (blast intro: op.empty_subsumes)
+                apply (rule op.empty_mono; assumption)
+      subgoal for x x'
+        by (auto dest: op.E_from_op_check_diag)
+              apply assumption
+             apply blast
+            apply blast
+      subgoal for a
+        by (auto dest: succs_correct reachable)
+          apply (simp; fail)
+      subgoal
+        using op.finite_reachable by (auto intro: reachable elim!: finite_subset[rotated])
+      subgoal for a a'
+        by (auto simp: empty_subsumes' dest: subsumes_key)
+      subgoal for a a'
+        by (auto simp: empty_subsumes' dest: subsumes_key)
+      subgoal premises prems for a
+      proof -
+        have
+          "set ((filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P Q) a)
+            = {x. (\<lambda>(l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> Q l' \<and> \<not> check_diag n M') a x}"
+          unfolding op.E_from_op_def succs_P_def using prems states'_states
+          by (safe; force dest!: reachable reachable_states simp: op.reachable_def)
+        then show ?thesis
+          by auto
+      qed
+      done
+  qed
+
+  sepref_register "PR_CONST Q"
+
+  lemma [sepref_fr_rules]:
+    "(return o Q_fun, RETURN o PR_CONST Q) \<in> location_assnk \<rightarrow>a id_assn"
+    using Q_fun by sepref_to_hoare (sep_auto simp: inv_rel_def b_rel_def fun_rel_def)
+
+  sepref_definition Q_impl is
+    "RETURN o (\<lambda> (l, M). Q l)" :: "state_assn'k \<rightarrow>a bool_assn"
+    by (rewrite in Q PR_CONST_def[symmetric]) sepref
+
+  interpretation leadsto: Leadsto_Search_Space_Key_Impl
+    "subsumes n" subsumes' location_assn fst a0 "\<lambda> _. False" "\<lambda> _. False" state_copy_impl
+    "\<lambda> (l, M). F l" "\<lambda> (l, M). Q l"
+    "\<lambda> (l, M). op.reachable (l, M) \<and> \<not> check_diag n M"
+    "\<lambda> (l, M). check_diag n M"
+    "filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P Q" "filter (\<lambda> (l, M). \<not>check_diag n M) o succs"
+    "\<lambda> (l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> \<not> check_diag n M'"
+    state_assn'
+    "succs_P_impl' Q_fun" a0_impl subsumes_impl "return o fst" succs_impl'
+    emptiness_check_impl F_impl Q_impl tracei
+    apply standard
+                       apply blast
+                      apply (blast intro: op.trans)
+    subgoal for a b a'
+      apply (drule op.mono[where a' = a'], auto dest: op.empty_mono[rotated])
+      apply (intro exI conjI)
+           apply (auto dest: op.empty_mono[rotated])
+      by (auto simp: empty_subsumes' dest: subsumes_key)
+                    apply blast
+    subgoal
+      using op.finite_reachable by (auto elim!: finite_subset[rotated])
+    subgoal premises prems for a
+    proof -
+      have
+        "set ((filter (\<lambda> (l, M). \<not>check_diag n M) o succs_P Q) a)
+            = {x. (\<lambda>(l, M) (l', M'). op.E_from_op (l, M) (l', M') \<and> Q l' \<and> \<not> check_diag n M') a x}"
+        unfolding op.E_from_op_def succs_P_def using prems states'_states
+        by (safe; force dest!: reachable_states[folded op.reachable_def])
+      also have "\<dots> =
+            {x. Subgraph_Node_Defs.E'
+             (\<lambda>x y.
+              (case x of (l, M) \<Rightarrow> \<lambda>(l', M'). op.E_from_op (l, M) (l', M') \<and> \<not> check_diag n M') y
+                \<and> \<not> (case y of (l, M) \<Rightarrow> check_diag n M) \<and> (case y of (l, M) \<Rightarrow> Q l)
+              )
+             (\<lambda>(l, M). op.reachable (l, M) \<and> \<not> check_diag n M) a x}"
+        unfolding Subgraph_Node_Defs.E'_def using prems by auto
+      finally show ?thesis .
+    qed
+                 apply blast
+                apply (clarsimp simp: empty_subsumes'; fail)
+               apply (rule location_assn_constraints)+
+            apply (rule liveness.refinements)
+           apply (rule liveness.refinements)
+          apply (unfold PR_CONST_def, rule succs_P_impl'.refine[OF Q_fun])
+    subgoal
+      by sepref_to_hoare sep_auto
+    by (rule
+        succs_impl'.refine liveness.refinements
+        emptiness_check_impl.refine[unfolded emptiness_check_def]
+        F_impl.refine Q_impl.refine tracei_refine
+        )+
+
+  definition
+    "leadsto_spec_alt = SPEC (\<lambda>r. r \<longleftrightarrow>
+       (\<exists>a. leadsto.A.search_space.reaches a0 a \<and>
+            \<not> (case a of (l, M) \<Rightarrow> check_diag n M) \<and>
+            (case a of (l, M) \<Rightarrow> F l) \<and> (case a of (l, M) \<Rightarrow> Q l) \<and>
+            (\<exists>b. leadsto.B.reaches a b \<and> leadsto.B.reaches1 b b))
+     )
+    "
+
+  lemmas leadsto_impl_hnr =
+    leadsto.leadsto_impl_hnr[
+      unfolded leadsto.leadsto_spec_alt_def, unfolded leadsto.reaches_cycle_def,
+      folded leadsto_spec_alt_def
+    ]
+
+  end (* Context for second predicate *)
+
+end (* Reachability Problem Impl Op *)
+
+
+subsection \<open>Implementation of the Invariant Precondition Check\<close>
+
+context TA_Impl_Op
+begin
+
+  sepref_register "init_dbm :: nat \<times> nat \<Rightarrow> _" :: "'e DBMEntry i_mtx"
+
+  sepref_register "unbounded_dbm n :: nat \<times> nat \<Rightarrow> int DBMEntry" :: "'b DBMEntry i_mtx"
+
+  lemmas [sepref_fr_rules] = dbm_subset_impl.refine
+
+  sepref_register "PR_CONST (dbm_subset n)" :: "'e DBMEntry i_mtx \<Rightarrow> 'e DBMEntry i_mtx \<Rightarrow> bool"
+
+  lemma [def_pat_rules]:
+    "dbm_subset $ n \<equiv> PR_CONST (dbm_subset n)"
+    by simp
+
+  context
+    notes [id_rules] = itypeI[of "PR_CONST n" "TYPE (nat)"]
+      and [sepref_import_param] = IdI[of n]
+  begin
+
+  sepref_definition start_inv_check_impl is
+    "uncurry0 (RETURN (start_inv_check :: bool))" :: "unit_assnk \<rightarrow>a bool_assn"
+    unfolding start_inv_check_def
+      FW''_def[symmetric] rev_map_fold reset'_upd_def inv_of_A_def[symmetric]
+    supply [sepref_fr_rules] = unbounded_dbm_impl.refine
+    by sepref
+
+  end (* End sepref setup *)
+
+end (* End of locale *)
+
+
+subsection \<open>Instantiation for the Concrete DBM Successor Operations\<close>
+
+lemma (in Graph_Defs) Alw_ev:
+  "Alw_ev \<phi> x" if "\<phi> x"
+  using that unfolding Alw_ev_def by (auto simp: holds.simps)
+
+context TA_Impl
+begin
+
+interpretation DBM_Impl n .
+
+context
+  notes [id_rules] = itypeI[of "PR_CONST n" "TYPE (nat)"]
+    and [sepref_import_param] = IdI[of n]
+begin
+
+  sepref_definition E_op'_impl is
+    "uncurry4 (\<lambda> l r. RETURN ooo E_op' l r)" :: "op_impl_assn"
+    unfolding E_op'_def FW''_def[symmetric] reset'_upd_def inv_of_A_def[symmetric] PR_CONST_def
+    by sepref
+
+  sepref_definition E_op''_impl is
+    "uncurry4 (\<lambda> l r. RETURN ooo E_op'' l r)" :: "op_impl_assn"
+    unfolding E_op''_def FW''_def[symmetric] reset'_upd_def inv_of_A_def[symmetric] filter_diag_def
+    by sepref
+
+end (* End sepref setup *)
+
+end (* TA Impl *)
+
+subsubsection \<open>Correctness Theorems\<close>
+
+context Reachability_Problem_Impl
+begin
+
+sublocale Reachability_Problem_Impl_Op
+  where loc_rel = loc_rel and f = "PR_CONST E_op''" and op_impl = E_op''_impl
+  unfolding PR_CONST_def by standard (rule E_op''_impl.refine)
+
+lemma E_op_F_reachable:
+  "op.F_reachable = E_op''.F_reachable" unfolding PR_CONST_def ..
+
+lemma (in -) state_set_eq[simp]:
+  "Simulation_Graphs_TA.state_set A = state_set (trans_of A)"
+  unfolding Simulation_Graphs_TA.state_set_def state_set_def trans_of_def ..
+
+lemma op_liveness_reaches_cycle_equiv:
+  "(\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> F (fst b))** a0 a \<and>
+   (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> F (fst b))++ a b
+  \<longleftrightarrow> op.liveness_pre.reaches a0 a \<and> op.liveness_pre.reaches1 a b" if "F l0"
+  using that by - (rule rtranclp_tranclp_equiv[of "F o fst"], auto simp: a0_def)
+
+lemma Alw_ev_impl_hnr:
+  "(uncurry0
+    (if F l0 then
+      dfs_map_impl'
+        (succs_P_impl' F_fun) a0_impl subsumes_impl (return \<circ> fst) state_copy_impl
+     else return False),
+   uncurry0 (SPEC (\<lambda>r. l0 \<in> state_set (trans_of A) \<longrightarrow>
+    r \<longleftrightarrow> \<not> (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> Alw_ev (\<lambda>(l, u). \<not> F l) (l0, u0)))))
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+  unfolding state_set_eq[symmetric]
+  apply (cases "F l0")
+  subgoal premises prems
+  proof -
+    define protected1 where "protected1 = E_op''.liveness_pre.reaches"
+    define protected2 where "protected2 = E_op''.liveness_pre.reaches1"
+    show ?thesis
+      using prems Alw_ev_mc[folded a0_def, of F, unfolded op_liveness_reaches_cycle_equiv[OF prems]]
+      apply (simp add: )
+      unfolding protected1_def[symmetric] protected2_def[symmetric]
+      using liveness_hnr[OF prems, to_hnr, unfolded hn_refine_def]
+        apply sepref_to_hoare
+      apply sep_auto
+      apply (erule cons_post_rule)
+      unfolding protected1_def[symmetric] protected2_def[symmetric]
+      by sep_auto
+  qed
+  subgoal
+    using Alw_ev_mc[folded a0_def, of F]
+    apply simp
+    by sepref_to_hoare sep_auto
+  done
+
+context
+    fixes Q :: "'s \<Rightarrow> bool" and Q_fun
+    assumes Q_fun: "(Q_fun, Q) \<in> inv_rel loc_rel states'"
+    assumes no_deadlock: "\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0)"
+begin
+
+lemma leadsto_spec_refine:
+  "leadsto_spec_alt Q
+  \<le> SPEC (\<lambda> r. \<not> r \<longleftrightarrow>
+    (\<nexists>x. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** (l0, init_dbm) x \<and>
+       F (fst x) \<and>
+       Q (fst x) \<and>
+       (\<exists>a. (\<lambda>a b. E_op''.E_from_op a b \<and>
+                   \<not> check_diag n (snd b) \<and> Q (fst b))**
+             x a \<and>
+            (\<lambda>a b. E_op''.E_from_op a b \<and>
+                   \<not> check_diag n (snd b) \<and> Q (fst b))++
+             a a))
+    )"
+proof -
+  have *:"
+    (\<lambda>x y. (case y of (l', M') \<Rightarrow> E_op''.E_from_op x (l', M') \<and> \<not> check_diag n M') \<and>
+    \<not> (case y of (l, M) \<Rightarrow> check_diag n M))
+    = (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))"
+    by (intro ext) auto
+  have **:
+    "(\<lambda>x y. (case y of (l', M') \<Rightarrow> E_op''.E_from_op x (l', M') \<and> \<not> check_diag n M') \<and>
+     (case y of (l, M) \<Rightarrow> Q l) \<and> \<not> (case y of (l, M) \<Rightarrow> check_diag n M))
+     = (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))"
+    by (intro ext) auto
+  have ***: "\<not> check_diag n b"
+    if "(\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** a0 (a, b)" for a b
+    using that by cases (auto simp: a0_def)
+  show ?thesis
+    unfolding leadsto_spec_alt_def[OF Q_fun]
+    unfolding PR_CONST_def a0_def[symmetric] by (auto dest: *** simp: * **)
+  qed
+
+lemma leadsto_impl_hnr:
+  "(uncurry0
+    (leadsto_impl state_copy_impl
+      (succs_P_impl' Q_fun) a0_impl subsumes_impl (return \<circ> fst)
+      succs_impl' emptiness_check_impl F_impl (Q_impl Q_fun) tracei),
+   uncurry0
+    (SPEC
+      (\<lambda>r. l0 \<in> state_set (trans_of A) \<longrightarrow>
+           (\<not> r) =
+           (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow>
+                  leadsto (\<lambda>(l, u). F l) (\<lambda>(l, u). \<not> Q l) (l0, u0)))))
+  \<in> unit_assnk \<rightarrow>a bool_assn"  
+  unfolding state_set_eq[symmetric]
+ using leadsto_impl_hnr[
+    OF Q_fun precond_a0,
+    FCOMP leadsto_spec_refine[THEN Id_SPEC_refine, THEN nres_relI], to_hnr, unfolded hn_refine_def]
+  using leadsto_mc[OF _ no_deadlock, of F Q]
+  apply (simp del: state_set_eq)
+  apply sepref_to_hoare
+  apply sep_auto
+  apply (erule cons_post_rule)
+  apply sep_auto
+  done
+
+end (* Context for leadsto predicate *)
+
+end (* End of Reachability Problem Impl *)
+
+
+subsection \<open>Instantiation for a Concrete Automaton\<close>
+
+datatype result = REACHABLE | UNREACHABLE | INIT_INV_ERR
+
+context Reachability_Problem_precompiled
+begin
+
+  sublocale Defs: Reachability_Problem_Impl_Defs
+    _ _ A 0 m "\<lambda> l i. if l < n \<and> i \<le> m then k ! l ! i else 0" "PR_CONST F" .
+
+  lemma
+    "(IArray xs) !! i = xs ! i"
+  by auto
+
+  text \<open>Definition of implementation auxiliaries (later connected to the automaton via proof)\<close>
+  (*
+  definition
+    "trans_impl l \<equiv> map (\<lambda> i. label i ((IArray trans) !! l ! i)) [0..<length (trans ! l)]"
+  *)
+
+  definition
+    "trans_map \<equiv> map (\<lambda> xs. map (\<lambda> i. label i (xs ! i)) [0..<length xs]) trans"
+
+  definition
+    "trans_impl l \<equiv> (IArray trans_map) !! l"
+
+  lemma trans_impl_alt_def:
+    "l < n
+    \<Longrightarrow> trans_impl l = map (\<lambda> i. label i ((IArray trans) !! l ! i)) [0..<length (trans ! l)]"
+  unfolding trans_impl_def trans_map_def by (auto simp: trans_length)
+
+  lemma state_set_n[intro, simp]:
+    "state_set (trans_of A) \<subseteq> {0..<n}"
+  unfolding state_set_def trans_of_def A_def T_def label_def using state_set trans_length
+  by (force dest: nth_mem)
+
+  lemma trans_impl_trans_of[intro, simp]:
+    "(trans_impl, trans_of A) \<in> transition_rel Defs.states"
+    using state_set_n n_gt_0 unfolding transition_rel_def trans_of_def A_def T_def
+    by (fastforce simp: trans_impl_alt_def)
+
+  definition "inv_fun l \<equiv> (IArray inv) !! l"
+
+  lemma inv_fun_inv_of[intro, simp]:
+    "(inv_fun, inv_of A) \<in> inv_rel Id Defs.states"
+  using state_set_n n_gt_0 unfolding inv_rel_def inv_fun_def inv_of_def A_def I_def
+  by auto
+
+  definition "final_fun = List.member final"
+
+  lemma final_fun_final[intro, simp]:
+    "(final_fun, F) \<in> inv_rel Id Defs.states"
+  using state_set_n unfolding F_def final_fun_def inv_rel_def by (auto simp: in_set_member)
+
+  lemma start_states[intro, simp]:
+    "0 \<in> Defs.states"
+  proof -
+    obtain g r l' where "trans ! 0 ! 0 = (g, r, l')" by (metis prod_cases3)
+    with start_has_trans n_gt_0 trans_length show ?thesis
+    unfolding state_set_def trans_of_def A_def T_def label_def by force
+  qed
+
+  lemma iarray_k':
+    "(IArray.sub (IArray (map (IArray o map int) k)), IArray o k') \<in> inv_rel Id Defs.states"
+    unfolding inv_rel_def k'_def
+    apply (clarsimp simp del: upt_Suc)
+    subgoal premises prems for j
+    proof -
+      have "j < n" using prems n_gt_0 state_set_n by fastforce
+      with k_length have "length (k ! j) = m + 1" by simp
+      with k_length(2) have
+        "map (\<lambda>i. if i \<le> m then k ! j ! i else 0) [0..<Suc m] = map (\<lambda>i. k ! j ! i) [0..<Suc m]"
+        by simp
+      also have "\<dots> = k ! j" using \<open>length (k ! j) = _\<close> by (simp del: upt_Suc) (metis List.map_nth)
+      also show ?thesis using \<open>j < n\<close> k_length(1)
+          apply simp
+          apply (subst calculation[symmetric])
+        by simp
+    qed
+    done
+
+  lemma trans_impl_refine_self:
+    "(trans_impl, trans_impl)
+      \<in> fun_rel_syn nat_rel (list_rel (Id \<times>r nat_rel \<times>r Id \<times>r nat_rel))"
+    by auto (metis IdI list_rel_id_simp relAPP_def)
+
+  sublocale Reachability_Problem_Impl
+    where A = A
+    and inv_fun = inv_fun
+    and F = "PR_CONST F"
+    and F_fun = final_fun
+    and trans_fun = trans_impl
+    and trans_impl = trans_impl
+    and ceiling = "IArray.sub (IArray (map (IArray o map int) k))"
+    and k = "\<lambda> l i. if l < n \<and> i \<le> m then k ! l ! i else 0"
+    and n = m
+    and loc_rel = Id
+    and l0 = "0::nat"
+    and l0i = 0
+    and show_state = "show"
+    and show_clock = "show"
+    and states' = Defs.states
+    unfolding PR_CONST_def using iarray_k' trans_impl_refine_self by - (standard, fastforce+)
+
+  subsubsection \<open>Correctness Theorems\<close>
+
+  lemma F_reachable_correct:
+    "op.F_reachable
+    \<longleftrightarrow> (\<exists> l' u u'. conv_A A \<turnstile>' \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final)"
+    using E_op''.E_from_op_reachability_check[of F_rel "PR_CONST F"] reachability_check
+    unfolding E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+    unfolding F_rel_def unfolding F_def by auto
+
+  definition
+    "reachability_checker_wl \<equiv>
+      worklist_algo2_impl subsumes_impl a0_impl F_impl succs_impl emptiness_check_impl"
+
+  definition
+    "reachability_checker' \<equiv>
+      pw_impl
+        (return o fst) state_copy_impl tracei
+        subsumes_impl a0_impl F_impl succs_impl emptiness_check_impl"
+
+  theorem reachability_check':
+    "(uncurry0 reachability_checker',
+      uncurry0 (
+        RETURN
+          (\<exists> l' u u'. conv_A A \<turnstile>' \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final)
+      )
+     )
+    \<in> unit_assnk \<rightarrow>a bool_assn"
+    using pw_impl_hnr_F_reachable
+    unfolding reachability_checker'_def F_reachable_correct .
+
+  corollary reachability_checker'_hoare:
+    "<emp> reachability_checker'
+    <\<lambda> r. \<up>(r \<longleftrightarrow>
+      (\<exists> l' u u'. conv_A A \<turnstile>' \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final))
+    >t"
+   apply (rule cons_post_rule)
+   using reachability_check'[to_hnr] apply (simp add: hn_refine_def)
+   by (sep_auto simp: pure_def)
+
+  definition reachability_checker where
+    "reachability_checker \<equiv> do
+      {
+        init_sat \<leftarrow> start_inv_check_impl;
+        if init_sat then do
+          { x \<leftarrow> reachability_checker';
+            return (if x then REACHABLE else UNREACHABLE)
+          }
+        else
+          return INIT_INV_ERR
+      }"
+
+  theorem reachability_check:
+    "(uncurry0 reachability_checker,
+       uncurry0 (
+        RETURN (
+          if start_inv_check
+          then
+            if
+              \<exists> l' u u'. conv_A A \<turnstile> \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final
+            then REACHABLE
+            else UNREACHABLE
+          else INIT_INV_ERR
+        )
+       )
+      )
+    \<in> unit_assnk \<rightarrow>a id_assn"
+  proof -
+    define check_A where
+      "check_A \<equiv> \<exists>l' u u'. conv_A A \<turnstile>' \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall>c \<in> {1..m}. u c = 0) \<and> l' \<in> set final"
+    define check_B where
+      "check_B \<equiv> \<exists>l' u u'. conv_A A \<turnstile> \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall>c \<in> {1..m}. u c = 0) \<and> l' \<in> set final"
+    note F_reachable_equiv' =
+      F_reachable_equiv
+      [unfolded F_def PR_CONST_def check_A_def[symmetric] check_B_def[symmetric]]
+    show ?thesis
+      unfolding reachability_checker_def
+      unfolding check_A_def[symmetric] check_B_def[symmetric]
+      using F_reachable_equiv'
+      supply reachability_check'
+        [unfolded check_A_def[symmetric], to_hnr, unfolded hn_refine_def, rule_format, sep_heap_rules]
+      supply start_inv_check_impl.refine
+        [to_hnr, unfolded hn_refine_def, rule_format, sep_heap_rules]
+      apply sepref_to_hoare
+      by (sep_auto simp: pure_def)
+  qed
+
+  corollary reachability_checker_hoare':
+    "<emp> reachability_checker
+    <\<lambda> r.
+    \<up>(r = (
+      if (\<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> u \<turnstile> inv_of (conv_A A) 0)
+        then
+          if
+            \<exists> l' u u'. conv_A A \<turnstile> \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final
+          then REACHABLE
+          else UNREACHABLE
+        else INIT_INV_ERR
+      ))
+    >t"
+   unfolding start_inv_check_correct[symmetric]
+   apply (rule cons_post_rule)
+   using reachability_check[to_hnr] apply (simp add: hn_refine_def)
+   by (sep_auto simp: pure_def)
+
+  corollary reachability_checker_hoare:
+    "<emp> reachability_checker
+    <\<lambda> r.
+    \<up>(
+        if \<not> (\<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> u \<turnstile> inv_of (conv_A A) 0)
+          then r = INIT_INV_ERR
+        else if \<exists> l' u u'. conv_A A \<turnstile> \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final
+          then r = REACHABLE
+        else r = UNREACHABLE
+      )
+    >t"
+   unfolding start_inv_check_correct[symmetric]
+   apply (rule cons_post_rule)
+   using reachability_check[to_hnr] apply (simp add: hn_refine_def)
+   by (sep_auto simp: pure_def)
+
+  paragraph \<open>Post-processing\<close>
+
+  ML \<open>
+    fun pull_let u t =
+      let
+        val t1 = abstract_over (u, t);
+        val r1 = Const (@{const_name "HOL.Let"}, dummyT) $ u $ Abs ("I", dummyT, t1);
+        val ct1 = Syntax.check_term @{context} r1;
+        val g1 =
+          Goal.prove @{context} [] [] (Logic.mk_equals (t, ct1))
+          (fn {context, ...} => EqSubst.eqsubst_tac context [0] [@{thm Let_def}] 1
+          THEN resolve_tac context [@{thm Pure.reflexive}] 1)
+      in g1 end;
+
+    fun get_rhs thm =
+      let
+        val Const ("Pure.eq", _) $ _ $ r = Thm.full_prop_of thm
+      in r end;
+
+    fun get_lhs thm =
+      let
+        val Const ("Pure.imp", _) $ (Const ("Pure.eq", _) $ l $ _) $ _ = Thm.full_prop_of thm
+      in l end;
+
+    fun pull_tac' u ctxt thm =
+      let
+        val l = get_lhs thm;
+        val rewr = pull_let u l;
+      in Local_Defs.unfold_tac ctxt [rewr] thm end;
+
+    fun pull_tac u ctxt = SELECT_GOAL (pull_tac' u ctxt) 1;
+  \<close>
+
+  ML \<open>
+    val th = @{thm succs_impl_def}
+    val r = get_rhs th;
+    val u1 = @{term "IArray.sub (IArray (map (IArray o map int) k))"};
+    val rewr1 = pull_let u1 r;
+    val r2 = get_rhs rewr1;
+    val u2 = @{term "inv_fun"};
+    val rewr2 = pull_let u2 r2;
+    val r3 = get_rhs rewr2;
+    val u3 = @{term "trans_impl"};
+    val rewr3 = pull_let u3 r3;
+    val mytac = fn ctxt => SELECT_GOAL (Local_Defs.unfold_tac ctxt [rewr1, rewr2, rewr3]) 1;
+  \<close>
+
+  lemma inv_fun_alt_def:
+    "inv_fun i \<equiv> let I = (IArray inv) in I !! i"
+  unfolding inv_fun_def by simp
+
+  lemma inv_fun_rewr:
+    "(let I0 = trans_impl; I = inv_fun; I1 = y in xx I0 I I1) \<equiv>
+     (let I0 = trans_impl; I = (IArray inv); I' = \<lambda> i. I !! i; I1 = y in xx I0 I' I1)"
+  unfolding inv_fun_def[abs_def] by simp
+
+  lemma inv_fun_rewr':
+    "(let I = inv_fun in xx I) \<equiv>
+     (let I = (IArray inv); I' = \<lambda> i. I !! i in xx I')"
+  unfolding inv_fun_def[abs_def] by simp
+
+  schematic_goal succs_impl_alt_def':
+    "succs_impl \<equiv> ?impl"
+  unfolding succs_impl_def
+   apply (tactic \<open>mytac @{context}\<close>)
+   unfolding inv_fun_rewr' trans_impl_def[abs_def]
+   apply (tactic \<open>pull_tac @{term "IArray trans_map"} @{context}\<close>)
+  by (rule Pure.reflexive)
+
+  schematic_goal succs_impl_alt_def:
+    "succs_impl \<equiv> ?impl"
+  unfolding succs_impl_def
+   apply (tactic \<open>pull_tac @{term "IArray.sub (IArray (map (IArray o map int) k))"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "inv_fun"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "trans_impl"} @{context}\<close>)
+   unfolding inv_fun_def[abs_def] trans_impl_def[abs_def]
+   apply (tactic \<open>pull_tac @{term "IArray inv"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "IArray trans_map"} @{context}\<close>)
+  by (rule Pure.reflexive)
+
+  lemma reachability_checker'_alt_def':
+    "reachability_checker' \<equiv>
+      let
+        key = return \<circ> fst;
+        sub = subsumes_impl;
+        copy = state_copy_impl;
+        start = a0_impl;
+        final = F_impl;
+        succs = succs_impl;
+        empty = emptiness_check_impl;
+        trace = tracei
+      in pw_impl key copy trace sub start final succs empty"
+  unfolding reachability_checker'_def by simp
+
+  schematic_goal reachability_checker_alt_def:
+    "reachability_checker \<equiv> ?impl"
+    unfolding reachability_checker_def
+    unfolding reachability_checker'_alt_def' succs_impl_def
+    unfolding E_op''_impl_def abstr_repair_impl_def abstra_repair_impl_def
+    unfolding start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+   apply (tactic \<open>pull_tac @{term "IArray.sub (IArray (map (IArray o map int) k))"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "inv_fun"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "trans_impl"} @{context}\<close>)
+   unfolding inv_fun_def[abs_def] trans_impl_def[abs_def]
+   apply (tactic \<open>pull_tac @{term "IArray inv"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "IArray trans_map"} @{context}\<close>)
+   unfolding trans_map_def label_def
+   unfolding init_dbm_impl_def a0_impl_def
+   unfolding F_impl_def
+   unfolding final_fun_def[abs_def]
+   unfolding subsumes_impl_def
+   unfolding emptiness_check_impl_def
+   unfolding state_copy_impl_def
+  by (rule Pure.reflexive)
+
+  schematic_goal reachability_checker_alt_def:
+    "reachability_checker \<equiv> ?impl"
+  unfolding succs_impl_def reachability_checker_def reachability_checker'_def
+   apply (tactic \<open>pull_tac @{term "IArray.sub (IArray (map (IArray o map int) k))"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "inv_fun"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "trans_impl"} @{context}\<close>)
+   unfolding inv_fun_def[abs_def] trans_impl_def[abs_def]
+   apply (tactic \<open>pull_tac @{term "IArray inv"} @{context}\<close>)
+   apply (tactic \<open>pull_tac @{term "IArray trans_map"} @{context}\<close>)
+  oops
+
+end (* End of locale *)
+
+subsubsection \<open>Check Preconditions\<close>
+context Reachability_Problem_precompiled_defs
+begin
+
+  abbreviation
+    "check_nat_subs \<equiv> \<forall> l < n. \<forall> (_, d) \<in> clkp_set' l. d \<ge> 0"
+
+  lemma check_nat_subs:
+    "check_nat_subs \<longleftrightarrow> (\<forall> l < n. snd ` clkp_set' l \<subseteq> \<nat>)"
+  unfolding Nats_def apply safe
+  subgoal for _ _ _ b using rangeI[of int "nat b"] by (auto 4 3)
+  by (auto 4 3)
+
+  definition
+    "check_pre \<equiv>
+      length inv = n \<and> length trans = n
+      \<and> length k = n \<and> (\<forall> l \<in> set k. length l = m + 1)
+      \<and> m > 0 \<and> n > 0 \<and> trans ! 0 \<noteq> []
+      \<and> (\<forall> l < n. \<forall> (c, d) \<in> clkp_set' l. k ! l ! c \<ge> nat d)
+      \<and> (\<forall> l < n. k ! l ! 0 = 0) \<and> (\<forall> l < n. \<forall> c \<in> {1..m}. k ! l ! c \<ge> 0)
+      \<and> check_nat_subs \<and> clk_set' = {1..m}
+      \<and> (\<forall> xs \<in> set trans. \<forall> (_, _, l) \<in> set xs. l < n)"
+
+  (* Can be optimized with better enumeration *)
+  abbreviation
+    "check_ceiling \<equiv>
+      \<forall> l < n. \<forall> (_, r, l') \<in> set (trans ! l). \<forall> c \<le> m. c \<notin> set r \<longrightarrow> k ! l ! c \<ge> k ! l' ! c"
+
+  lemma finite_clkp_set'[intro, simp]:
+    "finite (clkp_set' l)"
+  unfolding clkp_set'_def by auto
+
+  lemma check_axioms:
+    "Reachability_Problem_precompiled n m k inv trans \<longleftrightarrow> check_pre \<and> check_ceiling"
+  unfolding Reachability_Problem_precompiled_def check_pre_def check_nat_subs by auto
+
+end
+
+subsection \<open>Executable Checker\<close>
+
+lemmas Reachability_Problem_precompiled_defs.check_axioms[code]
+
+lemmas Reachability_Problem_precompiled_defs.clkp_set'_def[code]
+
+lemmas Reachability_Problem_precompiled_defs.clk_set'_def[code]
+
+lemmas Reachability_Problem_precompiled_defs.check_pre_def[code]
+
+lemmas Show_State_Defs.tracei_def[code]
+
+export_code Reachability_Problem_precompiled in SML module_name Test
+
+concrete_definition succs_impl' uses Reachability_Problem_precompiled.succs_impl_alt_def
+
+concrete_definition reachability_checker_impl
+  uses Reachability_Problem_precompiled.reachability_checker_alt_def
+
+export_code reachability_checker_impl in SML_imp module_name TA
+
+definition [code]:
+  "check_and_verify n m k I T final \<equiv>
+    if Reachability_Problem_precompiled n m k I T
+    then reachability_checker_impl m k I T final \<bind> (\<lambda> x. return (Some x))
+    else return None"
+
+lemmas [sep_heap_rules] = Reachability_Problem_precompiled.reachability_checker_hoare'
+
+theorem reachability_check:
+  "(uncurry0 (check_and_verify n m k I T final),
+    uncurry0 (
+       RETURN (
+        if (Reachability_Problem_precompiled n m k I T)
+        then Some (
+          if (\<forall> u. (\<forall> c \<in> {1..m}. u c = 0)
+              \<longrightarrow> u \<turnstile> inv_of (conv_A (Reachability_Problem_precompiled_defs.A n I T)) 0)
+          then
+            if
+              \<exists> l' u u'.
+              (conv_A (Reachability_Problem_precompiled_defs.A n I T)) \<turnstile> \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>
+              \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final
+            then REACHABLE
+            else UNREACHABLE
+          else INIT_INV_ERR
+          )
+        else None
+       )
+    )
+   )
+    \<in> unit_assnk \<rightarrow>a id_assn"
+proof -
+  define inv_pre where
+    "inv_pre \<equiv>
+    (\<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow>
+     u \<turnstile> inv_of (conv_A (Reachability_Problem_precompiled_defs.A n I T)) 0)"
+  define reach_check where
+    "reach_check =
+    (\<exists> l' u u'.
+      (conv_A (Reachability_Problem_precompiled_defs.A n I T)) \<turnstile>
+      \<langle>0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> l' \<in> set final)"
+  note [sep_heap_rules] =
+    Reachability_Problem_precompiled.reachability_checker_hoare'
+    [of n m k I T final,
+      unfolded inv_pre_def[symmetric],
+      unfolded reach_check_def[symmetric]
+      ]
+  show ?thesis
+    unfolding inv_pre_def[symmetric] reach_check_def[symmetric]
+    apply sepref_to_hoare
+    unfolding check_and_verify_def
+    by (sep_auto simp: reachability_checker_impl.refine[symmetric])
+qed
+
+export_code open check_and_verify checking SML
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Semantic_Refinement.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/Normalized_Zone_Semantics_Impl_Semantic_Refinement.thy
@@ -0,0 +1,1800 @@
+theory Normalized_Zone_Semantics_Impl_Semantic_Refinement
+  imports TA_Impl_Misc Floyd_Warshall.Floyd_Warshall
+    Difference_Bound_Matrices.FW_More
+    Normalized_Zone_Semantics_Impl
+    Worklist_Algorithms.Liveness_Subsumption
+    Munta_Base.TA_Syntax_Bundles
+begin
+
+section \<open>Semantic Refinement of the Reachability Checker\<close>
+
+unbundle no_library_syntax
+
+hide_const D
+
+hide_fact upd_def
+
+
+lemma FWI'_characteristic:
+  "canonical_subs n (I \<union> {a::nat}) (curry (FWI' M n a)) \<or> check_diag n (FWI' M n a)" if
+  "canonical_subs n I (curry M)" "I \<subseteq> {0..n}" "a \<le> n"
+  using fwi_characteristic[of n I "curry M", OF that]
+  unfolding check_diag_def neutral FWI'_def by simp
+
+lemma FWI'_check_diag_preservation:
+  "check_diag n (FWI' M n a)" if "check_diag n M"
+  using that fwi_mono[of _ n _ "curry M" a n n] unfolding check_diag_def FWI'_def FWI_def by force
+
+lemma diag_conv_M:
+  "\<forall>i\<le>n. conv_M D (i, i) \<le> 0" if "\<forall>i\<le>n. D (i, i) \<le> 0"
+  using that by auto (metis DBMEntry.simps(15) conv_dbm_entry_mono neutral of_int_0)
+
+lemma conv_M_diag:
+  "\<forall>i\<le>n. D (i, i) \<le> 0" if "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+  using that by (simp add: conv_dbm_entry_mono_rev neutral)
+
+lemma curry_conv_M_swap:
+  "(map_DBMEntry real_of_int \<circ>\<circ>\<circ> curry) M = curry (conv_M M)"
+  by (intro ext; simp)
+
+lemma canonical_subs_subset:
+  "canonical_subs n I' M" if "canonical_subs n I M" "I' \<subseteq> I"
+  using that unfolding canonical_subs_def by auto
+
+lemma n_eq_equiv:
+  "[M1]v,n = [M2]v,n" if "M1 =n M2"
+  using that unfolding DBM_zone_repr_def n_eq_def DBM_val_bounded_def by auto
+
+definition
+  "repair_pair n M a b = FWI' (FWI' M n b) n a"
+
+context TA_Start_Defs
+begin
+
+definition
+  "abstra_repair ac M = repair_pair n (abstra_upd ac M) 0 (constraint_clk ac)"
+
+definition
+  "filter_diag f M \<equiv> if check_diag n M then M((0,0) := Lt 0) else f M"
+
+definition abstr_repair where
+  "abstr_repair = fold (\<lambda> ac M. abstra_repair ac M)"
+
+definition
+  "E_op l r g l' M \<equiv>
+    let
+      M' = FW' (abstr_upd (inv_of A l) (up_canonical_upd M n)) n;
+      M'' = FW' (abstr_upd (inv_of A l') (reset'_upd (FW' (abstr_upd g M') n) n r 0)) n;
+      M''' = FW' (norm_upd M'' (k' l') n) n
+    in M'''"
+
+definition
+  "E_op' l r g l' M \<equiv>
+    let
+      M' = abstr_repair (inv_of A l) (up_canonical_upd M n);
+      M'' = abstr_repair (inv_of A l') (reset'_upd (abstr_repair g M') n r 0);
+      M''' = FW' (norm_upd M'' (k' l') n) n
+    in M'''"
+
+definition
+  "E_op'' l r g l' M \<equiv>
+    let
+      M1 = abstr_repair (inv_of A l) (up_canonical_upd M n);
+      M2 = filter_diag (\<lambda> M. abstr_repair g M) M1;
+      M3 = filter_diag (\<lambda> M. abstr_repair (inv_of A l') (reset'_upd M n r 0)) M2;
+      M4 = filter_diag (\<lambda> M. FW' (norm_upd M (k' l') n) n) M3
+    in M4"
+
+lemma check_diag_marker:
+  "check_diag n (M((0, 0) := Lt 0))"
+  unfolding check_diag_def by (auto simp: Lt_lt_LeI)
+
+lemma E_op''_alt_def:
+  "E_op'' l r g l' M \<equiv>
+    let
+      M' = abstr_repair (inv_of A l) (up_canonical_upd M n);
+      f1 = \<lambda> M. abstr_repair g M;
+      f2 = \<lambda> M. abstr_repair (inv_of A l') (reset'_upd M n r 0);
+      f3 = \<lambda> M. FW' (norm_upd M (k' l') n) n
+    in filter_diag (filter_diag f3 o filter_diag f2 o f1) M'" if "n > 0"
+  unfolding E_op''_def filter_diag_def
+  by (rule HOL.eq_reflection) (auto simp: Let_def check_diag_marker)
+
+end (* End of context for reachability problem defs *)
+
+lemma Bisimulation_Invariant_simulation_strengthen:
+  assumes "Bisimulation_Invariant A B sim PA PB"
+          "\<And> a b. sim a b \<Longrightarrow> PA a \<Longrightarrow> PB b \<Longrightarrow> R a b"
+    shows "Bisimulation_Invariant A B (\<lambda> a b. sim a b \<and> R a b) PA PB"
+proof -
+  interpret Bisimulation_Invariant A B sim PA PB by (rule assms)
+  show ?thesis
+    by (standard; (blast dest: A_B_step intro: assms(2) | blast dest: B_A_step intro: assms(2)))
+qed
+
+context TA_Start
+begin
+
+context
+  fixes E1 :: "'s \<times> _ \<Rightarrow> 's \<times> _ \<Rightarrow> bool"
+  assumes E_E1_step: "E a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E1 a c \<and> b \<sim> c)"
+  assumes E1_E_step: "E1 a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E a c \<and> b \<sim> c)"
+  assumes E1_wf_state[intro]: "wf_state a \<Longrightarrow> E1 a b \<Longrightarrow> wf_state b"
+begin
+
+lemma E1_steps_wf_state[intro]:
+  "wf_state b" if "E1** a b" "wf_state a"
+  using that by (induction rule: rtranclp_induct) auto
+
+lemma E_E1_step':
+  "(\<exists> b'. E1 b b' \<and> a' \<sim> b')" if "E a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that E_equiv[OF that] by (blast dest: E_E1_step)
+
+lemma E1_E_step':
+  "(\<exists> b'. E b b' \<and> a' \<sim> b')" if "E1 a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that
+  apply -
+  apply (drule E1_E_step, assumption)
+  apply safe
+  by (drule E_equiv; blast)
+
+lemma E_E1_steps:
+  "\<exists> b'. E1** b b' \<and> a' \<sim> b'" if "E** a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that
+  apply (induction rule: rtranclp_induct)
+   apply blast
+  apply clarsimp
+  apply (drule E_E1_step')
+     apply blast
+    prefer 2
+    apply blast
+   apply blast
+  by (auto intro: rtranclp.intros(2))
+
+lemma E1_E_steps:
+  "\<exists> b'. E** b b' \<and> a' \<sim> b'" if "E1** a a'" "wf_state a" "wf_state b" "a \<sim> b"
+  using that
+  apply (induction rule: rtranclp_induct)
+   apply blast
+  apply clarsimp
+  apply (drule E1_E_step')
+     apply blast
+    prefer 2
+    apply blast
+   apply blast
+  by (auto intro: rtranclp.intros(2))
+
+lemma E_E1_steps_empty:
+  "(\<exists> l' M'. E** a0 (l', M') \<and> [curry (conv_M M')]v,n = {}) \<longleftrightarrow>
+   (\<exists> l' M'. E1** a0 (l', M') \<and> [curry (conv_M M')]v,n = {})"
+  by (auto 4 4 simp: state_equiv_def dbm_equiv_def dest: E_E1_steps E1_E_steps)
+
+lemma E_E1_bisim:
+  "Bisimulation_Invariant E E1 (\<sim>) wf_state wf_state"
+  by standard (blast intro: state_equiv_sym dest: E1_E_step' intro: E_E1_step')+
+
+context
+  fixes P :: "'s \<times> _ \<Rightarrow> bool"
+    assumes P: "P a \<Longrightarrow> a \<sim> b \<Longrightarrow> wf_state a \<Longrightarrow> wf_state b \<Longrightarrow> P b"
+  begin
+
+lemma E_E1_steps_equiv:
+  "(\<exists> l' M'. E** a0 (l', M') \<and> P (l', M')) \<longleftrightarrow>
+   (\<exists> l' M'. E1** a0 (l', M') \<and> P (l', M'))"
+  apply safe
+  subgoal
+    by (frule E_E1_steps, safe; blast intro: P)
+  subgoal
+    by (frule E1_E_steps, safe; blast intro: P)
+  done
+
+lemma E_E1_bisim':
+  "Bisimulation_Invariant E E1 (\<lambda> a b. a \<sim> b \<and> (P a \<longleftrightarrow> P b)) wf_state wf_state"
+  by (rule Bisimulation_Invariant_simulation_strengthen; blast intro: state_equiv_sym P E_E1_bisim)
+
+end
+
+lemma E1_mono:
+  assumes "E1 (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+  shows "\<exists> M'. E1 (l,M) (l',M') \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+  using assms
+  apply -
+  apply (drule E1_E_step)
+   apply (simp add: wf_state_def; fail)
+  apply safe
+  apply (drule E_mono[where M = M], assumption+)
+  apply safe
+  by (drule E_E1_step; force simp: wf_state_def state_equiv_def dbm_equiv_def)
+
+lemma E1_wf_dbm[intro]:
+  "wf_dbm M \<Longrightarrow> E1 (l, M) (l', M') \<Longrightarrow> wf_dbm M'"
+  using E1_wf_state unfolding wf_state_def by auto
+
+(* Duplication (E_mono') *)
+lemma E1_mono':
+  assumes "E1 (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "dbm_subset n D M"
+  shows "\<exists> M'. E1 (l,M) (l',M') \<and> dbm_subset n D' M'"
+  using assms
+  apply -
+  apply (frule E1_mono[where M = M], assumption+)
+   apply (subst dbm_subset_correct'', assumption+)
+   apply (rule dbm_subset_conv, assumption)
+  apply safe
+  apply (subst (asm) dbm_subset_correct'')
+    apply (blast intro: dbm_subset_conv_rev)+
+  done
+
+end (* End of anonymous context *)
+
+lemma canonical'_up_canonical_upd:
+  assumes "canonical' M"
+  shows "canonical' (up_canonical_upd M n)"
+  using
+    up_canonical_preservation[OF assms] up_canonical_eq_up[OF assms]
+    up_canonical_upd_up_canonical'[of M n]
+  unfolding curry_def n_eq_def by auto
+
+lemma check_diag_up_canonical_upd:
+  "check_diag n (up_canonical_upd M n)" if "check_diag n M"
+  unfolding up_canonical_upd_def
+  apply (rule fold_acc_preserv'[where P = "check_diag n"])
+  using that by (auto simp: check_diag_def)
+
+lemma canonical_diag'_up_canonical_upd:
+  "canonical_diag' (up_canonical_upd M n)" if "canonical_diag' M"
+  using that by (blast dest: canonical'_up_canonical_upd check_diag_up_canonical_upd)
+
+lemma canonical_subs'_upd:
+  "canonical_subs' (I - {a, b}) (M((a, b) := c))" if "canonical_subs' I M" "c \<le> M (a, b)"
+  using that unfolding canonical_subs_def by (auto 4 4)
+
+lemma canonical'_conv_M_iff:
+  "canonical' (conv_M D) \<longleftrightarrow> canonical' D"
+  by (metis canonical_conv canonical_conv_rev)
+
+lemma canonical_subs'_subset:
+  "canonical_subs' I' M" if "canonical_subs' I M" "I' \<subseteq> I"
+  using that by (rule canonical_subs_subset)
+
+lemma wf_dbm_altD:
+  "canonical' D \<or> check_diag n D"
+  "\<forall>i\<le>n. D (i, i) \<le> 0"
+  "[curry (conv_M D)]v,n \<subseteq> V"
+  if "wf_dbm D"
+  using that
+    apply -
+    apply (drule wf_dbm_D, simp only: canonical'_conv_M_iff; fail)
+   apply (drule wf_dbm_D, auto intro!: conv_M_diag; fail)
+  apply (drule wf_dbm_D, erule valid_dbm.cases, simp; fail)
+  done
+
+lemmas valid_dbm_convI = valid_dbm.intros[OF _ dbm_int_conv]
+
+lemmas wf_dbm_altI = wf_dbm_I[OF _ diag_conv_M valid_dbm_convI]
+\<comment> \<open>Note that there is also the following:\<close>
+thm valid_dbm_def
+
+lemma wf_dbm_rule:
+  assumes "wf_dbm D"
+  assumes canonical:
+    "canonical' D \<Longrightarrow> \<forall>i\<le>n. D (i, i) \<le> 0 \<Longrightarrow> \<forall>i\<le>n. D (i, i) = 0 \<Longrightarrow> [curry (conv_M D)]v,n \<subseteq> V
+    \<Longrightarrow> canonical' (f D) \<or> check_diag n (f D)
+    "
+  assumes diag:
+    "check_diag n D \<Longrightarrow> \<forall>i\<le>n. D (i, i) \<le> 0 \<Longrightarrow> [curry (conv_M D)]v,n \<subseteq> V
+    \<Longrightarrow> check_diag n (f D)
+    "
+  assumes diag_le:
+    "canonical' D \<or> check_diag n D \<Longrightarrow> \<forall>i\<le>n. D (i, i) \<le> 0 \<Longrightarrow> [curry (conv_M D)]v,n \<subseteq> V
+    \<Longrightarrow> \<forall>i\<le>n. f D (i, i) \<le> 0
+    "
+  assumes V:
+    "canonical' D \<Longrightarrow> \<forall>i\<le>n. D (i, i) \<le> 0 \<Longrightarrow> [curry (conv_M D)]v,n \<subseteq> V
+    \<Longrightarrow> [curry (conv_M (f D))]v,n \<subseteq> V
+    "
+  shows "wf_dbm (f D)"
+proof -
+  from wf_dbm_altD[OF assms(1)] have facts:
+    "canonical' D \<or> check_diag n D" "\<forall>i\<le>n. D (i, i) \<le> 0" "[curry (conv_M D)]v,n \<subseteq> V"
+    .
+  from this(1) consider "canonical' D \<and> \<not> check_diag n D" | "check_diag n D"
+    by auto
+  then show ?thesis
+  proof cases
+    assume A: "canonical' D \<and> \<not> check_diag n D"
+    with facts(2) have "\<forall>i\<le>n. D (i, i) = 0"
+      unfolding check_diag_def neutral[symmetric] by fastforce
+    with A[THEN conjunct1] show ?thesis
+      by (intro wf_dbm_altI[unfolded canonical'_conv_M_iff] canonical diag_le V facts)
+  next
+    assume "check_diag n D"
+    then have "check_diag n (f D)" by (intro disjI2 diag facts)
+    then have "[curry (conv_M (f D))]v,n = {}" by (intro check_diag_conv_M check_diag_empty_spec)
+    with \<open>check_diag n (f D)\<close> show ?thesis
+      by - (rule wf_dbm_altI[unfolded canonical'_conv_M_iff];
+           (intro diag_le diag facts; fail | blast)+)
+  qed
+qed
+
+lemma canonical_subs'_abstra_upd:
+  "canonical_subs' (I - {0, constraint_clk ac}) (abstra_upd ac M)" if
+  "canonical_subs' I M" "constraint_clk ac > 0"
+  using that
+  apply (cases ac)
+  subgoal for c d
+    by (force
+        dest: canonical_subs'_upd[of I M "Lt d" c 0] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+  subgoal for c d
+    by (force
+        dest: canonical_subs'_upd[of I M "Le d" c 0] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+    defer
+  subgoal for c d
+    by (force
+        dest: canonical_subs'_upd[of I M "Lt (-d)" 0 c] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+  subgoal for c d
+    by (force
+        dest: canonical_subs'_upd[of I M "Le (-d)" 0 c] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+  subgoal for c d
+    apply (auto simp: min_def Let_def)
+       apply (force
+        dest: canonical_subs'_upd[of I M "Le (-d)" 0 c] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+      apply (force
+        dest: canonical_subs'_upd[of I M "Le d" c 0] simp: insert_commute min_def
+        intro: canonical_subs'_subset
+        )
+    subgoal
+      apply (drule canonical_subs'_upd[of I M "Le (-d)" 0 c])
+      by simp (metis fun_upd_apply fun_upd_triv gr_implies_not_zero old.prod.inject)
+    subgoal
+      by (drule canonical_subs'_upd[of I M "Le (-d)" 0 c];
+          force dest: canonical_subs'_upd[of _ _ "Le d" c 0])
+    done
+  done
+
+lemmas FWI_zone_equiv_spec = FWI_zone_equiv[OF surj_numbering]
+
+lemma repair_pair_equiv:
+  "M \<simeq> repair_pair n M a b" if "a \<le> n" "b \<le> n"
+  unfolding dbm_equiv_def repair_pair_def FWI'_def
+  apply (subst FWI_zone_equiv_spec[of b], rule that)
+  apply (subst FWI_zone_equiv_spec[of a], rule that)
+  by (simp add: FWI_def fwi_conv_M' fwi_conv_M'' curry_conv_M_swap)
+
+lemma repair_pair_canonical_diag_1:
+  "canonical_diag (repair_pair n M a b)" if "canonical_subs' ({0..n} - {a, b}) M" "a \<le> n" "b \<le> n"
+  using that
+  unfolding canonical'_conv_M_iff
+  unfolding repair_pair_def
+  apply -
+  apply (drule FWI'_characteristic[where a = b], force, assumption)
+  apply (erule disjE)
+   apply (drule FWI'_characteristic[where a = a], force, assumption)
+  unfolding canonical_alt_def
+  subgoal premises prems
+  proof -
+    have *: "{0..n} - {a, b} \<union> {b} \<union> {a} = {0..n}"
+      using that(2,3) by fastforce
+    with prems(3) show ?thesis by (simp only: *)
+  qed
+  by (auto intro: FWI'_check_diag_preservation)
+
+lemma repair_pair_canonical_diag_2:
+  "canonical_diag (repair_pair n M a b)" if "check_diag n M"
+  unfolding repair_pair_def using that by (auto intro: FWI'_check_diag_preservation)
+
+lemma repair_pair_diag_le:
+  "\<forall>i\<le>n. repair_pair n M 0 c (i, i) \<le> 0" if "\<forall>i\<le>n. M (i, i) \<le> 0"
+  using that
+  unfolding repair_pair_def FWI'_def
+  apply clarsimp
+  apply (rule order.trans[OF FWI_mono], assumption, assumption)
+  apply (rule order.trans[OF FWI_mono], assumption, assumption)
+  by simp
+
+lemma conv_abstra_upd:
+  "conv_M (abstra_upd ac M) = abstra_upd (conv_ac ac) (conv_M M)"
+  by (cases ac)
+     (auto 4 3 split: split_min
+       dest: not_le_imp_less conv_dbm_entry_mono_strict conv_dbm_entry_mono)
+
+lemma abstra_upd_conv_M_zone_equiv:
+  assumes "constraint_clk ac > 0" "constraint_clk ac \<le> n"
+  shows "[curry (conv_M (abstra_upd ac M))]v,n = {u. u \<turnstile>a conv_ac ac} \<inter> [curry (conv_M M)]v,n"
+  apply (subst conv_abstra_upd)
+  apply (subst abstra_upd_eq_abstra')
+   defer
+   apply (subst abstra_abstra'[symmetric, where v = v])
+    defer
+    apply (subst dbm_abstra_zone_eq[OF clock_numbering(1)])
+  using assms by (cases ac; auto simp: v_def)+
+
+lemma abstra_upd_conv_M_V:
+  assumes "[curry (conv_M M)]v,n \<subseteq> V" "constraint_clk ac > 0" "constraint_clk ac \<le> n"
+  shows "[curry (conv_M (abstra_upd ac M))]v,n \<subseteq> V"
+  using assms(1) by (auto simp: abstra_upd_conv_M_zone_equiv[OF assms(2-)])
+
+lemma abstra_repair_check_diag_preservation:
+  "check_diag n (abstra_repair ac M)" if "check_diag n M" "constraint_clk ac > 0"
+proof -
+  have "check_diag n (abstra_upd ac M)"
+    using that unfolding check_diag_def by (auto simp: abstra_upd_diag_preservation)
+  then show ?thesis
+    unfolding abstra_repair_def repair_pair_def by (auto intro: FWI'_check_diag_preservation)
+qed
+
+lemma wf_dbm_abstra_repair:
+  assumes "wf_dbm M" "constraint_clk ac > 0" "constraint_clk ac \<le> n"
+  shows "wf_dbm (abstra_repair ac M)"
+proof -
+  note facts = wf_dbm_altD[OF assms(1)]
+  let ?M = "abstra_upd ac M"
+  let ?MM = "abstra_repair ac M"
+  from assms(2) facts(2) have "\<forall>i\<le>n. ?M (i, i) \<le> 0"
+    by (auto simp: abstra_upd_diag_preservation)
+  then have diag: "\<forall>i\<le>n. ?MM (i, i) \<le> 0" unfolding abstra_repair_def by (rule repair_pair_diag_le)
+  from assms(2,3) facts(3) have "[curry (conv_M ?M)]v,n \<subseteq> V" by - (rule abstra_upd_conv_M_V)
+  with repair_pair_equiv[of 0 "constraint_clk ac" ?M] assms(3) have V:
+    "[curry (conv_M ?MM)]v,n \<subseteq> V"
+    unfolding dbm_equiv_def abstra_repair_def by simp
+  from facts(1) have "canonical_diag ?MM"
+  proof
+    assume "canonical' M"
+    from canonical_subs'_abstra_upd[OF this[unfolded canonical_alt_def] assms(2)] show
+      "canonical_diag ?MM"
+      unfolding abstra_repair_def by (intro repair_pair_canonical_diag_1 assms; simp)
+  next
+    assume "check_diag n M"
+    then have "check_diag n ?M"
+      unfolding check_diag_def using assms(2) by (auto simp: abstra_upd_diag_preservation)
+    then show "canonical_diag ?MM"
+      unfolding abstra_repair_def by (rule repair_pair_canonical_diag_2)
+  qed
+  with diag V show ?thesis by (intro wf_dbm_altI)
+qed
+
+lemma wf_dbm_abstra_repair_equiv:
+  assumes "constraint_clk ac \<le> n"
+  shows "abstra_repair ac M \<simeq> abstra_upd ac M"
+  by (simp add: abstra_repair_def assms dbm_equiv_sym repair_pair_equiv)
+
+lemma abstra_upd_equiv:
+  assumes "constraint_clk ac > 0" "constraint_clk ac \<le> n" "M \<simeq> M'"
+  shows "abstra_upd ac M \<simeq> abstra_upd ac M'"
+  using assms(3) abstra_upd_conv_M_zone_equiv[OF assms(1,2)] unfolding dbm_equiv_def by auto
+
+lemma wf_dbm_abstra_repair_equiv':
+  assumes "constraint_clk ac > 0" "constraint_clk ac \<le> n" "M \<simeq> M'"
+  shows "abstra_repair ac M \<simeq> abstra_upd ac M'"
+  using wf_dbm_abstra_repair_equiv[OF assms(2)] abstra_upd_equiv[OF assms] by blast
+
+lemma abstr_repair_check_diag_preservation:
+  "check_diag n (abstr_repair cc M)" if "check_diag n M" "\<forall> c \<in> constraint_clk ` set cc. c > 0"
+  using that unfolding abstr_repair_def
+  by (induction cc arbitrary: M) (auto intro!: abstra_repair_check_diag_preservation)
+
+lemma wf_dbm_abstr_repair_equiv:
+  assumes "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n" "M \<simeq> M'"
+  shows "abstr_repair cc M \<simeq> abstr_upd cc M'"
+  using assms unfolding abstr_repair_def abstr_upd_def
+  by (induction cc arbitrary: M M') (auto dest!: wf_dbm_abstra_repair_equiv')
+
+lemma wf_dbm_abstr_repair_equiv':
+  assumes "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n" "M \<simeq> M'"
+  shows "abstr_repair cc M \<simeq> abstr_repair cc M'"
+  using assms by (meson wf_dbm_abstr_repair_equiv dbm_equiv_sym dbm_equiv_trans dbm_equiv_refl)
+
+lemma wf_dbm_abstr_repair:
+  assumes "wf_dbm M" "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n"
+  shows "wf_dbm (abstr_repair cc M)"
+  using assms unfolding abstr_repair_def
+  by (induction cc arbitrary: M) (auto intro: wf_dbm_abstra_repair)
+
+lemma abstr_upd_conv_M_V:
+  assumes "[curry (conv_M M)]v,n \<subseteq> V" "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n"
+  shows "[curry (conv_M (abstr_upd cc M))]v,n \<subseteq> V"
+  using assms
+  unfolding abstr_upd_def
+  apply (induction cc arbitrary: M)
+   apply simp
+  subgoal premises prems
+    apply simp
+    apply (intro prems(1)[simplified] abstra_upd_conv_M_V[simplified])
+    using prems(2-) by auto
+  done
+
+lemma wf_dbm_FW'_abstr_upd:
+  "wf_dbm (FW' (abstr_upd cc M) n)" if "wf_dbm M" "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n"
+  apply (rule wf_dbm_rule[OF that(1)])
+  subgoal
+    unfolding check_diag_def neutral[symmetric] by (rule FW'_canonical)
+  subgoal
+    unfolding check_diag_def neutral[symmetric]
+    using that(2)
+    apply safe
+    apply (subst (asm) (2) abstr_upd_diag_preservation[symmetric, where M = M and cc = cc])
+    by (auto dest: FW'_neg_diag_preservation simp: collect_clks_def)
+  subgoal
+    apply (drule abstr_upd_diag_preservation'[where cc = cc])
+    by (simp add: FW'_diag_preservation collect_clks_def that(2))+
+  subgoal
+    by (subst FW'_zone_equiv) (intro abstr_upd_conv_M_V that)
+  done
+
+lemma up_canonical_upd_V:
+  "[curry (up_canonical_upd M n)]v,n \<subseteq> V" if "canonical' M" "[curry M]v,n \<subseteq> V"
+  apply (subst n_eq_equiv[OF up_canonical_upd_up_canonical'])
+  apply (subst up_canonical_equiv_up[OF that(1)])
+  apply (subst up_correct)
+   apply (rule clock_numbering(1))
+    by (rule up_V[OF that(2)])
+
+lemma n_eq_transfer[transfer_rule]: "eq_onp (\<lambda>x. x = n) n n"
+  by (simp add: eq_onp_def)
+
+lemma RI_up_canonical_upd:
+  "RI n (up_canonical_upd Mr n) (up_canonical_upd Mi n)" if "RI n Mr Mi"
+  supply [transfer_rule] = that
+  by transfer_prover
+
+lemma wf_dbm_up_canonical_upd:
+  "wf_dbm (up_canonical_upd M n)" if "wf_dbm M"
+  apply (rule wf_dbm_rule[OF that])
+  subgoal
+    by (intro canonical_diag'_up_canonical_upd disjI1)
+  subgoal
+    by (simp add: check_diag_def up_canonical_upd_diag_preservation)
+  subgoal
+    by (simp add: up_canonical_upd_diag_preservation)
+  subgoal
+    unfolding canonical'_conv_M_iff[symmetric]
+    apply (drule up_canonical_upd_V, assumption)
+    by (subst (asm) RI_zone_equiv[OF RI_up_canonical_upd[OF RI_conv_M]])
+  done
+
+lemma reset_resets_spec:
+  "[reset M n i d]v,n = {u(i := d) |u. u \<in> [M]v,n}" if "i > 0" "i \<le> n" "0 \<le> d"
+proof -
+  from that have v_i[simp]: "v i = i" by (simp add: v_def)
+  show ?thesis
+    apply (subst reset_resets[OF _ clock_numbering(1), of i, unfolded v_i])
+    using clock_numbering(2) that by fastforce+
+qed
+
+lemma reset_canonical_n_eq_reset_canonical_upd:
+  "reset_canonical (curry M) i d =n curry (reset_canonical_upd M n i d)" if "i > 0"
+  using that unfolding n_eq_def by (auto simp: reset_canonical_upd_reset_canonical')
+
+lemma RI_reset_canonical_upd:
+  "RI n (reset_canonical_upd Mr n i (real_of_int d)) (reset_canonical_upd Mi n i d)" if
+  "RI n Mr Mi" "i \<le> n"
+proof -
+  have [transfer_rule]: "ri (real_of_int d) d"
+    by (simp only: ri_def)
+  from that have [transfer_rule]: "eq_onp (\<lambda>x. x < Suc n) i i" by (simp add: eq_onp_def)
+  note [transfer_rule] = that(1)
+  show ?thesis by transfer_prover
+qed
+
+lemma reset_canonical_upd_correct:
+  "[curry (reset_canonical_upd D n i d)]v,n = {u(i := d) |u. u \<in> [curry D]v,n}" if
+  "canonical' D" "i > 0" "i \<le> n" "0 \<le> d"
+  apply (subst n_eq_equiv[OF reset_canonical_n_eq_reset_canonical_upd, symmetric])
+  defer
+   apply (subst reset_reset_canonical[symmetric])
+      prefer 5
+      apply (subst reset_resets_spec)
+  using that clock_numbering(1) by auto
+
+lemma reset_canonical_upd_correct_conv_M:
+  "[curry ((map_DBMEntry real_of_int \<circ>\<circ>\<circ> reset_canonical_upd M n) i d)]v,n =
+    {u(i := real_of_int d) |u. u \<in> [curry (conv_M M)]v,n}"
+  if "i > 0" "i \<le> n" "0 \<le> d" "canonical' M"
+  apply (subst RI_zone_equiv[OF RI_reset_canonical_upd[OF RI_conv_M], symmetric], rule that)
+  apply (subst reset_canonical_upd_correct)
+  unfolding canonical'_conv_M_iff using that by auto
+
+lemma wf_dbm_reset_canonical_upd:
+  "wf_dbm (reset_canonical_upd M n i d)" if "wf_dbm M" "i > 0" "i \<le> n" "0 \<le> d"
+  apply (rule wf_dbm_rule[OF that(1)])
+  subgoal
+    by (intro reset_canonical_upd_canonical disjI1 that)
+  subgoal
+    by (metis check_diag_def reset_canonical_upd_diag_preservation that(2))
+  subgoal
+    by (simp add: reset_canonical_upd_diag_preservation that(2))
+  subgoal
+    unfolding canonical'_conv_M_iff[symmetric]
+    apply (subst RI_zone_equiv[OF RI_reset_canonical_upd[OF RI_conv_M], symmetric], rule that)
+    apply (subst reset_canonical_upd_correct)
+    using that unfolding V_def by auto
+  done
+
+lemma reset_canonical_upd_equiv':
+  "reset_canonical_upd D n i d \<simeq> reset_canonical_upd M n i d" if
+  "D \<simeq> M" "i > 0" "i \<le> n" "0 \<le> d" "canonical' D" "canonical' M"
+  using that unfolding dbm_equiv_def
+  apply (subst reset_canonical_upd_correct_conv_M)
+      prefer 5
+      apply (subst reset_canonical_upd_correct_conv_M)
+  by auto
+
+lemma reset_canonical_upd_check_diag_preservation:
+  "check_diag n (reset_canonical_upd D n i d)" if "check_diag n D" "i > 0"
+  by (metis check_diag_def reset_canonical_upd_diag_preservation that)
+
+lemma reset_canonical_upd_equiv:
+  "reset_canonical_upd D n i d \<simeq> reset_canonical_upd M n i d" if
+  "D \<simeq> M" "i > 0" "i \<le> n" "0 \<le> d" "canonical_diag' D" "canonical_diag' M"
+proof -
+  have *: "check_diag n D \<longleftrightarrow> check_diag n M"
+    using that unfolding dbm_equiv_def
+      apply -
+    apply (subst canonical_check_diag_empty_iff[symmetric], assumption)
+    apply (subst canonical_check_diag_empty_iff[symmetric], assumption)
+      by simp
+  from that(5) consider "canonical' D" "\<not> check_diag n D" | "check_diag n D"
+    by auto
+  then show ?thesis
+  proof cases
+    case 1
+    with * that(6) have "canonical' D" "canonical' M"
+      by auto
+    with that show ?thesis
+      by (auto intro: reset_canonical_upd_equiv')
+  next
+    case 2
+    then have "check_diag n D" "check_diag n M" unfolding * by assumption+
+    then have
+      "check_diag n (reset_canonical_upd D n i d)" "check_diag n (reset_canonical_upd M n i d)"
+      by - (intro reset_canonical_upd_check_diag_preservation that(2), assumption)+
+    then show ?thesis by (auto dest!: check_diag_empty_spec check_diag_conv_M simp: dbm_equiv_def)
+  qed
+qed
+
+lemma reset'_upd_check_diag_preservation:
+  "check_diag n (reset'_upd M n cs d)" if "check_diag n M" "\<forall> i \<in> set cs. i > 0"
+  using that unfolding reset'_upd_def
+  by (induction cs arbitrary: M) (auto intro: reset_canonical_upd_check_diag_preservation)
+
+lemma wf_dbm_reset'_upd:
+  "wf_dbm (reset'_upd M n cs d)" if "wf_dbm M" "\<forall> i \<in> set cs. i > 0 \<and> i \<le> n" "0 \<le> d"
+  using that unfolding reset'_upd_def
+  by (induction cs arbitrary: M) (auto intro: wf_dbm_reset_canonical_upd)
+
+lemma reset'_upd_equiv:
+  "reset'_upd D n cs d \<simeq> reset'_upd M n cs d" if
+  "D \<simeq> M" "\<forall> i \<in> set cs. i > 0 \<and> i \<le> n" "0 \<le> d" "wf_dbm M" "wf_dbm D"
+  using that unfolding reset'_upd_def
+  apply (induction cs arbitrary: D M)
+   apply (simp; fail)
+  subgoal premises prems for c cs D M
+    apply simp
+    apply (rule prems(1))
+        apply (rule reset_canonical_upd_equiv)
+    using prems(2-)
+    by (auto intro: wf_dbm_reset_canonical_upd dest: wf_dbm_altD)
+  done
+
+lemma E_E_op: "E = (\<lambda> (l, M) (l', M'''). \<exists> g a r. A \<turnstile> l \<longrightarrow>g,a,r l' \<and> M''' = E_op l r g l' M)"
+  unfolding E_op_def E_alt_def by simp
+
+end (* End of locale for Reachability Problem *)
+
+locale E_From_Op_Defs = TA_Start_Defs _ l0 for l0 :: "'s" +
+  fixes f :: "'s
+   \<Rightarrow> nat list
+      \<Rightarrow> (nat, int) acconstraint list
+         \<Rightarrow> 's
+            \<Rightarrow> (nat \<times> nat \<Rightarrow> int DBMEntry)
+               \<Rightarrow> nat \<times> nat \<Rightarrow> int DBMEntry"
+begin
+
+definition "E_from_op = (\<lambda> (l, M) (l', M'''). \<exists> g a r. A \<turnstile> l \<longrightarrow>g,a,r l' \<and> M''' = f l r g l' M)"
+
+end
+
+locale E_From_Op = TA_Start + E_From_Op_Defs
+
+locale E_From_Op_Bisim = E_From_Op +
+  assumes op_bisim:
+    "A \<turnstile> l \<longrightarrow>g,a,r l' \<Longrightarrow> wf_dbm M \<Longrightarrow> f l r g l' M \<simeq> E_op l r g l' M"
+    and op_wf:
+    "A \<turnstile> l \<longrightarrow>g,a,r l' \<Longrightarrow> wf_dbm M \<Longrightarrow> wf_dbm (f l r g l' M)"
+begin
+
+lemma E_E_from_op_step:
+  "E a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E_from_op a c \<and> b \<sim> c)"
+  unfolding E_E_op E_from_op_def wf_state_def state_equiv_def
+  by (blast intro: op_bisim dbm_equiv_sym)
+
+lemma E_from_op_E_step:
+  "E_from_op a b \<Longrightarrow> wf_state a \<Longrightarrow> (\<exists> c. E a c \<and> b \<sim> c)"
+  unfolding E_E_op E_from_op_def wf_state_def state_equiv_def by (blast intro: op_bisim)
+
+lemma E_from_op_wf_state:
+  "wf_state a \<Longrightarrow> E_from_op a b \<Longrightarrow> wf_state b"
+  unfolding E_E_op E_from_op_def wf_state_def state_equiv_def by (blast intro: op_wf)
+
+lemma E_E_from_op_steps_empty:
+  "(\<exists>l' M'. E** a0 (l', M') \<and> [curry (conv_M M')]v,n = {}) \<longleftrightarrow>
+   (\<exists>l' M'. E_from_op** a0 (l', M') \<and> [curry (conv_M M')]v,n = {})"
+  by (rule E_E1_steps_empty[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state])
+
+theorem E_from_op_reachability_check:
+  "(\<exists> l' D'. E** a0 (l', D') \<and> F_rel (l', D'))
+  \<longleftrightarrow> (\<exists> l' D'. E_from_op** a0 (l', D') \<and> F_rel (l', D'))"
+  if F_rel_def: "F_rel = (\<lambda>(l, D). F l \<and> \<not> check_diag n D)"
+  by (rule E_E1_steps_equiv[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state])
+    (auto
+      simp: F_rel_def state_equiv_def wf_state_def dbm_equiv_def
+      dest!:
+        check_diag_empty_spec[OF check_diag_conv_M]
+        canonical_check_diag_empty_iff[OF wf_dbm_altD(1)]
+    )
+
+lemma E_from_op_mono:
+  assumes "E_from_op (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "[curry (conv_M D)]v,n \<subseteq> [curry (conv_M M)]v,n"
+  shows "\<exists> M'. E_from_op (l,M) (l',M') \<and> [curry (conv_M D')]v,n \<subseteq> [curry (conv_M M')]v,n"
+  using assms by - (rule E1_mono[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state]; blast)
+
+lemma E_from_op_mono':
+  assumes "E_from_op (l,D) (l',D')"
+    and   "wf_dbm D" "wf_dbm M"
+    and "dbm_subset n D M"
+  shows "\<exists> M'. E_from_op (l,M) (l',M') \<and> dbm_subset n D' M'"
+  using assms by - (rule E1_mono'[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state]; blast)
+
+lemma E_from_op_bisim:
+  "Bisimulation_Invariant E E_from_op (\<sim>) wf_state wf_state"
+  by (rule E_E1_bisim[OF E_E_from_op_step E_from_op_E_step E_from_op_wf_state])
+
+end (* End of context for bisimilarity *)
+
+locale E_From_Op_Finite = E_From_Op +
+  assumes step_FW_norm:
+    "A \<turnstile> l \<longrightarrow>g,a,r l' \<Longrightarrow> dbm_default (curry M) n \<Longrightarrow> dbm_int (curry M) n
+    \<Longrightarrow> \<not> check_diag n (f l r g l' M)
+    \<Longrightarrow> \<exists> M'.
+      f l r g l' M = FW' (norm_upd M' (k' l') n) n
+      \<and> dbm_default (curry M') n \<and> dbm_int (curry M') n"
+  and check_diag:
+    "A \<turnstile> l \<longrightarrow>g,a,r l' \<Longrightarrow> check_diag n M \<Longrightarrow> check_diag n (f l r g l' M)"
+begin
+
+(* Duplication *)
+lemma steps_FW_norm:
+  assumes "E_from_op** (l, D) (l', D')"
+    and "dbm_default (curry D) n"
+    and "dbm_int (curry D) n"
+  shows "l' = l \<and> D' = D \<or> check_diag n D' \<or> (\<exists>M. D' = FW' (norm_upd M (k' l') n) n \<and>
+             ((\<forall>i>n. \<forall>j. curry M i j = 0) \<and> (\<forall>j>n. \<forall>i. curry M i j = 0)) \<and> dbm_int (curry M) n)"
+using assms proof (induction "(l', D')" arbitrary: l' D')
+  case base then show ?case by auto
+next
+  case (step y)
+  obtain l'' D'' where [simp]: "y = (l'', D'')"
+    by force
+  from step.hyps(3)[OF this] step.prems have "
+    l'' = l \<and> D'' = D \<or>
+    check_diag n D'' \<or>
+    (\<exists>M. D'' = FW' (norm_upd M (k' l'') n) n \<and>
+         ((\<forall>i>n. \<forall>j. curry M i j = 0) \<and> (\<forall>j>n. \<forall>i. curry M i j = 0)) \<and> dbm_int (curry M) n)"
+    by auto
+  then show ?case
+  proof (standard, goal_cases)
+    case 1
+    with step_FW_norm[OF _ step.prems(1,2)] step.hyps(2) show ?thesis
+      unfolding E_from_op_def by auto
+  next
+    case 2
+    then show ?case
+    proof (standard, goal_cases)
+      case 1
+      with step.hyps(2) show ?case
+        unfolding E_from_op_def by (auto simp: check_diag)
+    next
+      case 2
+      then obtain M where M:
+        "D'' = FW' (norm_upd M (k' l'') n) n" "dbm_default (curry M) n" "dbm_int (curry M) n"
+        by auto
+      from step.hyps(2) show ?case
+        unfolding E_from_op_def
+        apply simp
+        apply (elim exE conjE)
+        subgoal premises prems for g a r
+        proof (cases "check_diag n (f l'' r g l' (FW' (norm_upd M (k' l'') n) n))")
+          case False
+          with prems show ?thesis
+          using M
+          apply -
+          apply (drule
+            step_FW_norm[
+              of l'' g a r l',
+              rotated,
+              OF FW'_default[OF norm_upd_default]
+                 FW'_int_preservation[OF norm_upd_int_preservation[where k =  "k' l''"]]
+            ]
+            , assumption)
+          defer
+            apply (rule length_k'; fail)
+          apply (simp; fail)
+          unfolding k'_def by auto
+        next
+          case True
+          with \<open>D' = _\<close> \<open>D'' = _\<close> show ?thesis by auto
+        qed
+        done
+    qed
+  qed
+qed
+
+(* Duplication *)
+lemma E_closure_finite:
+  "finite {x. E_from_op** a0 x \<and> \<not> check_diag n (snd x)}"
+proof -
+  have k': "map int (map (k l) [0..<Suc n]) = k' l" for l unfolding k'_def by auto
+  have *:
+    "(l, M) = a0 \<or> check_diag n M
+     \<or> (\<exists>M'. M = FW' (norm_upd M' (k' l) n) n \<and> dbm_default (curry M') n)"
+    if "E_from_op** a0 (l, M)" for l M
+    using that unfolding E_closure a0_def
+    apply -
+    apply (drule steps_FW_norm)
+    unfolding init_dbm_def by (auto simp: neutral)
+  moreover have **: "l \<in> locations" if "E_from_op** a0 (l, M)" for l M
+    using that unfolding E_closure locations_def
+    apply (induction "(l, M)")
+     apply (simp add: a0_def; fail)
+    by (force simp: E_from_op_def)
+  have
+    "{x. E_from_op** a0 x \<and> \<not> check_diag n (snd x)} \<subseteq>
+     {a0} \<union>
+     (locations \<times> {FW' (norm_upd M (k' l) n) n | l M. l \<in> locations \<and> dbm_default (curry M) n})"
+    (is "_ \<subseteq> _ \<union> ?S")
+    apply safe
+     apply (rule **, assumption)
+    apply (frule *)
+    by (auto intro: **)
+  moreover have "finite ?S"
+    using FW'_normalized_integral_dbms_finite' finite_locations
+    using [[simproc add: finite_Collect]]
+    by (auto simp: k'_def finite_project_snd)
+  ultimately show ?thesis by (auto intro: finite_subset)
+qed
+
+end (* End of context for finiteness *)
+
+locale E_From_Op_Bisim_Finite = E_From_Op_Bisim + E_From_Op_Finite
+begin
+
+  lemma E_from_op_check_diag:
+    "check_diag n M'" if "check_diag n M" "E_from_op (l, M) (l', M')"
+    using that unfolding E_from_op_def by (blast intro: check_diag)
+
+  lemma reachable_wf_dbm:
+    "wf_dbm M" if "E_from_op** a0 (l, M)"
+  using that proof (induction "(l, M)" arbitrary: l M)
+    case base
+    then show ?case unfolding a0_def by simp
+  next
+    case (step y l' M')
+    obtain l M where [simp]: "y = (l, M)" by force
+    from E_from_op_wf_state[OF _ step(2)] step(1,3) show ?case
+      unfolding wf_state_def by auto
+  qed
+
+end
+
+locale E_From_Op_Bisim_Finite_Reachability = E_From_Op_Bisim_Finite + Reachability_Problem
+begin
+
+  sublocale Search_Space E_from_op a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M"
+    apply standard
+    subgoal for a
+      apply (rule prod.exhaust[of a])
+      by (auto simp add: subsumes_simp_1 dbm_subset_refl)
+    subgoal for a b c
+      apply (rule prod.exhaust[of a], rule prod.exhaust[of b], rule prod.exhaust[of c])
+      subgoal for l1 M1 l2 M2 l3 M3
+        apply simp
+        apply (cases "check_diag n M1")
+         apply (simp add: subsumes_def; fail)
+        apply (simp add: subsumes_def)
+        by (meson check_diag_subset dbm_subset_def dbm_subset_trans)
+      done
+    subgoal for a b a'
+      apply (rule prod.exhaust[of a], rule prod.exhaust[of b], rule prod.exhaust[of a'])
+      apply safe
+      by (drule E_from_op_mono';
+          fastforce simp: E_def subsumes_def dbm_subset_def Graph_Start_Defs.reachable_def
+          intro!: reachable_wf_dbm)
+    subgoal
+      unfolding F_rel_def subsumes_def by auto
+    subgoal
+      using check_diag_subset unfolding subsumes_def dbm_subset_def by auto
+    subgoal
+      using E_from_op_check_diag by auto
+    unfolding F_rel_def subsumes_def
+    unfolding check_diag_def pointwise_cmp_def
+    by fastforce
+
+  sublocale Search_Space_finite E_from_op a0 F_rel "subsumes n" "\<lambda> (l, M). check_diag n M"
+    by standard
+       (auto intro: finite_subset[OF _ E_closure_finite] simp: Graph_Start_Defs.reachable_def)
+
+  sublocale liveness_pre:
+    Liveness_Search_Space_pre
+    "\<lambda> (l, M) (l', M'). E_from_op (l, M) (l', M') \<and> F l \<and> F l' \<and> \<not> check_diag n M'" a0 "\<lambda> _. False"
+    "subsumes n" "\<lambda> (l, M). \<not> check_diag n M \<and> reachable (l, M)"
+    apply standard
+        apply blast
+       apply (blast intro: trans)
+      apply safe
+     apply (frule mono)
+         apply assumption
+        apply assumption
+       apply assumption
+      apply (simp; fail)
+     apply safe
+     apply (intro exI conjI)
+       prefer 3
+       apply assumption
+      apply safe
+    subgoal a
+      using empty_mono by auto
+    subgoal
+      using reachable_step by blast
+    subgoal
+      by (metis subsumes_simp_2)
+    subgoal
+      by (metis subsumes_simp_2)
+    subgoal
+      by (rule a)
+    subgoal
+      using finite_reachable by (auto intro: finite_subset[rotated])
+    done
+
+end (* End of context for finiteness and bisimilarity *)
+
+context TA_Start
+begin
+
+lemma norm_step_dbm_equiv:
+  "FW' (norm_upd D (k' l') n) n \<simeq> FW' (norm_upd M (k' l') n) n" if "D \<simeq> M" "wf_dbm D" "wf_dbm M"
+  unfolding dbm_equiv_def
+  by (subst norm_step_correct'[OF that(2,3,1)], subst norm_step_correct'[OF that(3,3)]) auto
+
+lemma empty_dbm_marker:
+  "[curry (conv_M (M((0, 0) := Lt 0)))]v,n = {}"
+proof -
+  have "check_diag n (M((0, 0) := Lt 0))"
+    by (simp add: check_diag_marker)
+  then show ?thesis
+    using canonical_check_diag_empty_iff by blast
+qed
+
+lemma wf_dbm_marker:
+  "wf_dbm (M((0, 0) := Lt 0))" (is "wf_dbm ?M") if "wf_dbm M"
+proof -
+  have "check_diag n ?M"
+    by (simp add: check_diag_marker)
+  moreover from this have "[curry (conv_M ?M)]v,n \<subseteq> V"
+    by (simp add: empty_dbm_marker)
+  moreover have "\<forall>i\<le>n. (M((0, 0) := Lt 0)) (i, i) \<le> 0"
+    using that by (auto simp: Lt_lt_LeI neutral dest: wf_dbm_altD)
+  ultimately show ?thesis
+    by (intro wf_dbm_altI) (simp add: check_diag_marker; fail)+
+qed
+
+lemma filter_diag_wf_dbm:
+  assumes "\<And> M. wf_dbm M \<Longrightarrow> wf_dbm (f M)" "wf_dbm M"
+  shows "wf_dbm (filter_diag f M)"
+  unfolding filter_diag_def using assms by (auto intro: wf_dbm_marker)
+
+lemma
+  assumes "A \<turnstile> l \<longrightarrow>g,a,r l'" "wf_dbm M"
+    shows E_op'_wf: "wf_dbm (E_op' l r g l' M)" and E_op''_wf: "wf_dbm (E_op'' l r g l' M)"
+proof -
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). 0 < c \<and> c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set (inv_of A l'). 0 < c \<and> c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set g. 0 < c \<and> c \<le> n"
+    using clock_range collect_clocks_clk_set[OF assms(1)] unfolding collect_clks_def by blast
+  moreover have "\<forall>i\<in>set r. 0 < i \<and> i \<le> n"
+    using clock_range reset_clk_set[OF assms(1)] unfolding collect_clks_def by blast
+  moreover note side_conds = calculation assms(2)
+  note wf_intros =
+    norm_step_wf_dbm wf_dbm_abstr_repair wf_dbm_reset'_upd wf_dbm_up_canonical_upd
+    filter_diag_wf_dbm
+  show "wf_dbm (E_op' l r g l' M)"
+    unfolding E_op'_def by simp (intro wf_intros side_conds order.refl)
+  show "wf_dbm (E_op'' l r g l' M)"
+    unfolding E_op''_def by simp (intro wf_intros side_conds order.refl)
+qed
+
+lemma wf_dbm_abstr_repair_equiv_FW:
+  assumes "\<forall> c \<in> constraint_clk ` set cc. c > 0 \<and> c \<le> n" "M \<simeq> M'"
+  shows "abstr_repair cc M \<simeq> FW' (abstr_upd cc M') n"
+  using wf_dbm_abstr_repair_equiv[OF assms] by (simp add: dbm_equiv_def FW'_zone_equiv)
+
+lemma E_op'_bisim:
+  "E_op' l r g l' M \<simeq> E_op l r g l' M" if "A \<turnstile> l \<longrightarrow>g,a,r l'" "wf_dbm M"
+proof -
+  note intros =
+    norm_step_dbm_equiv wf_dbm_abstr_repair_equiv_FW[rotated] reset'_upd_equiv dbm_equiv_refl
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). 0 < c \<and> c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set (inv_of A l'). 0 < c \<and> c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set g. 0 < c \<and> c \<le> n"
+    using clock_range collect_clocks_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover have "\<forall>i\<in>set r. 0 < i \<and> i \<le> n"
+    using clock_range reset_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover note side_conds = calculation that(2)
+  note wf_intros =
+    norm_step_wf_dbm wf_dbm_abstr_repair wf_dbm_reset'_upd wf_dbm_up_canonical_upd
+    wf_dbm_FW'_abstr_upd
+  show ?thesis unfolding E_op'_def E_op_def by simp (intro intros wf_intros side_conds order.refl)
+qed
+
+lemma filter_diag_equiv:
+  assumes "\<And> M. check_diag n M \<Longrightarrow> check_diag n (f M)"
+  shows "filter_diag f M \<simeq> f M"
+proof (cases "check_diag n M")
+  case True
+  then have "check_diag n (f M)" by (rule assms)
+  with True show ?thesis
+    unfolding filter_diag_def dbm_equiv_def
+    by (auto dest: check_diag_conv_M check_diag_empty_spec simp: empty_dbm_marker)
+next
+  case False
+  then show ?thesis by (auto simp: filter_diag_def)
+qed
+
+lemma E_op''_bisim:
+  "E_op'' l r g l' M \<simeq> E_op' l r g l' M" if "A \<turnstile> l \<longrightarrow>g,a,r l'" "wf_dbm M"
+proof -
+  note intros =
+    norm_step_dbm_equiv dbm_equiv_refl dbm_equiv_trans[OF filter_diag_equiv, rotated]
+    wf_dbm_abstr_repair_equiv' reset'_upd_equiv
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). 0 < c \<and> c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have
+    "\<forall>c\<in>constraint_clk ` set (inv_of A l'). 0 < c \<and> c \<le> n"
+    "\<forall>c\<in>constraint_clk ` set (inv_of A l'). 0 < c"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast+
+  moreover have "\<forall>c\<in>constraint_clk ` set g. 0 < c \<and> c \<le> n" "\<forall>c\<in>constraint_clk ` set g. 0 < c"
+    using clock_range collect_clocks_clk_set[OF that(1)] unfolding collect_clks_def by blast+
+  moreover have "\<forall>i\<in>set r. 0 < i \<and> i \<le> n" "\<forall>i\<in>set r. 0 < i"
+    using clock_range reset_clk_set[OF that(1)] unfolding collect_clks_def by blast+
+  moreover note side_conds = calculation that(2)
+  note wf_intros =
+    wf_dbm_abstr_repair wf_dbm_reset'_upd wf_dbm_up_canonical_upd filter_diag_wf_dbm
+  note check_diag_intros =
+    norm_step_check_diag_preservation reset'_upd_check_diag_preservation
+    abstr_repair_check_diag_preservation norm_step_check_diag_preservation
+  show ?thesis unfolding E_op''_def E_op'_def
+    by simp (intro intros check_diag_intros side_conds wf_intros order.refl)
+qed
+
+lemma E_op''_bisim':
+  "E_op'' l r g l' M \<simeq> E_op l r g l' M" if "A \<turnstile> l \<longrightarrow>g,a,r l'" "wf_dbm M"
+  using that by (blast intro: E_op''_bisim E_op'_bisim)
+
+lemma abstra_upd_default:
+  assumes "dbm_default (curry M) n" "constraint_clk ac \<le> n"
+  shows "dbm_default (curry (abstra_upd ac M)) n"
+  using assms by (auto simp: abstra_upd_out_of_bounds1 abstra_upd_out_of_bounds2)
+
+lemma FWI_default:
+  assumes "dbm_default (curry M) n"
+  shows "dbm_default (curry (FWI' M n i)) n"
+  using assms unfolding FWI'_def FWI_def by (simp add: fwi_out_of_bounds1 fwi_out_of_bounds2)
+
+lemma abstra_repair_default:
+  assumes "dbm_default (curry M) n" "constraint_clk ac \<le> n"
+  shows "dbm_default (curry (abstra_repair ac M)) n"
+  using assms unfolding abstra_repair_def repair_pair_def by (intro abstra_upd_default FWI_default)
+
+lemma abstr_repair_default:
+  assumes "dbm_default (curry M) n" "\<forall>c\<in>collect_clks cc. c \<le> n"
+  shows "dbm_default (curry (abstr_repair cc M)) n"
+  using assms
+  unfolding abstr_repair_def
+  apply (induction cc arbitrary: M)
+   apply (simp; fail)
+  by (drule abstra_repair_default; auto)
+
+lemma abstra_repair_int:
+  assumes "dbm_int (curry M) n" "constraint_clk ac \<le> n" "snd (constraint_pair ac) \<in> \<int>"
+  shows "dbm_int (curry (abstra_repair ac M)) n"
+  using assms unfolding abstra_repair_def repair_pair_def FWI'_def FWI_def
+  by simp (intro abstra_upd_int_preservation fwi_int_preservation; (assumption | simp))
+
+(* Unused *)
+lemma abstr_repair_int:
+  assumes
+    "dbm_int (curry M) n" "\<forall>c\<in>collect_clks cc. c \<le> n" "\<forall> (_,d) \<in> collect_clock_pairs cc. d \<in> \<int>"
+  shows "dbm_int (curry (abstr_repair cc M)) n"
+  using assms
+  unfolding abstr_repair_def
+  apply (induction cc arbitrary: M)
+   apply (simp; fail)
+  by (drule abstra_repair_int; auto simp: collect_clock_pairs_def)
+
+lemma E_op'_FW_norm:
+  "\<exists> M'.
+    E_op' l r g l' M = FW' (norm_upd M' (k' l') n) n
+    \<and> dbm_default (curry M') n \<and> dbm_int (curry M') n"
+  if "A \<turnstile> l \<longrightarrow>g,a,r l'" "dbm_default (curry M) n"
+proof -
+  note default_intros = up_canonical_upd_default reset'_upd_default abstr_repair_default
+  (* Duplication *)
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have
+    "\<forall>c\<in>constraint_clk ` set (inv_of A l'). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set g. c \<le> n"
+    using clock_range collect_clocks_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover have "\<forall>i\<in>set r. i \<le> n"
+    using clock_range reset_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover note side_conds = calculation that(2)
+  let ?M = "
+    curry (
+       abstr_repair (inv_of A l')
+         (reset'_upd (abstr_repair g (abstr_repair (inv_of A l) (up_canonical_upd M n))) n r 0))"
+  have "dbm_default ?M n"
+    by (intro default_intros[unfolded collect_clks_def] side_conds)
+  moreover have "dbm_int ?M n"
+    by (auto simp: Ints_def)
+  ultimately show ?thesis unfolding E_op'_def by auto
+qed
+
+lemma filter_diag_default:
+  assumes "\<And> M. dbm_default (curry M) n \<Longrightarrow> dbm_default (curry (f M)) n" "dbm_default (curry M) n"
+  shows "dbm_default (curry (filter_diag f M)) n"
+  unfolding filter_diag_def using assms by auto
+
+lemma E_op''_FW_norm:
+  "\<exists> M'.
+    E_op'' l r g l' M = FW' (norm_upd M' (k' l') n) n
+    \<and> dbm_default (curry M') n \<and> dbm_int (curry M') n"
+  if "A \<turnstile> l \<longrightarrow>g,a,r l'" "dbm_default (curry M) n" "\<not> check_diag n (E_op'' l r g l' M)"
+proof -
+  note default_intros =
+    filter_diag_default up_canonical_upd_default reset'_upd_default abstr_repair_default
+  (* Duplication *)
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have
+    "\<forall>c\<in>constraint_clk ` set (inv_of A l'). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set g. c \<le> n"
+    using clock_range collect_clocks_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover have "\<forall>i\<in>set r. i \<le> n"
+    using clock_range reset_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover note side_conds = calculation that(2)
+  let ?M = "
+    filter_diag (\<lambda>M. abstr_repair (inv_of A l') (reset'_upd M n r 0))
+      (filter_diag (abstr_repair g) (abstr_repair (inv_of A l) (up_canonical_upd M n)))"
+  have "dbm_default (curry ?M) n"
+    by (intro default_intros[unfolded collect_clks_def] side_conds)
+  moreover have "dbm_int (curry ?M) n"
+    by (auto simp: Ints_def)
+  ultimately show ?thesis
+    using that(3) unfolding E_op''_def
+    by (auto simp: Let_def filter_diag_def intro: check_diag_marker)
+qed
+
+lemma E_op''_check_diag_aux:
+  "check_diag n (abstr_repair (inv_of A l) (up_canonical_upd M n))" if "check_diag n M"
+  apply (intro abstr_repair_check_diag_preservation check_diag_up_canonical_upd)
+  using that collect_clks_inv_clk_set[of A l] clock_range unfolding collect_clks_def by blast+
+
+lemma E_op''_check_diag:
+  "check_diag n (E_op'' l r g l' M)" if "check_diag n M"
+  using that E_op''_check_diag_aux unfolding E_op''_def filter_diag_def
+  by (auto simp: Let_def intro: check_diag_marker)
+
+sublocale E_op'': E_From_Op_Bisim_Finite _ _ _ _ E_op''
+  by standard (rule E_op''_bisim' E_op''_wf E_op''_FW_norm E_op''_check_diag; assumption)+
+
+end (* TA Start *)
+
+context Reachability_Problem
+begin
+
+sublocale E_op'': E_From_Op_Bisim_Finite_Reachability _ _ _ _ E_op'' ..
+
+end
+
+context TA_Start_Defs
+begin
+
+abbreviation step_impl' ("\<langle>_, _\<rangle> \<leadsto>_ \<langle>_, _\<rangle>" [61,61,61] 61)
+where
+  "\<langle>l, Z\<rangle> \<leadsto>a \<langle>l'', Z'''\<rangle> \<equiv> \<exists> l' Z' Z''.
+    A \<turnstile>I \<langle>l, Z\<rangle> \<leadsto>n,\<tau> \<langle>l', Z'\<rangle>
+    \<and> A \<turnstile>I \<langle>l', Z'\<rangle> \<leadsto>n,\<upharpoonleft>a \<langle>l'', Z''\<rangle>
+    \<and> FW' (norm_upd Z'' (k' l'') n) n = Z'''"
+
+sublocale Graph_Defs "\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>" .
+
+end (* Reachability Problem Defs *)
+
+
+paragraph \<open>Misc\<close>
+
+lemma finite_conv_A:
+  "finite (trans_of (conv_A A))" if "finite (trans_of A)"
+  using that unfolding trans_of_def conv_A_def by (cases A) auto
+
+lemma real_of_int_inj:
+  "inj real_of_int"
+  by standard auto
+
+lemma map_DBMEntry_real_of_int_inj:
+  "a = b" if "map_DBMEntry real_of_int a = map_DBMEntry real_of_int b"
+proof -
+  have "inj (map_DBMEntry real_of_int)"
+    by (intro DBMEntry.inj_map real_of_int_inj)
+  with that show ?thesis
+    by - (erule injD)
+qed
+
+lemma (in -) n_eq_conv_MI:
+  "curry D =n (curry D')" if "curry (conv_M D) =n curry (conv_M D')"
+  using that unfolding n_eq_def by (auto intro: map_DBMEntry_real_of_int_inj)
+
+lemma (in TA_Start_No_Ceiling) check_diag_conv_M_iff:
+  "check_diag n D \<longleftrightarrow> check_diag n (conv_M D)"
+  using check_diag_conv_M check_diag_conv_M_rev by fast
+
+lemma (in Regions) \<R>_regions_distinct:
+  "R = R'" if "u \<in> R" "u \<in> R'" "R \<in> \<R> l" "R' \<in> \<R> l"
+  using that unfolding \<R>_def by clarsimp (rule valid_regions_distinct; auto)
+
+lemma step_impl_delay_loc_eq:
+  "l' = l" if "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l', D'\<rangle>"
+  using that by cases auto
+
+
+context TA_Start_No_Ceiling
+begin
+
+paragraph \<open>@{term init_dbm}\<close>
+
+lemma init_dbm_semantics:
+  "u \<in> [(curry init_dbm :: real DBM)]v,n \<longleftrightarrow> (\<forall>c\<in>{1..n}. u c = 0)"
+  by (safe elim!: init_dbm_semantics' init_dbm_semantics'')
+
+lemma init_dbm_non_empty:
+  "[(curry init_dbm :: real DBM)]v,n \<noteq> {}"
+proof -
+  let ?u = "\<lambda> c. 0 :: real"
+  have "?u \<in> [curry init_dbm]v,n"
+    by (rule init_dbm_semantics'', auto)
+  then show ?thesis by auto
+qed
+
+end
+
+context TA_Start
+begin
+
+sublocale TA: Regions_TA v n "Suc n" "{1..<Suc n}" k "conv_A A"
+  by (standard; intro valid_abstraction' finite_conv_A finite_trans_of_finite_state_set finite_trans)
+
+paragraph \<open>@{term dbm_default}\<close>
+
+lemma dbm_default_n_eq_eqI:
+  assumes "dbm_default (curry D) n" "dbm_default (curry D') n" "curry D' =n curry D"
+  shows "D' = D"
+  using assms
+    apply (intro ext)
+      apply safe
+      subgoal for a b
+        by (cases "a \<le> n"; cases "b \<le> n"; simp add: n_eq_def)
+      done
+
+lemma E_op''_dbm_default':
+  "dbm_default (curry (E_op'' l r g l' M)) n"
+  if "A \<turnstile> l \<longrightarrow>g,a,r l'" "dbm_default (curry M) n"
+proof -
+  note default_intros =
+    filter_diag_default up_canonical_upd_default reset'_upd_default abstr_repair_default
+  (* Duplication *)
+  have "\<forall>c\<in>constraint_clk ` set (inv_of A l). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l] unfolding collect_clks_def by blast
+  moreover have
+    "\<forall>c\<in>constraint_clk ` set (inv_of A l'). c \<le> n"
+    using clock_range collect_clks_inv_clk_set[of A l'] unfolding collect_clks_def by blast
+  moreover have "\<forall>c\<in>constraint_clk ` set g. c \<le> n"
+    using clock_range collect_clocks_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover have "\<forall>i\<in>set r. i \<le> n"
+    using clock_range reset_clk_set[OF that(1)] unfolding collect_clks_def by blast
+  moreover note side_conds = calculation that(2)
+  let ?M = "
+    filter_diag (\<lambda>M. abstr_repair (inv_of A l') (reset'_upd M n r 0))
+      (filter_diag (abstr_repair g) (abstr_repair (inv_of A l) (up_canonical_upd M n)))"
+  have "dbm_default (curry ?M) n"
+    by (intro default_intros[unfolded collect_clks_def] side_conds)
+  then have "dbm_default (curry (filter_diag (\<lambda>M. FW' (norm_upd M (k' l') n) n) ?M)) n"
+    by - (rule default_intros, intro FW'_default norm_upd_default)
+  then show ?thesis
+    unfolding E_op''_def by (auto simp: Let_def filter_diag_def)
+qed
+
+lemma E_op''_dbm_default:
+  assumes
+    "E_op''.E_from_op (l, D) (l', D')"
+    "dbm_default (curry D) n"
+  shows
+    "dbm_default (curry D') n"
+  using assms E_op''_dbm_default' unfolding E_op''.E_from_op_def by auto
+
+
+paragraph \<open>@{term check_diag}\<close>
+
+lemma canonical_empty_check_diag':
+  assumes "wf_dbm D" "[curry (conv_M D)]v,n = {}"
+  shows "check_diag n D"
+  apply (rule canonical_empty_check_diag[OF _ assms(2)])
+  using wf_dbm_D(1)[OF assms(1)] unfolding check_diag_def neutral by auto
+
+paragraph \<open>Stronger invariant on DBMs\<close>
+
+definition
+  "dbm_inv D \<equiv>
+    canonical' (conv_M D) \<and> \<not> check_diag n D \<and> (\<forall>i\<le>n. conv_M D (i, i) \<le> 0)
+    \<and> valid_dbm (curry (conv_M D))
+    \<and> dbm_default (curry D) n"
+
+lemma dbm_inv_equvi_eqI:
+  assumes "dbm_inv D" "dbm_inv D'" "[curry (conv_M D')]v,n = [curry (conv_M D)]v,n"
+  shows "D = D'"
+proof -
+  from assms have "wf_dbm D" "\<not> check_diag n D"
+    unfolding dbm_inv_def wf_dbm_def by auto
+  then have *: "[curry (conv_M D)]v,n \<noteq> {}"
+    by (auto dest!: canonical_empty_check_diag')
+  from assms have
+    "curry (conv_M D) =n curry (conv_M D')"
+    apply -
+    apply (rule canonical_eq_upto[OF clock_numbering(1) surj_numbering _ _ * assms(3)[symmetric]])
+    unfolding dbm_inv_def
+       apply (simp; fail)+
+    subgoal
+      unfolding check_diag_conv_M_iff
+      unfolding check_diag_def neutral[symmetric] by fastforce
+    subgoal
+      unfolding check_diag_conv_M_iff
+      unfolding check_diag_def neutral[symmetric] by fastforce
+    done
+  then have "curry D =n curry D'"
+    by (rule n_eq_conv_MI)
+  moreover from assms have "dbm_default (curry D) n" "dbm_default (curry D') n"
+    unfolding dbm_inv_def by simp+
+  ultimately show ?thesis
+    by - (rule dbm_default_n_eq_eqI)
+qed
+
+lemma dbm_invI:
+  "dbm_inv D" if "wf_dbm D" "\<not> check_diag n D" "dbm_default (curry D) n"
+  using that unfolding dbm_inv_def wf_dbm_def by auto
+
+lemma init_dbm_dbm_inv[intro, simp]:
+  "dbm_inv init_dbm"
+proof -
+  have "\<not> check_diag n init_dbm"
+    unfolding init_dbm_def check_diag_def by auto
+  moreover have "dbm_default (curry init_dbm) n"
+    unfolding init_dbm_def neutral by auto
+  ultimately show ?thesis
+    using wf_dbm_init_dbm by - (rule dbm_invI)
+qed
+
+
+subsubsection \<open>Bridging the semantic gap\<close>
+
+lemma step_impl'_E:
+  "(\<lambda> (l, D) (l', D'). \<exists> a. \<langle>l, D\<rangle> \<leadsto>a \<langle>l', D'\<rangle>) = E"
+  unfolding E_def by (intro ext) auto
+
+lemma step_z_norm''_step_impl'_equiv:
+  "Bisimulation_Invariant
+     (\<lambda> (l, D) (l', D'). \<exists> a. step_z_norm'' (conv_A A) l D a l' D')
+     (\<lambda>(l, D) (l', D'). \<exists>a. \<langle>l, D\<rangle> \<leadsto>a \<langle>l', D'\<rangle>)
+     (\<lambda>(l, M) (l', D). l = l' \<and> [curry (conv_M D)]v,n = [M]v,n)
+     (\<lambda>(l, y). valid_dbm y)
+     wf_state"
+proof (standard, goal_cases)
+  case prems: (1 a b a')
+  obtain l M l' M' l1 D where unfolds[simp]: "a = (l, M)" "b = (l', M')" "a' = (l1, D)"
+    by force+
+  from prems have [simp]: "l1 = l"
+    by auto
+  from prems obtain a1 where
+    "step_z_norm'' (conv_A A) l M a1 l' M'"
+    by auto
+  then obtain l2 M1 where steps:
+    "conv_A A \<turnstile> \<langle>l, M\<rangle> \<leadsto>v,n,\<tau> \<langle>l2, M1\<rangle>"
+    "step_z_norm' (conv_A A) l2 M1 \<upharpoonleft>a1 l' M'"
+    unfolding step_z_norm''_def by auto
+  from step_z_dbm_equiv'[OF steps(1), of "curry (conv_M D)"] prems(2-) obtain M2 where
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l2, M2\<rangle>" "wf_dbm D" "[M1]v,n = [M2]v,n"
+    by (auto simp: wf_state_def)
+  with step_impl_complete''_improved[OF this(1)] obtain D2 where D2:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l2, D2\<rangle>" "[curry (conv_M D2)]v,n = [M1]v,n"
+    by auto
+  from step_impl_wf_dbm[OF D2(1) \<open>wf_dbm D\<close>] have "wf_dbm D2" .
+  then have *:
+    "canonical' (conv_M D2) \<or> check_diag n D2"
+    "\<forall>i\<le>n. conv_M D2 (i, i) \<le> 0"
+    "valid_dbm (curry (conv_M D2))"
+    using wf_dbm_D by blast+
+  have "valid_dbm M1"
+    using prems steps(1) by - (rule step_z_valid_dbm', auto)
+  from step_z_norm_equiv'[OF steps(2), OF this *(3) sym[OF D2(2)]] obtain D3 where D3:
+    "step_z_norm' (conv_A A) l2 (curry (conv_M D2)) \<upharpoonleft>a1 l' D3"
+    "[M']v,n = [D3]v,n"
+    by atomize_elim
+  from step_impl_norm_complete''[OF D3(1) *(3,1,2)] obtain D4 where D4:
+    "A \<turnstile>I \<langle>l2, D2\<rangle> \<leadsto>n,\<upharpoonleft>a1 \<langle>l', D4\<rangle>"
+    "[curry (conv_M (FW' (norm_upd D4 (k' l') n) n))]v,n = [D3]v,n"
+    by auto
+  with D2(1) have "\<langle>l, D\<rangle> \<leadsto>a1 \<langle>l', FW' (norm_upd D4 (k' l') n) n\<rangle>"
+    by auto
+  with D4(2) D3(2) show ?case
+    by force
+next
+  case prems: (2 a a' b')
+  obtain l M l' D' l1 D where unfolds[simp]: "a = (l, M)" "b' = (l', D')" "a' = (l1, D)"
+    by force+
+  from prems have [simp]: "l1 = l"
+    by auto
+  from prems obtain a1 where "\<langle>l, D\<rangle> \<leadsto>a1 \<langle>l', D'\<rangle>"
+    by auto
+  then obtain D1 D2 where steps:
+    "A \<turnstile>I \<langle>l, D\<rangle> \<leadsto>n,\<tau> \<langle>l, D1\<rangle>" "A \<turnstile>I \<langle>l, D1\<rangle> \<leadsto>n,\<upharpoonleft>a1 \<langle>l', D2\<rangle>"
+    "D' = FW' (norm_upd D2 (k' l') n) n"
+    by (auto dest: step_impl_delay_loc_eq)
+  from prems have "wf_dbm D"
+    by (auto simp: wf_state_def)
+  with steps have "wf_dbm D1"
+    by (blast intro: step_impl_wf_dbm)
+  from \<open>wf_dbm D\<close> have
+    "canonical' (conv_M D) \<or> check_diag n D"
+    "\<forall>i\<le>n. conv_M D (i, i) \<le> 0"
+    "valid_dbm (curry (conv_M D))"
+    using wf_dbm_D by blast+
+  from step_impl_sound'[OF steps(1) this] obtain M2 where M2:
+    "conv_A A \<turnstile> \<langle>l, curry (conv_M D)\<rangle> \<leadsto>v,n,\<tau> \<langle>l, M2\<rangle>"
+    "[curry (conv_M D1)]v,n = [M2]v,n"
+    by auto
+  from step_impl_sound_wf[OF steps(2) \<open>wf_dbm D1\<close>] obtain M3 where M3:
+    "step_z_norm' (conv_A A) l (curry (conv_M D1)) \<upharpoonleft>a1 l' M3"
+    "[curry (conv_M (FW' (norm_upd D2 (k' l') n) n))]v,n = [M3]v,n"
+    by auto
+  from step_z_dbm_equiv'[OF M2(1), of M] prems(2) obtain M2' where M2':
+    "conv_A A \<turnstile> \<langle>l, M\<rangle> \<leadsto>v,n,\<tau> \<langle>l, M2'\<rangle>" "[M2]v,n = [M2']v,n"
+    by auto
+  have "valid_dbm (curry (conv_M D1))" "valid_dbm M2'"
+    subgoal
+      using \<open>wf_dbm D1\<close> wf_dbm_D(3) by auto
+    subgoal
+      using prems M2'(1) by - (rule step_z_valid_dbm', auto)
+    done
+  from step_z_norm_equiv'[OF M3(1), of M2', OF this] M2(2) M2'(2) obtain M3' where
+    "step_z_norm' (conv_A A) l M2' (\<upharpoonleft>a1) l' M3'" "[M3]v,n = [M3']v,n"
+    by auto
+  with M2'(1) M3(2) steps(3) show ?case
+    unfolding step_z_norm''_def by auto
+next
+  case (3 a b)
+  then show ?case
+    unfolding step_z_norm''_def using step_z_norm_valid_dbm'_spec step_z_valid_dbm' by auto
+next
+  case (4 a b)
+  then show ?case
+    by (clarsimp simp: norm_step_wf_dbm step_impl_wf_dbm wf_state_def)
+qed
+
+context
+  assumes "l0 \<in> state_set A"
+begin
+
+lemma l0_state_set:
+  "l0 \<in> state_set (conv_A A)"
+  using \<open>l0 \<in> state_set A\<close> unfolding state_set_def trans_of_def conv_A_def by (cases A; force)
+
+interpretation start:
+  Regions_TA_Start_State v n "Suc n" "{1..<Suc n}" k "conv_A A"
+  l0 "[(curry init_dbm :: real DBM)]v,n"
+  apply standard
+  subgoal
+    by (rule l0_state_set)
+  subgoal
+    using valid_dbm_V' by blast
+  subgoal
+    by (rule init_dbm_non_empty)
+  done
+
+lemma norm_final_bisim:
+  "Bisimulation_Invariant
+     (\<lambda>(l, D) (l', D'). \<exists>a. step_z_norm'' (conv_A A) l D a l' D')
+     E_op''.E_from_op
+     (\<lambda> (l, M) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = [M]v,n)
+     (\<lambda>(l, y). valid_dbm y) wf_state"
+  by (rule
+    Bisimulation_Invariant_sim_replace[OF
+      Bisimulation_Invariant_composition[OF
+        step_z_norm''_step_impl'_equiv[unfolded step_impl'_E] E_op''.E_from_op_bisim
+      ]
+    ])
+    (auto simp add: state_equiv_def dbm_equiv_def)
+
+lemma norm_final_bisim_empty:
+  "Bisimulation_Invariant
+     (\<lambda>(l, D) (l', D'). \<exists>a. step_z_norm'' (conv_A A) l D a l' D' \<and> [D']v,n \<noteq> {})
+     (\<lambda> a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))
+     (\<lambda> (l, M) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = [M]v,n)
+     (\<lambda>(l, y). valid_dbm y) wf_state"
+  by (rule
+    Bisimulation_Invariant_filter[OF
+      norm_final_bisim, of "\<lambda> (l, D). [D]v,n \<noteq> {}" "\<lambda> a. \<not> check_diag n (snd a)"]
+    )
+    (auto
+      intro!: canonical_empty_check_diag' simp: wf_state_def
+      dest: check_diag_empty_spec[OF check_diag_conv_M]
+    )
+
+lemma beta_final_bisim_empty:
+  "Bisimulation_Invariant
+     (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {})
+     (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))
+     (\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z)
+     (\<lambda>_. True) wf_state"
+  by (rule
+    Bisimulation_Invariant_sim_replace[OF
+      Bisimulation_Invariant_composition[OF
+        bisim[OF global_clock_numbering' valid_abstraction'] norm_final_bisim_empty
+      ]
+    ]
+    )
+    (auto dest!: wf_dbm_D(3) simp: wf_state_def)
+
+lemma beta_final_bisim_empty_strong:
+  "Bisimulation_Invariant
+     (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {})
+     (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))
+     (\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z)
+     (\<lambda>_. True) (\<lambda> (l, D). dbm_inv D)"
+  apply (rule Bisimulation_Invariant_strengthen_post'[OF beta_final_bisim_empty])
+  subgoal for a b
+    unfolding dbm_inv_def wf_state_def wf_dbm_def by (fastforce dest!: E_op''_dbm_default)
+  subgoal for a
+    unfolding wf_state_def dbm_inv_def wf_dbm_def by auto
+  done
+
+interpretation bisim:
+  Bisimulation_Invariant
+    "\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {}"
+    "\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b)"
+    "\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z"
+    "\<lambda>_. True" "\<lambda> (l, D). dbm_inv D"
+  by (rule beta_final_bisim_empty_strong)
+
+context
+  fixes P Q :: "'s \<Rightarrow> bool" \<comment> \<open>The state property we want to check\<close>
+begin
+
+lemma beta_final_bisim_empty_Q:
+  "Bisimulation_Invariant
+     (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')
+     (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))
+     (\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z)
+     (\<lambda>_. True) (\<lambda> (l, D). dbm_inv D)"
+  by (rule
+      Bisimulation_Invariant_filter[
+        OF beta_final_bisim_empty_strong, of "\<lambda> (l, _). Q l" "\<lambda> (l, _). Q l"
+      ]
+    ) auto
+
+interpretation bisim_Q:
+  Bisimulation_Invariant
+    "\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l'"
+    "\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b)"
+    "\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z"
+    "\<lambda>_. True" "(\<lambda> (l, D). dbm_inv D)"
+  by (rule beta_final_bisim_empty_Q)
+
+interpretation bisims_Q:
+  Bisimulation_Invariants
+    "\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l'"
+    "\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b)"
+    "\<lambda>(l, Z) (l', D'). l' = l \<and> [curry (conv_M D')]v,n = Z"
+    "\<lambda>_. True" "\<lambda>_. True" "\<lambda> (l, D). dbm_inv D" "\<lambda> (l, D). dbm_inv D"
+  by (intro Bisimulation_Invariant_Bisimulation_Invariants beta_final_bisim_empty_Q)
+
+lemma leadsto_sem_equiv:
+  "(\<forall>x0\<in>start.a0. leadsto (start.\<phi> P) ((Not \<circ>\<circ> start.\<psi>) Q) x0)
+  = (\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> leadsto (\<lambda> (l, u). P l) (\<lambda> (l, u). \<not> Q l) (l0, u0))
+  "
+proof -
+  have unfold: "(\<lambda>(l, u). P l) = (\<lambda> x. P (fst x))" "(\<lambda>(l, u). \<not> Q l) = (\<lambda>x. \<not> Q (fst x))"
+    by auto
+  show ?thesis
+    unfolding
+      start.a0_def from_R_def init_dbm_semantics start.\<phi>_def start.\<psi>_def comp_def unfold
+    by auto
+qed
+
+lemma Alw_ev_sem_equiv:
+  "(\<forall>x0\<in>start.a0. Alw_ev ((Not \<circ>\<circ> start.\<phi>) Q) x0)
+  = (\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> Alw_ev (\<lambda> (l, u). \<not> Q l) (l0, u0))"
+proof -
+  have unfold: "(\<lambda>(l, u). \<not> Q l) = (\<lambda>x. \<not> Q (fst x))"
+    by auto
+  show ?thesis
+    unfolding start.a0_def from_R_def init_dbm_semantics start.\<phi>_def unfold
+    unfolding comp_def by auto
+qed
+
+lemma leadsto_mc2:
+  "(\<exists>x.
+    TA.reaches (l0, [curry init_dbm]v,n) x \<and>
+    P (fst x) \<and>
+    Q (fst x) \<and>
+    (\<exists>a. (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')** x a \<and>
+         (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')++ a a))
+   \<longleftrightarrow>
+   (\<exists>x.
+    (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** (l0, init_dbm) x \<and>
+    P (fst x) \<and>
+    Q (fst x) \<and>
+    (\<exists>a. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))** x a \<and>
+         (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))++ a a))
+  "
+  apply safe
+   apply (drule bisim.A_B.simulation_reaches[where b = "(l0, init_dbm)"], (simp; fail)+)
+   apply clarify
+   apply (drule bisim_Q.A_B.simulation_reaches)
+      apply blast
+     apply blast
+    apply blast
+   apply clarify
+   apply (frule bisims_Q.A_B.reaches1_unique[rotated], blast+)
+    apply (auto dest: dbm_inv_equvi_eqI; fail)
+   apply force
+  apply (drule bisim.B_A.simulation_reaches[where b = "(l0, [curry init_dbm]v,n)"], (simp; fail)+)
+  apply (drule bisim_Q.B_A.simulation_reaches)
+     apply blast
+    apply blast
+   apply blast
+  apply (drule bisims_Q.B_A.reaches1_unique[rotated]; force)
+  done
+
+lemma Alw_ev_mc2:
+  "Q l0 \<and>
+  (\<exists>a. (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')**
+    (l0, [curry init_dbm]v,n) a \<and>
+  (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')++ a a)
+   \<longleftrightarrow>
+   Q l0 \<and>
+   (\<exists>a. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))** (l0, init_dbm) a \<and>
+        (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))++ a a)
+  "
+  apply safe
+  subgoal for a b
+    apply (drule bisim_Q.A_B.simulation_reaches[where b = "(l0, init_dbm)"])
+       apply force
+      apply blast
+     apply blast
+    apply clarify
+    apply (frule bisims_Q.A_B.reaches1_unique[rotated], blast+)
+     apply (auto dest: dbm_inv_equvi_eqI; fail)
+    apply force
+    done
+  subgoal for a b
+    apply (drule bisim_Q.B_A.simulation_reaches[where b = "(l0, [curry init_dbm]v,n)"])
+       apply (simp; fail)
+      apply blast
+     apply blast
+    apply (drule bisims_Q.B_A.reaches1_unique[rotated], auto)
+    done
+  done
+
+lemmas Alw_ev_mc = Alw_ev_sem_equiv[symmetric,
+    unfolded start.Alw_ev_mc1[of Q, unfolded Graph_Start_Defs.reachable_def],
+    unfolded Alw_ev_mc2
+    ]
+
+context
+  assumes no_deadlock: "\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0)"
+begin
+
+lemma no_deadlock':
+  "\<forall>x0\<in>start.a0. \<not> deadlock x0"
+  unfolding start.a0_def from_R_def using no_deadlock by (auto dest: init_dbm_semantics')
+
+lemma leadsto_mc1:
+  "(\<forall>u0. (\<forall>c \<in> {1..n}. u0 c = 0) \<longrightarrow> leadsto (\<lambda> (l, u). P l) (\<lambda> (l, u). \<not> Q l) (l0, u0)) =
+   (\<nexists>x. TA.reaches (l0, [curry init_dbm]v,n) x \<and>
+       P (fst x) \<and>
+       Q (fst x) \<and>
+       (\<exists>a. (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')** x a \<and>
+            (\<lambda>(l, Z) (l', Z'). step_z_beta' (conv_A A) l Z l' Z' \<and> Z' \<noteq> {} \<and> Q l')++ a a))"
+  unfolding leadsto_sem_equiv[symmetric] by (rule start.leadsto_mc1[OF no_deadlock'])
+
+lemmas leadsto_mc = leadsto_mc1[unfolded leadsto_mc2]
+
+end (* No deadlock *)
+
+end (* State properties *)
+  
+end (* Start State *)
+
+(* Unused *)
+lemma cla_init_dbm_id:
+  "cla l0 ([curry init_dbm]v,n) = ([curry init_dbm]v,n)"
+proof -
+  let ?u = "\<lambda> c. 0 :: real"
+  let ?I = "\<lambda> _. Const 0"
+  let ?R = "region X ?I {}"
+  have "valid_region X (k l0) ?I {}"
+    by standard auto
+  then have "?R \<in> \<R> l0"
+    unfolding \<R>_def X_alt_def by blast
+  have region: "u \<in> ?R" if "u \<in> [curry init_dbm]v,n" for u
+    using that unfolding init_dbm_semantics X_def by - (standard; fastforce)
+  have iff: "u \<in> ?R \<longleftrightarrow> u \<in> [curry init_dbm]v,n" for u
+    apply standard
+    subgoal
+      unfolding init_dbm_semantics X_def[symmetric]
+      by (auto elim: Regions.intv_elem.cases elim!: Regions.region.cases)
+    by (rule region)
+  have single: "R = ?R" if "u \<in> [curry init_dbm]v,n" "u \<in> R" "R \<in> \<R> l0" for u R
+    using that \<open>?R \<in> \<R> l0\<close> by - (rule \<R>_regions_distinct; auto intro: region)
+  show ?thesis
+    using \<open>?R \<in> _\<close>
+    unfolding cla_def
+    apply safe
+     apply (drule single, assumption+)
+     apply (subst (asm) iff[symmetric], simp; fail)
+    apply (frule region)
+    by auto
+qed
+
+end (* TA Start *)
+
+context Reachability_Problem_precompiled
+begin
+
+lemma start_in_state_set:
+  "0 \<in> state_set A"
+  unfolding state_set_def A_def T_def using n_gt_0 start_has_trans by fastforce
+
+thm leadsto_mc[OF start_in_state_set]
+
+end (* Reachability Problem Pre-compiled *)
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/Sepref_Acconstraint.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/Sepref_Acconstraint.thy
@@ -0,0 +1,333 @@
+theory Sepref_Acconstraint
+  imports Refine_Imperative_HOL.IICF Timed_Automata.Timed_Automata
+begin
+
+  subsection \<open>Refinement Assertion\<close>
+  fun acconstraint_assn where
+(*
+    "acconstraint_assn A (E1 x) (E1 x') = A x x'"
+  | "acconstraint_assn A (E2 x) (E2 x') = A x x'"
+  | "acconstraint_assn A (E3) (E3) = emp"
+*)
+  "acconstraint_assn A B (LT x y) (LT x' y') = A x x' * B y y'"
+  | "acconstraint_assn A B (LE x y) (LE x' y') = A x x' * B y y'"
+  | "acconstraint_assn A B (EQ x y) (EQ x' y') = A x x' * B y y'"
+  | "acconstraint_assn A B (GE x y) (GE x' y') = A x x' * B y y'"
+  | "acconstraint_assn A B (GT x y) (GT x' y') = A x x' * B y y'"
+(*
+  | "acconstraint_assn A (E5 x y) (E5 x' y') = bool_assn x x' * A y y'"
+*)
+  | "acconstraint_assn _ _ _ _ = false"
+
+
+  fun acconstraint_relp where
+  "acconstraint_relp A B (LT x y) (LT x' y') \<longleftrightarrow> A x x' \<and> B y y'"
+  | "acconstraint_relp A B (LE x y) (LE x' y') \<longleftrightarrow> A x x' \<and> B y y'"
+  | "acconstraint_relp A B (EQ x y) (EQ x' y') \<longleftrightarrow> A x x' \<and> B y y'"
+  | "acconstraint_relp A B (GE x y) (GE x' y') \<longleftrightarrow> A x x' \<and> B y y'"
+  | "acconstraint_relp A B (GT x y) (GT x' y') \<longleftrightarrow> A x x' \<and> B y y'"
+  | "acconstraint_relp _ _ _ _ \<longleftrightarrow> False"
+
+  definition [to_relAPP]: "acconstraint_rel A B \<equiv> p2rel (acconstraint_relp (rel2p A) (rel2p B))"
+  
+  lemma aconstraint_assn_pure_conv[constraint_simps]:
+    "acconstraint_assn (pure A) (pure B) \<equiv> pure (\<langle>A,B\<rangle> acconstraint_rel)"
+  apply (rule eq_reflection)
+  apply (intro ext)
+  subgoal for a b
+  apply (cases a; cases b; simp add: acconstraint_rel_def pure_def p2rel_def rel2p_def)
+  done
+  done
+
+  lemmas [sepref_import_rewrite, sepref_frame_normrel_eqs, fcomp_norm_unfold] =
+    aconstraint_assn_pure_conv[symmetric]
+
+  text \<open>You might want to prove some properties\<close>
+
+  text \<open>A pure-rule is required to enable recovering of invalidated data that was not stored on the heap\<close>
+  lemma acconstraint_assn_pure[constraint_rules]: "is_pure A \<Longrightarrow> is_pure B \<Longrightarrow> is_pure (acconstraint_assn A B)"
+    apply (auto simp: is_pure_iff_pure_assn)
+    apply (rename_tac x x')
+    apply (case_tac x; case_tac x'; simp add: pure_def)
+    done
+
+  text \<open>An identitiy rule is required to easily prove trivial refinement theorems\<close>    
+  lemma acconstraint_assn_id[simp]: "acconstraint_assn id_assn id_assn = id_assn"
+    apply (intro ext)
+    subgoal for x y by (cases x; cases y; simp add: pure_def)
+    done
+
+  text \<open>With congruence condition\<close>  
+  lemma acconstraint_match_cong[sepref_frame_match_rules]: 
+    "\<lbrakk>\<And>x y. \<lbrakk>x\<in>set1_acconstraint e; y\<in>set1_acconstraint e'\<rbrakk> \<Longrightarrow> hn_ctxt A x y \<Longrightarrow>t hn_ctxt A' x y\<rbrakk>
+    \<Longrightarrow> \<lbrakk>\<And>x y. \<lbrakk>x\<in>set2_acconstraint e; y\<in>set2_acconstraint e'\<rbrakk> \<Longrightarrow> hn_ctxt B x y \<Longrightarrow>t hn_ctxt B' x y\<rbrakk>
+    \<Longrightarrow> hn_ctxt (acconstraint_assn A B) e e' \<Longrightarrow>t hn_ctxt (acconstraint_assn A' B') e e'"
+    by (cases e; cases e'; simp add: hn_ctxt_def entt_star_mono)
+      
+
+  lemma acconstraint_merge_cong[sepref_frame_merge_rules]:
+    assumes "\<And>x y. \<lbrakk>x\<in>set1_acconstraint e; y\<in>set1_acconstraint e'\<rbrakk> \<Longrightarrow> hn_ctxt A x y \<or>A hn_ctxt A' x y \<Longrightarrow>t hn_ctxt Am x y"
+    assumes "\<And>x y. \<lbrakk>x\<in>set2_acconstraint e; y\<in>set2_acconstraint e'\<rbrakk> \<Longrightarrow> hn_ctxt B x y \<or>A hn_ctxt B' x y \<Longrightarrow>t hn_ctxt Bm x y"
+    shows "hn_ctxt (acconstraint_assn A B) e e' \<or>A hn_ctxt (acconstraint_assn A' B') e e' \<Longrightarrow>t hn_ctxt (acconstraint_assn Am Bm) e e'"
+    apply (blast intro: entt_disjE acconstraint_match_cong entt_disjD1[OF assms(1)] entt_disjD2[OF assms(1)] entt_disjD1[OF assms(2)] entt_disjD2[OF assms(2)])
+    done
+
+  text \<open>Propagating invalid\<close>  
+  lemma entt_invalid_acconstraint: "hn_invalid (acconstraint_assn A B) e e' \<Longrightarrow>t hn_ctxt (acconstraint_assn (invalid_assn A) (invalid_assn B)) e e'"
+    apply (simp add: hn_ctxt_def invalid_assn_def[abs_def])
+    apply (rule enttI)
+    apply clarsimp
+    apply (cases e; cases e'; auto simp: mod_star_conv pure_def) 
+    done
+
+  lemmas invalid_acconstraint_merge[sepref_frame_merge_rules] = gen_merge_cons[OF entt_invalid_acconstraint]
+
+  subsection \<open>Constructors\<close>  
+  text \<open>Constructors need to be registered\<close>
+  sepref_register LT LE EQ GE GT
+  
+  text \<open>Refinement rules can be proven straightforwardly on the separation logic level (method @{method sepref_to_hoare})\<close>
+  (*
+  lemma [sepref_fr_rules]: "(return o E1,RETURN o E1) \<in> Ad \<rightarrow>a acconstraint_assn A"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(return o E2,RETURN o E2) \<in> Ad \<rightarrow>a acconstraint_assn A"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(uncurry0 (return E3),uncurry0 (RETURN E3)) \<in> unit_assnk \<rightarrow>a acconstraint_assn A"
+    by sepref_to_hoare sep_auto
+  *)
+  lemma [sepref_fr_rules]: "(uncurry (return oo LT),uncurry (RETURN oo LT)) \<in> Ad*aBd \<rightarrow>a acconstraint_assn A B"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(uncurry (return oo LE),uncurry (RETURN oo LE)) \<in> Ad*aBd \<rightarrow>a acconstraint_assn A B"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(uncurry (return oo EQ),uncurry (RETURN oo EQ)) \<in> Ad*aBd \<rightarrow>a acconstraint_assn A B"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(uncurry (return oo GE),uncurry (RETURN oo GE)) \<in> Ad*aBd \<rightarrow>a acconstraint_assn A B"
+    by sepref_to_hoare sep_auto
+  lemma [sepref_fr_rules]: "(uncurry (return oo GT),uncurry (RETURN oo GT)) \<in> Ad*aBd \<rightarrow>a acconstraint_assn A B"
+    by sepref_to_hoare sep_auto
+  (*
+  lemma [sepref_fr_rules]: "(uncurry (return oo E4),uncurry (RETURN oo E4)) \<in> Ad*aAd \<rightarrow>a acconstraint_assn A"
+    by sepref_to_hoare sep_auto
+  *)
+  (*
+  lemma [sepref_fr_rules]: "(uncurry (return oo E5),uncurry (RETURN oo E5)) \<in> bool_assnk*aAd \<rightarrow>a acconstraint_assn A"
+    by sepref_to_hoare (sep_auto simp: pure_def)
+  *)
+
+  subsection \<open>Destructor\<close>  
+  text \<open>There is currently no automation for destructors, so all the registration boilerplate 
+    needs to be done manually\<close>
+
+  text \<open>Set ups operation identification heuristics\<close>
+  sepref_register case_acconstraint
+
+  text \<open>In the monadify phase, this eta-expands to make visible all required arguments\<close>
+  lemma [sepref_monadify_arity]: "case_acconstraint \<equiv> \<lambda>2f1 f2 f3 f4 f5 x. SP case_acconstraint$(\<lambda>2x. f1$x)$(\<lambda>2x. f2$x)$f3$(\<lambda>2x y. f4$x$y)$(\<lambda>2x y. f5$x$y)$x"
+    by simp
+
+  text \<open>This determines an evaluation order for the first-order operands\<close>  
+  lemma [sepref_monadify_comb]: "case_acconstraint$f1$f2$f3$f4$f5$x \<equiv> (\<bind>)$(EVAL$x)$(\<lambda>2x. SP case_acconstraint$f1$f2$f3$f4$f5$x)" by simp
+
+  text \<open>This enables translation of the case-distinction in a non-monadic context.\<close>  
+  lemma [sepref_monadify_comb]: "EVAL$(case_acconstraint$(\<lambda>2x y. f1 x y)$(\<lambda>2x y. f2 x y)$(\<lambda>2x y. f3 x y)$(\<lambda>2x y. f4 x y)$(\<lambda>2x y. f5 x y)$x) 
+    \<equiv> (\<bind>)$(EVAL$x)$(\<lambda>2x. SP case_acconstraint$(\<lambda>2x y. EVAL $ f1 x y)$(\<lambda>2x y. EVAL $ f2 x y)$(\<lambda>2x y. EVAL $ f3 x y)$(\<lambda>2x y. EVAL $ f4 x y)$(\<lambda>2x y. EVAL $ f5 x y)$x)"
+    apply (rule eq_reflection)
+    by (simp split: acconstraint.splits)
+
+  text \<open>Auxiliary lemma, to lift simp-rule over \<open>hn_ctxt\<close>\<close>  
+  lemma acconstraint_assn_ctxt: "acconstraint_assn A B x y = z \<Longrightarrow> hn_ctxt (acconstraint_assn A B) x y = z"
+    by (simp add: hn_ctxt_def)
+
+  text \<open>The cases lemma first extracts the refinement for the datatype from the precondition.
+    Next, it generate proof obligations to refine the functions for every case. 
+    Finally the postconditions of the refinement are merged. 
+
+    Note that we handle the
+    destructed values separately, to allow reconstruction of the original datatype after the case-expression.
+
+    Moreover, we provide (invalidated) versions of the original compound value to the cases,
+    which allows access to pure compound values from inside the case.
+    \<close>  
+  lemma acconstraint_cases_hnr:
+    fixes A B and e :: "('a,'b) acconstraint" and e' :: "('ai,'bi) acconstraint"
+    defines [simp]: "INVe \<equiv> hn_invalid (acconstraint_assn A B) e e'"
+    assumes FR: "\<Gamma> \<Longrightarrow>t hn_ctxt (acconstraint_assn A B) e e' * F"
+    (*
+    assumes E1: "\<And>x1 x1a. \<lbrakk>e = E1 x1; e' = E1 x1a\<rbrakk> \<Longrightarrow> hn_refine (hn_ctxt A x1 x1a * INVe * F) (f1' x1a) (hn_ctxt A1' x1 x1a * hn_ctxt XX1 e e' * \<Gamma>1') R (f1 x1)"
+    assumes E2: "\<And>x2 x2a. \<lbrakk>e = E2 x2; e' = E2 x2a\<rbrakk> \<Longrightarrow> hn_refine (hn_ctxt A x2 x2a * INVe * F) (f2' x2a) (hn_ctxt A2' x2 x2a * hn_ctxt XX2 e e' * \<Gamma>2') R (f2 x2)"
+    assumes E3: "\<lbrakk>e = E3; e' = E3\<rbrakk> \<Longrightarrow> hn_refine F f3' \<Gamma>3' R f3"
+    *)
+    assumes LT: "\<And>x41 x42 x41a x42a.
+       \<lbrakk>e = LT x41 x42; e' = LT x41a x42a\<rbrakk>
+       \<Longrightarrow> hn_refine
+            (hn_ctxt A x41 x41a * hn_ctxt B x42 x42a * INVe * F)
+            (f1' x41a x42a)
+            (hn_ctxt A1a' x41 x41a * hn_ctxt B1b' x42 x42a * hn_ctxt XX1 e e' * \<Gamma>1') R
+            (f1 x41 x42)"
+    assumes LE: "\<And>x41 x42 x41a x42a.
+       \<lbrakk>e = LE x41 x42; e' = LE x41a x42a\<rbrakk>
+       \<Longrightarrow> hn_refine
+            (hn_ctxt A x41 x41a * hn_ctxt B x42 x42a * INVe * F)
+            (f2' x41a x42a)
+            (hn_ctxt A2a' x41 x41a * hn_ctxt B2b' x42 x42a * hn_ctxt XX2 e e' * \<Gamma>2') R
+            (f2 x41 x42)"
+    assumes EQ: "\<And>x41 x42 x41a x42a.
+       \<lbrakk>e = EQ x41 x42; e' = EQ x41a x42a\<rbrakk>
+       \<Longrightarrow> hn_refine
+            (hn_ctxt A x41 x41a * hn_ctxt B x42 x42a * INVe * F)
+            (f3' x41a x42a)
+            (hn_ctxt A3a' x41 x41a * hn_ctxt B3b' x42 x42a * hn_ctxt XX3 e e' * \<Gamma>3') R
+            (f3 x41 x42)"
+    assumes GE: "\<And>x41 x42 x41a x42a.
+       \<lbrakk>e = GE x41 x42; e' = GE x41a x42a\<rbrakk>
+       \<Longrightarrow> hn_refine
+            (hn_ctxt A x41 x41a * hn_ctxt B x42 x42a * INVe * F)
+            (f4' x41a x42a)
+            (hn_ctxt A4a' x41 x41a * hn_ctxt B4b' x42 x42a * hn_ctxt XX4 e e' * \<Gamma>4') R
+            (f4 x41 x42)"
+    assumes GT: "\<And>x41 x42 x41a x42a.
+       \<lbrakk>e = GT x41 x42; e' = GT x41a x42a\<rbrakk>
+       \<Longrightarrow> hn_refine
+            (hn_ctxt A x41 x41a * hn_ctxt B x42 x42a * INVe * F)
+            (f5' x41a x42a)
+            (hn_ctxt A5a' x41 x41a * hn_ctxt B5b' x42 x42a * hn_ctxt XX5 e e' * \<Gamma>5') R
+            (f5 x41 x42)"
+    (*
+    assumes E5: "\<And>x51 x52 x51a x52a.
+       \<lbrakk>e = E5 x51 x52; e' = E5 x51a x52a\<rbrakk>
+       \<Longrightarrow> hn_refine (hn_ctxt bool_assn x51 x51a * hn_ctxt A x52 x52a * INVe * F) (f5' x51a x52a)
+            (hn_ctxt bool_assn x51 x51a * hn_ctxt A5' x52 x52a * hn_ctxt XX5 e e' * \<Gamma>5') R (f5 x51 x52)"
+    *)
+    assumes MERGE1a[unfolded hn_ctxt_def]:
+      "\<And>x x'. hn_ctxt A1a' x x' \<or>A hn_ctxt A2a' x x' \<or>A hn_ctxt A3a' x x' \<or>A hn_ctxt A4a' x x' \<or>A hn_ctxt A5a' x x' \<Longrightarrow>t hn_ctxt A' x x'"
+    assumes MERGE1b[unfolded hn_ctxt_def]:
+      "\<And>x x'. hn_ctxt B1b' x x' \<or>A hn_ctxt B2b' x x' \<or>A hn_ctxt B3b' x x' \<or>A hn_ctxt B4b' x x' \<or>A hn_ctxt B5b' x x' \<Longrightarrow>t hn_ctxt B' x x'"
+    assumes MERGE2[unfolded hn_ctxt_def]: "\<Gamma>1' \<or>A \<Gamma>2' \<or>A \<Gamma>3' \<or>A \<Gamma>4' \<or>A \<Gamma>5' \<Longrightarrow>t \<Gamma>'"
+    shows "hn_refine \<Gamma> (case_acconstraint f1' f2' f3' f5' f4' e') (hn_ctxt (acconstraint_assn A' B') e e' * \<Gamma>') R (case_acconstraint$(\<lambda>2x y. f1 x y)$(\<lambda>2x y. f2 x y)$(\<lambda>2x y. f3 x y)$(\<lambda>2x y. f5 x y)$(\<lambda>2x y. f4 x y)$e)"
+    
+    apply (rule hn_refine_cons_pre[OF FR])
+    apply1 extract_hnr_invalids
+    apply (cases e; cases e'; simp add: acconstraint_assn.simps[THEN acconstraint_assn_ctxt])
+    subgoal 
+      apply (rule hn_refine_cons[OF _ LT _ entt_refl]; assumption?)
+      applyS (simp add: hn_ctxt_def)
+      apply (rule entt_star_mono)
+      apply1 (rule entt_fr_drop)
+      apply (rule entt_star_mono)
+
+      apply1 (rule entt_trans[OF _ MERGE1a])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE1b])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE2])
+      applyS (simp add: entt_disjI1' entt_disjI2')
+    done
+    subgoal 
+      apply (rule hn_refine_cons[OF _ LE _ entt_refl]; assumption?)
+      applyS (simp add: hn_ctxt_def)
+      apply (rule entt_star_mono)
+      apply1 (rule entt_fr_drop)
+      apply (rule entt_star_mono)
+
+      apply1 (rule entt_trans[OF _ MERGE1a])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE1b])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE2])
+      applyS (simp add: entt_disjI1' entt_disjI2')
+    done
+    subgoal 
+      apply (rule hn_refine_cons[OF _ EQ _ entt_refl]; assumption?)
+      applyS (simp add: hn_ctxt_def)
+      apply (rule entt_star_mono)
+      apply1 (rule entt_fr_drop)
+      apply (rule entt_star_mono)
+
+      apply1 (rule entt_trans[OF _ MERGE1a])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE1b])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE2])
+      applyS (simp add: entt_disjI1' entt_disjI2')
+    done
+    subgoal
+      apply (rule hn_refine_cons[OF _ GT _ entt_refl]; assumption?)
+      applyS (simp add: hn_ctxt_def)
+      apply (rule entt_star_mono)
+      apply1 (rule entt_fr_drop)
+      apply (rule entt_star_mono)
+
+      apply1 (rule entt_trans[OF _ MERGE1a])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE1b])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE2])
+      applyS (simp add: entt_disjI1' entt_disjI2')
+    done
+    subgoal 
+      apply (rule hn_refine_cons[OF _ GE _ entt_refl]; assumption?)
+      applyS (simp add: hn_ctxt_def)
+      apply (rule entt_star_mono)
+      apply1 (rule entt_fr_drop)
+      apply (rule entt_star_mono)
+
+      apply1 (rule entt_trans[OF _ MERGE1a])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE1b])
+      applyS (simp add: hn_ctxt_def entt_disjI1' entt_disjI2')
+
+      apply1 (rule entt_trans[OF _ MERGE2])
+      applyS (simp add: entt_disjI1' entt_disjI2')
+    done
+  done
+
+  text \<open>After some more preprocessing (adding extra frame-rules for non-atomic postconditions, 
+    and splitting the merge-terms into binary merges), this rule can be registered\<close>
+  lemmas [sepref_comb_rules] = acconstraint_cases_hnr[sepref_prep_comb_rule]
+
+  subsection \<open>Regression Test\<close>
+(*
+
+  definition "test \<equiv> do {
+    let x = E1 True;
+
+    _ \<leftarrow> case x of
+      E1 _ \<Rightarrow> RETURN x  (* Access compound inside case *)
+    | _ \<Rightarrow> RETURN E3;  
+
+    (* Now test with non-pure *)
+    let a = op_array_replicate 4 (3::nat);
+    let x = E5 False a;
+    
+    _ \<leftarrow> case x of
+      E1 _ \<Rightarrow> RETURN (0::nat)
+    | E2 _ \<Rightarrow> RETURN 1
+    | E3 \<Rightarrow> RETURN 0
+    | E4 _ _ \<Rightarrow> RETURN 0
+    | E5 _ a \<Rightarrow> mop_list_get a 0;
+
+    (* Rely on that compound still exists (it's components are only read in the case above) *)
+    case x of
+      E1 a \<Rightarrow> do {mop_list_set a 0 0; RETURN (0::nat)}
+    | E2 _ \<Rightarrow> RETURN 1
+    | E3 \<Rightarrow> RETURN 0
+    | E4 _ _ \<Rightarrow> RETURN 0
+    | E5 _ _ \<Rightarrow> RETURN 0
+  }"
+
+  sepref_definition foo is "SYNTH (uncurry0 test) (unit_assnk \<rightarrow>a nat_assn)"      
+    unfolding test_def
+    supply [[goals_limit=1]]
+    by sepref
+*)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/TA_DBM_Operations_Impl.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/TA_DBM_Operations_Impl.thy
@@ -0,0 +1,217 @@
+theory TA_DBM_Operations_Impl
+  imports
+    Timed_Automata.TA_DBM_Operations
+    Difference_Bound_Matrices.DBM_Operations_Impl_Refine
+    Sepref_Acconstraint
+    Munta_Base.TA_Syntax_Bundles
+begin
+
+unbundle no_library_syntax
+
+section \<open>Constraints\<close>
+
+(*
+fun abstra' :: "(nat, 't::time) acconstraint \<Rightarrow> 't DBM \<Rightarrow> 't DBM"
+where
+  "abstra' (EQ c d) M =
+    (\<lambda> i j . if i = 0 \<and> j = c then min (M i j) (Le (-d)) else if i = c \<and> j = 0 then min (M i j) (Le d) else M i j)" |
+  "abstra' (LT c d) M =
+    (\<lambda> i j . if i = c \<and> j = 0 then min (M i j) (Lt d) else M i j)" |
+  "abstra' (LE c d) M =
+    (\<lambda> i j . if i = c \<and> j = 0 then min (M i j) (Le d) else M i j)" |
+  "abstra' (GT c d) M =
+    (\<lambda> i j. if i = 0 \<and> j = c then min (M i j) (Lt (- d)) else M i j)" |
+  "abstra' (GE c d) M =
+    (\<lambda> i j. if i = 0 \<and> j = c then min (M i j) (Le (- d)) else M i j)"
+*)
+
+definition "abstra' ac M = abstra ac M id"
+
+lemma abstra_abstra':
+  assumes "v (constraint_clk ac) = constraint_clk ac"
+  shows "abstra ac M v = abstra' ac M"
+unfolding abstra'_def using assms by (cases ac; fastforce)
+
+definition "abstr' cc M = abstr cc M id"
+
+lemma abstr_abstr':
+  assumes "\<forall> c \<in> collect_clks cc. v c = c"
+  shows "abstr cc M v = abstr' cc M"
+unfolding abstr'_def using assms
+by (auto simp: collect_clks_def intro: fold_cong abstra_abstra'[unfolded abstra'_def])
+
+lemma And_abstr:
+  assumes "clock_numbering' v n" "\<forall> c \<in> collect_clks cc. v c \<le> n"
+  shows "[And M (abstr cc (\<lambda> i j. \<infinity>) v)]v,n = [abstr cc M v]v,n"
+using And_correct dbm_abstr_zone_eq[OF assms] dbm_abstr_zone_eq2[OF assms] by blast
+
+lemma min_inf_r:
+  "min a \<infinity> = a"
+  by (cases a) (auto split: split_min simp: less_eq dbm_le_def)
+
+lemma min_inf_l:
+  "min \<infinity> b = b"
+  by (cases b) (auto split: split_min simp: less_eq dbm_le_def)
+
+lemma And_abstra_commute:
+  assumes "clock_numbering' v n" "\<forall> c \<in> collect_clks cc. v c \<le> n"
+  shows "And M (abstra ac (\<lambda> i j. \<infinity>) v) = abstra ac M v"
+  by (cases ac) (auto simp: min_inf_l min_inf_r any_le_inf intro!: ext)
+
+(* Could be proven by using the facts that abstra commutes with itself and And*)
+lemma
+  assumes "clock_numbering' v n" "\<forall> c \<in> collect_clks cc. v c \<le> n"
+  shows "And M (abstr cc (\<lambda> i j. \<infinity>) v) = abstr cc M v"
+apply (induction cc)
+apply simp
+apply (simp add: min_inf_r)
+oops
+
+lemma abstra_canonical_diag_preservation:
+  assumes "clock_numbering v"
+  shows "(abstra ac M v) i i = M i i"
+using assms by (cases ac) (auto dest: clock_numberingD)
+
+lemma abstr_canonical_diag_preservation:
+  assumes "clock_numbering v"
+  shows "(abstr cc M v) i i = M i i"
+using assms
+apply (induction cc arbitrary: M)
+apply (simp; fail)
+ subgoal for ac cc M
+ using abstra_canonical_diag_preservation[where v = v and ac = ac] by auto
+done
+
+fun abstra_upd :: "(nat, 't::{linorder,uminus}) acconstraint \<Rightarrow> 't DBM' \<Rightarrow> 't DBM'"
+where
+  "abstra_upd (EQ c d) M =
+    (let
+      m0c = op_mtx_get M (0, c);
+      mc0 = op_mtx_get M(c, 0)
+     in M((0, c) := min m0c (Le (-d)), (c, 0) := min mc0 (Le d)))" |
+  "abstra_upd (LT c d) M =
+    M((c, 0) := min (op_mtx_get M (c, 0)) (Lt d))" |
+  "abstra_upd (LE c d) M =
+    M((c, 0) := min (op_mtx_get M (c, 0)) (Le d))" |
+  "abstra_upd (GT c d) M =
+    M((0, c) := min (op_mtx_get M ((0, c))) (Lt (- d)))" |
+  "abstra_upd (GE c d) M =
+    M((0, c) := min (op_mtx_get M (0, c)) (Le (- d)))"
+
+lemma abstra_upd_alt_def:
+  "abstra_upd ac M = case_acconstraint
+    (\<lambda> c d. M((c, 0) := min (op_mtx_get M (c, 0)) (Lt d)))
+    (\<lambda> c d. M((c, 0) := min (op_mtx_get M (c, 0)) (Le d)))
+    (\<lambda> c d. let m0c = op_mtx_get M (0, c); mc0 = op_mtx_get M(c, 0) in M((0, c) := min m0c (Le (-d)), (c, 0) := min mc0 (Le d)))
+    (\<lambda> c d. M((0, c) := min (op_mtx_get M ((0, c))) (Lt (- d))))
+    (\<lambda> c d. M((0, c) := min (op_mtx_get M (0, c)) (Le (- d))))
+    ac"
+by (cases ac) auto
+
+lemma abstra_upd_eq_abstra':
+  assumes "constraint_clk ac > 0"
+  shows "curry (abstra_upd ac M) = abstra' ac (curry M)"
+unfolding abstra'_def using assms by (cases ac; fastforce)
+
+lemma abstra_upd_int_preservation:
+  assumes "snd (constraint_pair ac) \<in> \<int>" "dbm_int (curry M) n"
+  shows "dbm_int (curry (abstra_upd ac M)) n"
+using assms by (cases ac; simp split: split_min)
+
+definition abstr_upd where
+  "abstr_upd = fold (\<lambda> ac M. abstra_upd ac M)"
+
+lemma abstr_upd_alt_def:
+  "RETURN oo abstr_upd = (\<lambda> cc M. nfoldli cc (\<lambda> _. True) (\<lambda> ac M. RETURN (abstra_upd ac M)) M)"
+by (intro ext, simp add: abstr_upd_def fold_eq_nfoldli)
+
+
+lemma abstr_upd_abstr':
+  assumes "\<forall> c \<in> collect_clks cc. c > 0"
+  shows "curry (abstr_upd cc M) = abstr' cc (curry M)"
+unfolding abstr_upd_def abstr'_def using assms
+by (induction cc arbitrary: M) (auto simp: abstra_abstra' abstra_upd_eq_abstra')
+
+lemma abstra_upd_out_of_bounds1:
+  assumes "constraint_clk ac \<le> n" "i > n"
+  shows "(abstra_upd ac M) (i, j) = M (i, j)"
+using assms by (cases ac) auto
+
+lemma abstra_upd_out_of_bounds2:
+  assumes "constraint_clk ac \<le> n" "j > n"
+  shows "(abstra_upd ac M) (i, j) = M (i, j)"
+using assms by (cases ac) auto
+
+lemma abstr_upd_out_of_bounds1:
+  assumes "\<forall> c \<in> collect_clks cc. c \<le> n" "i > n"
+  shows "(abstr_upd cc M) (i, j) = M (i, j)"
+using assms by (induction cc arbitrary: M) (auto simp: abstr_upd_def abstra_upd_out_of_bounds1)
+
+lemma abstr_upd_out_of_bounds2:
+  assumes "\<forall> c \<in> collect_clks cc. c \<le> n" "j > n"
+  shows "(abstr_upd cc M) (i, j) = M (i, j)"
+using assms by (induction cc arbitrary: M) (auto simp: abstr_upd_def abstra_upd_out_of_bounds2)
+
+lemma abstr_upd_int_preservation:
+  assumes "\<forall> (_,d) \<in> collect_clock_pairs cc. d \<in> \<int>" "dbm_int (curry M) n"
+  shows "dbm_int (curry (abstr_upd cc M)) n"
+unfolding abstr_upd_def using assms
+proof (induction cc arbitrary: M)
+  case Nil then show ?case by simp
+next
+  case (Cons c cc)
+  show ?case
+   apply simp
+   apply (rule Cons.IH[simplified])
+   defer
+   apply (rule abstra_upd_int_preservation[simplified])
+  using Cons.prems unfolding collect_clock_pairs_def by auto
+qed
+
+(*
+fun abstra_upd :: "('c, 't::time) acconstraint \<Rightarrow> 't DBM' \<Rightarrow> ('c \<Rightarrow> nat) \<Rightarrow> 't DBM'"
+where
+  "abstra_upd (EQ c d) M v =
+    (let m0c = op_mtx_get M (0, v c); mc0 = op_mtx_get M(v c, 0) in M((0, v c) := min m0c (Le (-d)), (v c, 0) := min mc0 (Le d)))" |
+  "abstra_upd (LT c d) M v =
+    M((v c, 0) := min (op_mtx_get M (v c, 0)) (Lt d))" |
+  "abstra_upd (LE c d) M v =
+    M((v c, 0) := min (op_mtx_get M (v c, 0)) (Le d))" |
+  "abstra_upd (GT c d) M v =
+    M((0, v c) := min (op_mtx_get M ((0, v c))) (Lt (- d)))" |
+  "abstra_upd (GE c d) M v =
+    M((0, v c) := min (op_mtx_get M (0, v c)) (Le (- d)))"
+
+lemma abstra_upd_eq_abstra:
+  assumes "clock_numbering v"
+  shows "curry (abstra_upd ac M v) = abstra ac (curry M) v"
+using assms by (cases ac; fastforce dest: clock_numberingD)
+*)
+
+context
+  fixes n :: nat
+begin
+
+interpretation DBM_Impl n .
+
+sepref_definition abstra_upd_impl is
+  "uncurry (RETURN oo abstra_upd)" ::
+  "(acconstraint_assn clock_assn id_assn)k *a mtx_assnd \<rightarrow>a mtx_assn"
+unfolding abstra_upd_alt_def zero_clock_def[symmetric] by sepref
+
+sepref_register abstra_upd ::
+  "(nat, ('a :: {linordered_cancel_ab_monoid_add,uminus,heap})) acconstraint
+  \<Rightarrow> 'a DBMEntry i_mtx \<Rightarrow> 'a DBMEntry i_mtx"
+
+lemmas [sepref_fr_rules] = abstra_upd_impl.refine
+
+sepref_definition abstr_upd_impl is
+  "uncurry (RETURN oo abstr_upd)"
+  :: "(list_assn (acconstraint_assn clock_assn id_assn))k *a mtx_assnd \<rightarrow>a mtx_assn"
+  unfolding abstr_upd_alt_def by sepref
+
+end
+
+export_code abstr_upd_impl abstra_upd_impl in SML_imp
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/TA_Impl_Misc.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/TA_Impl_Misc.thy
@@ -0,0 +1,71 @@
+theory TA_Impl_Misc
+  imports Main
+begin
+
+subsection \<open>Invariants on folds\<close>
+
+lemma fold_acc_preserv:
+  assumes "\<And> x acc. P acc \<Longrightarrow> P (f x acc)" "P acc"
+  shows "P (fold f xs acc)"
+  using assms(2) by (induction xs arbitrary: acc) (auto intro: assms(1))
+
+lemma fold_acc_preserv':
+  assumes "\<And> x acc. x \<in> set xs \<Longrightarrow> P acc \<Longrightarrow> P (f x acc)" "P acc"
+  shows "P (fold f xs acc)"
+  using assms by (induction xs arbitrary: acc) auto
+
+lemma fold_acc_ev_preserv':
+  fixes x
+  assumes
+    "\<And> x acc. P acc \<Longrightarrow> P (f x acc)" "\<And> x acc. Q acc \<Longrightarrow> Q (f x acc)"
+    "\<And> acc. Q acc \<Longrightarrow> P (f x acc)" "x \<in> set xs" "Q acc"
+  shows "P (fold f xs acc) \<and> Q (fold f xs acc)"
+  using assms(4,5) apply (induction xs arbitrary: acc)
+  using assms(1,2,3) by (auto intro: fold_acc_preserv)
+
+lemma fold_acc_ev_preserv:
+  fixes x
+  assumes
+    "\<And> x acc. P acc \<Longrightarrow> Q acc \<Longrightarrow> P (f x acc)" "\<And> x acc. Q acc \<Longrightarrow> Q (f x acc)"
+    "\<And> acc. Q acc \<Longrightarrow> P (f x acc)" "x \<in> set xs" "Q acc"
+  shows "P (fold f xs acc) \<and> Q (fold f xs acc)"
+  apply (rule fold_acc_ev_preserv'[where P = "\<lambda> acc. P acc \<and> Q acc" and Q = Q, THEN conjunct1])
+  by (auto intro: assms)
+
+lemmas fold_ev_preserv = fold_acc_ev_preserv[where Q = "\<lambda> _. True", simplified]
+
+lemma fold_evD':
+  assumes "\<not> P acc" "P (fold f xs acc)"
+    shows "\<exists> x ys zs. xs = ys @ x # zs \<and> \<not> P (fold f ys acc) \<and> P (f x (fold f ys acc))"
+  using assms
+  apply (induction xs arbitrary: acc)
+   apply (simp; fail)
+  subgoal premises prems for x xs acc
+    apply (cases "P (f x acc)")
+     using prems(2-) apply fastforce
+     apply (drule prems(1))
+     using prems apply simp
+     apply clarsimp
+     subgoal for xa ys zs
+       apply (rule exI[where x = xa])
+       apply (rule exI[where x = "x # ys"])
+       by fastforce
+     done
+   done
+
+lemma fold_evD:
+  assumes
+    "P y (fold f xs acc)" "\<not> P y acc" "\<And> acc x. \<not> P y acc \<Longrightarrow> Q acc \<Longrightarrow> P y (f x acc) \<Longrightarrow> x = y"
+    "Q acc" "\<And> acc x. Q acc \<Longrightarrow> Q (f x acc)" "\<And> acc x. \<not> P y acc \<Longrightarrow> Q acc \<Longrightarrow> P y (f x acc) \<Longrightarrow> R y"
+  shows "\<exists> ys zs. xs = ys @ y # zs \<and> \<not> P y (fold f ys acc) \<and> P y (f y (fold f ys acc)) \<and> R y"
+proof -
+  from fold_evD'[OF assms(2,1)] obtain x ys zs where *:
+    "xs = ys @ x # zs" "\<not> P y (fold f ys acc)" "P y (f x (fold f ys acc))"
+    by auto
+  moreover from assms(4-) have "Q (fold f ys acc)" by (auto intro: fold_acc_preserv)
+  ultimately show ?thesis using assms(3,6) by auto
+qed
+
+lemmas fold_evD'' = fold_evD[where R = "\<lambda> _. True", simplified]
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/TA_Impl/TA_More.thy
--- /dev/null
+++ thys/Munta_Model_Checker/TA_Impl/TA_More.thy
@@ -0,0 +1,220 @@
+theory TA_More
+  imports
+    Timed_Automata.Timed_Automata
+    Timed_Automata.TA_DBM_Operations
+    Refine_Imperative_HOL.IICF
+begin
+
+section \<open>Mixed Additional Material on Timed Automata\<close>
+
+text \<open>
+  Additional Lemmas on Timed Automata that do not belong to a particular other part of the
+  formalization.
+\<close>
+
+fun extend_cc where
+  "extend_cc cc [] = cc"
+| "extend_cc cc (c # cs) = GE c 0 # extend_cc cc cs"
+
+lemma collect_clock_pairs_extend_cc:
+  "collect_clock_pairs (extend_cc cc cs) = collect_clock_pairs cc \<union> {(c, 0) | c. c \<in> set cs}"
+  by (induction cs) (auto simp: collect_clock_pairs_def)
+
+definition
+  "extend_ta A xs \<equiv> (trans_of A, \<lambda> l. extend_cc (inv_of A l) xs)"
+
+(* Interestingly, this does not go through without the IICF simpset *)
+lemma clk_set_extend_ta:
+  "clk_set (extend_ta A xs) = clk_set A \<union> set xs"
+  unfolding extend_ta_def clkp_set_def inv_of_def trans_of_def
+  by (auto simp: collect_clki_def collect_clock_pairs_extend_cc)
+
+lemma extend_cc_iff:
+  "u \<turnstile> extend_cc cc cs \<longleftrightarrow> u \<turnstile> cc" if "\<forall> c. u c \<ge> 0"
+  using that by (induction cs) (force simp: clock_val_def)+
+
+lemma [simp]:
+  "trans_of (extend_ta A cs) = trans_of A"
+  unfolding extend_ta_def trans_of_def by simp
+
+lemma [simp]:
+  "inv_of (extend_ta A cs) l = extend_cc (inv_of A l) cs"
+  unfolding extend_ta_def inv_of_def by simp
+
+lemma reset_V:
+  "\<forall> c. ([r\<rightarrow>0]u) c \<ge> 0" if "\<forall> c. u c \<ge> 0"
+  using that apply safe
+  subgoal for c
+    by (cases "c \<in> set r") (auto simp add: reset_set1 reset_set11)
+  done
+
+lemma delay_V:
+  "\<forall> c. (u \<oplus> d) c \<ge> 0" if "\<forall> c. u c \<ge> 0" "(d :: 't :: time) \<ge> 0"
+  using that unfolding cval_add_def by auto
+
+context
+  notes [elim!]  = step.cases step'.cases step_t.cases step_z.cases
+begin
+
+lemma extend_ta_iff:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle> \<longleftrightarrow> extend_ta A cs \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>" if "\<forall> c. u c \<ge> 0"
+  using that apply safe
+     apply (force simp: extend_cc_iff reset_V intro: step_a.intros elim!: step_a.cases)
+    apply (force simp add: extend_cc_iff delay_V)
+   apply (cases rule: step_a.cases, assumption)
+   apply (rule step.intros, rule step_a.intros)
+  unfolding extend_ta_def apply (force simp: trans_of_def)
+     apply assumption+
+  prefer 2
+  apply assumption
+  subgoal
+    by (metis extend_cc_iff inv_of_def prod.sel(2) reset_V)
+  subgoal
+    by (metis delay_V extend_cc_iff inv_of_def prod.sel(2) step_t step_t.intros)
+  done
+    (*  apply simp
+apply (force simp add: extend_cc_iff)
+   apply (auto simp add: reset_V extend_cc_iff intro: step_a.intros)
+  using extend_cc_iff delay_V by fas *)
+
+lemma steps_iffI:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<longleftrightarrow> A' \<turnstile> \<langle>l, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>" if
+  "\<And> l u l' u'. P u \<Longrightarrow> A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle> \<longleftrightarrow> A' \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>"
+  "\<And> (A :: ('a, 'c, 't :: time, 's) ta) l u l' u'. A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle> \<Longrightarrow> P u \<Longrightarrow> P u'" "P u" for
+  A A' :: "('a, 'c, 't :: time, 's) ta"
+  apply standard
+   apply (insert that(3))
+  subgoal
+    by (induction A\<equiv> A _ _ _ _ rule: steps.induct) (auto simp: that(1,2))
+  by (induction A\<equiv> A' _ _ _ _ rule: steps.induct) (auto simp: that(2) that(1)[symmetric])
+
+lemma step_V:
+  "\<forall> c. u' c \<ge> 0" if "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>" "\<forall> c. u c \<ge> 0"
+  using that by (auto simp: reset_V delay_V elim: step_a.cases)
+
+end (* End of context for aggressive elimination rules *)
+
+lemma extend_ta_iff_steps:
+  "A \<turnstile> \<langle>l, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<longleftrightarrow> extend_ta A cs \<turnstile> \<langle>l, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>" if "\<forall> c. u c \<ge> 0"
+  apply (intro steps_iffI extend_ta_iff)
+    apply assumption
+  by (rule step_V, assumption+, rule that)
+
+
+
+lemma collect_clkvt_alt_def:
+  "collect_clkvt T = \<Union>((set o fst \<circ> snd \<circ> snd \<circ> snd) ` T)"
+unfolding collect_clkvt_def by fastforce
+
+lemma ta_collect_clkt_alt_def:
+  "collect_clkt S = \<Union> (collect_clock_pairs ` (fst o snd) ` S)"
+unfolding Timed_Automata.collect_clkt_def by fastforce
+
+lemma ta_collect_clki_alt_def:
+  "collect_clki I = \<Union> (collect_clock_pairs ` I ` UNIV)"
+unfolding Timed_Automata.collect_clki_def by auto
+
+lemma constraint_clk_constraint_pair:
+  "constraint_clk ac = fst (constraint_pair ac)"
+by (cases ac) auto
+
+lemma collect_clks_inv_clk_set:
+  "collect_clks (inv_of A l) \<subseteq> clk_set A"
+  unfolding
+    clkp_set_def collect_clki_def collect_clks_def collect_clock_pairs_def
+  by (auto simp: constraint_clk_constraint_pair)
+
+lemma collect_clocks_clk_set:
+  assumes
+    "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  shows
+    "collect_clks g \<subseteq> clk_set A"
+  using assms unfolding clkp_set_def collect_clks_id collect_clkt_def by fastforce
+
+lemma reset_clk_set:
+  assumes
+    "A \<turnstile> l \<longrightarrow>g,a,r l'"
+  shows
+    "set r \<subseteq> clk_set A"
+using assms by (fastforce simp: clkp_set_def collect_clkvt_def)
+
+lemma ac_iff:
+  "u1 \<turnstile>a ac \<longleftrightarrow> u2 \<turnstile>a ac" if
+  "u1 (fst (constraint_pair ac)) = u2 (fst (constraint_pair ac))"
+  using that by (cases ac) auto
+
+lemma ac_iff':
+  "u1 \<turnstile>a ac \<longleftrightarrow> u2 \<turnstile>a ac" if
+  "u1 (constraint_clk ac) = u2 (constraint_clk ac)"
+  using that by (cases ac) auto
+
+lemma cc_iff:
+  "u1 \<turnstile> cc \<longleftrightarrow> u2 \<turnstile> cc" if "\<forall> (c, d) \<in> collect_clock_pairs cc. u1 c = u2 c"
+  using that
+  by (auto 4 3 simp: ac_iff[of u1 _ u2] list_all_iff collect_clock_pairs_def clock_val_def)
+
+lemma cc_iff':
+  "u1 \<turnstile> cc \<longleftrightarrow> u2 \<turnstile> cc" if "\<forall> c \<in> collect_clks cc. u1 c = u2 c"
+  using that
+  by (auto simp: ac_iff'[of u1 _ u2] list_all_iff collect_clks_def clock_val_def)
+
+lemma step_t_bisim:
+  "\<exists> u2'. A \<turnstile> \<langle>l, u2\<rangle> \<rightarrow>d \<langle>l', u2'\<rangle> \<and> (\<forall> c. c \<in> clk_set A \<longrightarrow> u1' c = u2' c)"
+  if assms: "A \<turnstile> \<langle>l, u1\<rangle> \<rightarrow>d \<langle>l', u1'\<rangle>" "\<forall> c. c \<in> clk_set A \<longrightarrow> u1 c = u2 c"
+  using that(1)
+  apply cases
+  apply (subst (asm) cc_iff'[of _ _ "u2 \<oplus> d"])
+  subgoal
+    using collect_clks_inv_clk_set[of A l] assms(2) by (auto simp: cval_add_def)
+  using assms(2) by (force simp: cval_add_def)
+
+lemma step_a_bisim:
+  "\<exists> u2'. A \<turnstile> \<langle>l, u2\<rangle> \<rightarrow>a \<langle>l', u2'\<rangle> \<and> (\<forall> c. c \<in> clk_set A \<longrightarrow> u1' c = u2' c)"
+  if assms: "A \<turnstile> \<langle>l, u1\<rangle> \<rightarrow>a \<langle>l', u1'\<rangle>" "\<forall> c. c \<in> clk_set A \<longrightarrow> u1 c = u2 c"
+  using that(1)
+  apply cases
+  subgoal for g r
+    apply (subst (asm) cc_iff'[of _ _ "u2"])
+    subgoal
+      using assms(2) by (force dest: collect_clocks_clk_set)
+    apply (subst (asm) (2) cc_iff'[of _ _ "[r\<rightarrow>0]u2"])
+    subgoal
+      apply rule
+      subgoal for c
+        using collect_clks_inv_clk_set[of A l'] assms(2) by (cases "c \<in> set r"; force)
+      done
+    apply (intro exI conjI)
+     apply (rule, assumption+)
+     apply (simp; fail)
+    apply rule
+    subgoal for c
+      using collect_clks_inv_clk_set[of A l'] assms(2) by (cases "c \<in> set r"; force)
+    done
+  done
+
+lemma step'_bisim:
+  "\<exists> u2'. A \<turnstile>' \<langle>l, u2\<rangle> \<rightarrow> \<langle>l', u2'\<rangle> \<and> (\<forall> c. c \<in> clk_set A \<longrightarrow> u1' c = u2' c)"
+  if assms: "A \<turnstile>' \<langle>l, u1\<rangle> \<rightarrow> \<langle>l', u1'\<rangle>" "\<forall> c. c \<in> clk_set A \<longrightarrow> u1 c = u2 c"
+  using that(1)
+  apply cases
+  apply (drule step_t_bisim[OF _ that(2)])
+  apply clarify
+  apply (drule step_a_bisim, assumption)
+  apply auto
+  done
+
+lemma ta_bisimulation:
+  "Bisimulation_Invariant
+  (\<lambda> (l, u) (l', u'). A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+  (\<lambda> (l, u) (l', u'). A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+  (\<lambda> (l, u) (l', u'). l' = l \<and> (\<forall> c. c \<in> clk_set A \<longrightarrow> u c = u' c))
+  (\<lambda> _. True) (\<lambda> _. True)
+  "
+  apply standard
+  subgoal for a b a'
+    by clarify (drule step'_bisim; auto)
+  subgoal for a b a'
+    by clarify (drule step'_bisim; auto)
+  ..
+
+end (* Theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/Program_Analysis.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/Program_Analysis.thy
@@ -0,0 +1,1408 @@
+theory Program_Analysis
+  imports
+    UPPAAL_State_Networks_Impl
+    Munta_Base.More_Methods
+begin
+
+fun steps_approx :: "nat \<Rightarrow> 't instrc option list \<Rightarrow> addr \<Rightarrow> addr set" where
+  "steps_approx 0 prog pc = (if pc < length prog then {pc} else {})" |
+  "steps_approx (Suc n) prog pc =
+    (
+      if pc \<ge> length prog
+      then {}
+      else
+        case prog ! pc of
+          None \<Rightarrow> {pc}
+        | Some cmd \<Rightarrow>
+          let succs =
+            (
+              case cmd of
+                CEXP ac \<Rightarrow> {pc + 1}
+              | INSTR instr \<Rightarrow> (
+                  case instr of
+                    CALL   \<Rightarrow> {..<length prog}
+                  | RETURN \<Rightarrow> {..<length prog}
+                  | JMPZ pc' \<Rightarrow> {pc + 1, pc'}
+                  | HALT \<Rightarrow> {}
+                  |    _ \<Rightarrow> {pc + 1})
+            )
+          in {pc} \<union> \<Union> (steps_approx n prog ` succs)
+    )
+  "
+
+lemma bounded_less_simp[simp]:
+  "\<forall>q\<in>{..<p::nat}. P q \<equiv> \<forall> q < p. P q"
+  by (rule eq_reflection) auto
+
+context
+  fixes prog :: "int instrc option list"
+    and strip :: "real instrc \<Rightarrow> instr"
+  assumes instr_id[simp]:
+    "strip (INSTR cmd) = cmd"
+    "strip (CEXP ac) \<notin> {CALL, RETURN, HALT} \<union> (JMPZ ` UNIV)"
+begin
+
+(* We could move this to a different file but rather would like to introduce a private namespace
+   on the spot *)
+private definition [simp]:
+  "P' \<equiv> map_option strip o (conv_prog (\<lambda> i. if i < length prog then prog ! i else None))"
+
+lemma steps_out_of_range':
+  assumes "steps P' n (pc, st, s, f, rs) (pc', st', s', f', rs')" "pc \<ge> length prog"
+  shows "pc' = pc"
+  using assms by cases (auto simp: striptp_def)
+
+lemmas steps_out_of_range = steps_out_of_range'[unfolded striptp_def P'_def]
+
+lemma steps_steps_approx':
+  assumes "steps P' n (pc, st, s, f, rs) (pc', st', s', f', rs')" "pc' < length prog"
+  shows "pc' \<in> steps_approx n prog pc"
+proof -
+  {
+    fix pc :: nat
+      and st :: "int list"
+      and m :: "int list"
+      and f :: bool
+      and rs :: "nat list"
+      and a :: nat
+      and aa :: "int list"
+      and ab :: "int list"
+      and ac :: bool
+      and b :: "nat list"
+      and n :: nat
+      and za :: "int instrc"
+      and x2 :: "int instrc"
+    assume A:
+      "UPPAAL_Asm.step (strip (map_instrc real_of_int za)) (pc, st, m, f, rs)
+      = Some (a, aa, ab, ac, b)"
+      "UPPAAL_Asm.steps (map_option strip \<circ>
+        (\<lambda>pc. conv_prog (If (pc < length prog) (prog ! pc)) None)) n
+        (a, aa, ab, ac, b) (pc', st', s', f', rs')"
+      "pc' \<in> steps_approx n prog a"
+      "pc' < length prog"
+      "(if pc < length prog then prog ! pc else None) = Some za"
+      "prog ! pc = Some x2"
+      "\<not> (\<exists>x\<in>case x2::int instrc of INSTR (JMPZ pc') \<Rightarrow> {pc + 1, pc'}
+          | INSTR CALL \<Rightarrow> {..<length prog} | INSTR instr.RETURN \<Rightarrow> {..<length prog}
+          | INSTR HALT \<Rightarrow> {} | INSTR _ \<Rightarrow> {pc + 1}
+          | CEXP ac \<Rightarrow> {pc + 1}. pc' \<in> steps_approx n prog x)"
+    have "pc' = pc"
+    proof (cases x2)
+      case (INSTR x1)
+      with A show "pc' = pc"
+        apply (cases x1)
+                         apply (solves \<open>auto elim!: UPPAAL_Asm.step.elims split: if_split_asm\<close>)+
+            apply (auto elim!: UPPAAL_Asm.step.elims split: if_split_asm)[]
+        subgoal for q _
+          apply (cases "q < length prog")
+           apply force
+          apply (drule steps_out_of_range; simp; fail)
+          done
+        subgoal for q _
+          apply (cases "q < length prog")
+           apply force
+          apply (drule steps_out_of_range; simp; fail)
+          done
+           apply (auto elim!: UPPAAL_Asm.step.elims split: if_split_asm)[]
+        subgoal for q
+          apply (cases "q + 1 < length prog")
+           apply force
+          apply (drule steps_out_of_range; simp)
+          done
+        subgoal for q
+          apply (cases "q + 1 < length prog")
+           apply force
+          apply (drule steps_out_of_range; simp)
+          done
+          apply (solves \<open>auto elim!: UPPAAL_Asm.step.elims split: if_split_asm\<close>)+
+        done
+    next
+      case (CEXP x2)
+      with A show ?thesis
+        using instr_id by (fastforce split: if_split_asm elim!: UPPAAL_Asm.step.elims)
+    qed
+  } note * = this
+  show ?thesis
+    using assms
+    apply (
+        induction P' n "(pc, st, s, f, rs)" "(pc', st', s', f', rs')"
+        arbitrary: pc st s f rs rule: steps.induct
+        )
+     apply (solves \<open>simp split: option.split\<close>)
+    apply (clarsimp, rule conjI)
+     apply (simp add: striptp_def split: if_split_asm; fail)
+    apply (clarsimp split: option.split, rule conjI)
+    apply (solves \<open>auto simp add: striptp_def split: if_split_asm\<close>)
+    apply safe
+    by (rule *)
+qed
+
+lemmas steps_steps_approx = steps_steps_approx'[unfolded P'_def]
+
+end (* End of context for fixed program *)
+
+
+context
+  fixes prog :: "int instrc option list"
+begin
+
+private abbreviation "P i \<equiv> if i < length prog then prog ! i else None"
+
+lemma stepsc_out_of_range:
+  assumes "stepsc (conv_prog P) n u (pc, st, s, f, rs) (pc', st', s', f', rs')" "pc \<ge> length prog"
+  shows "pc' = pc"
+  using assms by cases auto
+
+lemma stepsc_steps_approx:
+  assumes "stepsc (conv_prog P) n u (pc, st, s, f, rs) (pc', st', s', f', rs')" "pc' < length prog"
+  shows "pc' \<in> steps_approx n prog pc"
+proof -
+  {
+    fix u :: "nat \<Rightarrow> real"
+      and pc :: nat
+      and st :: "int list"
+      and m :: "int list"
+      and f :: bool
+      and rs :: "nat list"
+      and a :: nat
+      and aa :: "int list"
+      and ab :: "int list"
+      and ac :: bool
+      and b :: "nat list"
+      and n :: nat
+      and z :: "int instrc"
+    assume A: "stepc (map_instrc real_of_int z) u (pc, st, m, f, rs) = Some (a, aa, ab, ac, b)"
+      "stepsc (conv_prog P) n u (a, aa, ab, ac, b) (pc', st', s', f', rs')"
+      "pc' \<in> steps_approx n prog a"
+      "pc' < length prog"
+      "P pc = Some z"
+    have "\<forall>x2. prog ! pc = Some x2 \<longrightarrow> pc' = pc \<or>
+    (\<exists>x\<in>case x2 of INSTR (JMPZ pc') \<Rightarrow> {pc + 1, pc'} | INSTR CALL \<Rightarrow> {..<length prog}
+      | INSTR instr.RETURN \<Rightarrow> {..<length prog} | INSTR HALT \<Rightarrow> {} | INSTR _ \<Rightarrow> {pc + 1}
+      | CEXP ac \<Rightarrow> {pc + 1}. pc' \<in> steps_approx n prog x)"
+    proof(cases z)
+      case (INSTR x1)
+      with A show ?thesis
+        apply (simp split: option.split_asm if_split_asm)
+        apply (cases x1)
+                         apply (simp split: if_split_asm; fail)
+                       apply (solves \<open>auto elim!: UPPAAL_Asm.step.elims split: if_split_asm\<close>)+
+            apply (auto elim!: UPPAAL_Asm.step.elims split: if_split_asm)[]
+        subgoal for q
+          apply (cases "q < length prog")
+           apply force
+          apply (drule stepsc_out_of_range; simp)
+          done
+        subgoal for q
+          apply (cases "q < length prog")
+           apply force
+          apply (drule stepsc_out_of_range; simp)
+          done
+           apply (auto elim!: UPPAAL_Asm.step.elims split: if_split_asm)[]
+        subgoal for q
+          apply (cases "q + 1 < length prog")
+           apply force
+          apply (drule stepsc_out_of_range; simp)
+          done
+        subgoal for q
+          apply (cases "q + 1 < length prog")
+           apply force
+          apply (drule stepsc_out_of_range; simp)
+          done
+          apply (drule stepsc_out_of_range; simp)
+         apply (solves \<open>auto elim!: UPPAAL_Asm.step.elims split: if_split_asm\<close>)+
+        done
+    next
+      case (CEXP x2)
+      with A show ?thesis
+        by (force split: if_split_asm)
+    qed
+  } note * = this
+  show ?thesis
+    using assms
+    apply (
+        induction "conv_prog P" n u "(pc, st, s, f, rs)" "(pc', st', s', f', rs')"
+        arbitrary: pc st s f rs rule: stepsc.induct
+        )
+     apply (simp split: option.split)
+    apply clarsimp
+    apply (rule conjI)
+     apply (simp split: if_split_asm; fail)
+    apply (clarsimp split: option.split)
+    apply (rule conjI)
+     apply (simp split: if_split_asm; fail)
+    by (rule *)
+qed
+
+definition
+  "time_indep_check pc n \<equiv>
+    \<forall> pc' \<in> steps_approx n prog pc. pc' < length prog
+    \<longrightarrow> (case prog ! pc' of Some cmd \<Rightarrow> is_instr cmd | _ \<Rightarrow> True)"
+
+lemma time_indep_overapprox:
+  assumes
+    "time_indep_check pc n"
+  shows "time_indep (conv_prog P) n (pc, st, s, f, rs)"
+proof -
+  { fix pc' st' s' f' rs' cmd u
+    assume A:
+      "stepsc (conv_prog local.P) n u (pc, st, s, f, rs) (pc', st', s', f', rs')"
+      "conv_prog local.P pc' = Some cmd"
+    have "is_instr cmd"
+    proof (cases "pc' < length prog")
+      case True
+      with A(2) obtain cmd' where "prog ! pc' = Some cmd'" by auto
+      with True stepsc_steps_approx[OF A(1)] A(2) assms show ?thesis
+        by (cases cmd') (auto simp: time_indep_check_def)
+    next
+      case False
+      with A(2) show ?thesis by (auto split: option.split_asm)
+    qed
+  }
+  then show ?thesis unfolding time_indep_def by blast
+qed
+
+end (* End of context for fixed program *)
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  definition
+    "collect_cexp' pc = {ac. Some (CEXP ac) \<in> ((!) prog) ` steps_approx max_steps prog pc}"
+
+  definition "clkp_set'' i l \<equiv>
+    collect_clock_pairs (inv ! i ! l) \<union>
+    \<Union> ((\<lambda> (g, _). constraint_pair ` collect_cexp' g) ` set (trans ! i ! l))"
+
+  definition
+    "collect_cexp = {ac. Some (CEXP ac) \<in> set prog}"
+
+  definition
+    "collect_store' pc =
+    {(c, x). Some (INSTR (STOREC c x)) \<in> ((!) prog) ` steps_approx max_steps prog pc}"
+
+end
+
+lemma visited_resets_mono:
+  "set r \<subseteq> set r'" if "visited P n (pc, st, s, f, r) (pc', st', s', f', r') pcs"
+  using that
+  apply (
+    induction P \<equiv> P n "(pc, st, s, f, r :: nat list)" "(pc', st', s', f', r')" pcs
+    arbitrary: pc st s f r  rule: visited.induct
+    )
+   apply blast
+  by (erule step.elims; force split: option.splits if_splits elim!: step.elims)+
+
+lemma visitedc_resets_mono:
+  "set r \<subseteq> set r'" if "visitedc P n u (pc, st, s, f, r) (pc', st', s', f', r') pcs"
+  using that
+  apply (
+    induction P \<equiv> P n u "(pc, st, s, f, r :: nat list)" "(pc', st', s', f', r')" pcs
+    arbitrary: pc st s f r  rule: visitedc.induct
+    )
+   apply blast
+  by (erule stepc.elims; force split: option.splits if_splits elim!: step.elims)+
+
+lemma visited_reset:
+  "\<exists> x. \<exists> pc \<in> set pcs. Some (STOREC c x) = P pc"
+  if "visited P n (pc, st, s, f, r) (pc', st', s', f', r') pcs" "c \<in> set r' - set r"
+  using that
+  apply (
+      induction P \<equiv> P n "(pc, st, s, f, r :: nat list)" "(pc', st', s', f', r')" pcs
+      arbitrary: pc st s f r rule: visited.induct
+      )
+   apply blast
+  by (erule step.elims; force split: option.split_asm if_split_asm elim!: step.elims)
+
+lemma visitedc_reset:
+  "\<exists> x. \<exists> pc \<in> set pcs. Some (INSTR (STOREC c x)) = P pc"
+  if "visitedc P n u (pc, st, s, f, r) (pc', st', s', f', r') pcs" "c \<in> set r' - set r"
+  using that
+  apply (
+      induction P \<equiv> P n u "(pc, st, s, f, r :: nat list)" "(pc', st', s', f', r')" pcs
+      arbitrary: pc st s f r rule: visitedc.induct
+      )
+   apply blast
+  by (erule stepc.elims; force split: option.split_asm if_split_asm elim!: step.elims)
+
+lemma visited_fuel_mono:
+  "visited P n' s s' pcs" if "visited P n s s' pcs" "n' \<ge> n"
+  using that
+  apply (induction arbitrary: n')
+  subgoal for _ _ _ n'
+    by (cases n') (auto intro: visited.intros)
+  subgoal for _ _ _ _ _ _ _ _ _ _ _ n'
+    by (cases n') (auto intro: visited.intros)
+  done
+
+lemma visitedc_fuel_mono:
+  "visitedc P n' u s s' pcs" if "visitedc P n u s s' pcs" "n' \<ge> n"
+  using that
+  apply (induction arbitrary: n')
+  subgoal for _ _ _ _ n'
+    by (cases n') (auto intro: visitedc.intros)
+  subgoal for _ _ _ _ _ _ _ _ _ _ _ _ n'
+    by (cases n') (auto intro: visitedc.intros)
+  done
+
+lemma visted_split:
+  assumes "visited P n (pc, st, s, f, r) s'' (pcs' @ pc' # pcs)"
+  obtains st' s' f' r' where
+    "visited P n (pc, st, s, f, r) (pc', st', s', f', r') pcs"
+    "visited P n (pc', st', s', f', r') s'' (pcs' @ [pc'])"
+  apply atomize_elim
+  using assms
+  proof (induction _ _ _ _ "pcs' @ pc' # pcs" arbitrary: pcs rule: visited.induct)
+    case (1 prog n u start)
+    then show ?case by simp
+  next
+    case prems: (2 cmd pc st m f rs s prog n s' pcs pcs'')
+    show ?case
+    proof (cases "pcs'' = []")
+      case True
+      with prems show ?thesis by (blast intro: visited.intros)
+    next
+      case False
+      with \<open>pcs @ [pc] = _\<close> obtain pcs3 where "pcs'' = pcs3 @ [pc]"
+        by (metis append_butlast_last_id last_ConsR last_append last_snoc list.distinct(1))
+      with prems obtain st' s'a f' r' where
+        "visited prog n s (pc', st', s'a, f', r') pcs3"
+        "visited prog n (pc', st', s'a, f', r') s' (pcs' @ [pc'])"
+        by (auto 9 2)
+      with prems \<open>pcs'' = _\<close> show ?thesis by (auto 4 6 intro: visited.intros visited_fuel_mono)
+    qed
+  qed
+
+lemma visited_steps':
+  assumes "visited P n (pc, st, s, f, r) s'' pcs" "pc' \<in> set pcs"
+  obtains st' s' f' r' where "steps P n (pc, st, s, f, r) (pc', st', s', f', r')"
+  using assms by (force dest!: split_list dest: visited_steps elim: visted_split)
+
+lemma vistedc_split:
+  assumes "visitedc P n u (pc, st, s, f, r) s'' (pcs' @ pc' # pcs)"
+  obtains st' s' f' r' where
+    "visitedc P n u (pc, st, s, f, r) (pc', st', s', f', r') pcs"
+    "visitedc P n u (pc', st', s', f', r') s'' (pcs' @ [pc'])"
+  apply atomize_elim
+  using assms
+  proof (induction _ _ _ _ _ "pcs' @ pc' # pcs" arbitrary: pcs rule: visitedc.induct)
+    case (1 prog n u start)
+    then show ?case by simp
+  next
+    case prems: (2 cmd u pc st m f rs s prog n s' pcs pcs'')
+    show ?case
+    proof (cases "pcs'' = []")
+      case True
+      with prems show ?thesis by (blast intro: visitedc.intros)
+    next
+      case False
+      with \<open>pcs @ [pc] = _\<close> obtain pcs3 where "pcs'' = pcs3 @ [pc]"
+        by (metis append_butlast_last_id last_ConsR last_append last_snoc list.distinct(1))
+      with prems obtain st' s'a f' r' where
+        "visitedc prog n u s (pc', st', s'a, f', r') pcs3"
+        "visitedc prog n u (pc', st', s'a, f', r') s' (pcs' @ [pc'])"
+        by (auto 9 2)
+      with prems \<open>pcs'' = _\<close> show ?thesis by (auto 4 6 intro: visitedc.intros visitedc_fuel_mono)
+    qed
+  qed
+
+lemma visitedc_stepsc':
+  assumes "visitedc P n u (pc, st, s, f, r) s'' pcs" "pc' \<in> set pcs"
+  obtains st' s' f' r' where "stepsc P n u (pc, st, s, f, r) (pc', st', s', f', r')"
+  using assms  by (force dest!: split_list dest: visitedc_stepsc elim: vistedc_split)
+
+lemma steps_fuel_mono:
+  "steps P n' s s'" if "steps P n s s'" "n' \<ge> n"
+  using that
+  apply (induction arbitrary: n')
+  subgoal for _ _ _ n'
+    by (cases n') auto
+  subgoal for _ _ _ _ _ _ _ _ _ _ n'
+    by (cases n') auto
+  done
+
+lemma exec_steps':
+  assumes "exec prog n s pcs = Some (s', pcs')" "pc' \<in> set pcs' - set pcs"
+  obtains st m f rs where "steps prog n s (pc', st, m, f, rs)"
+  apply atomize_elim
+  using assms
+proof (induction P \<equiv> prog n s pcs arbitrary: s' rule: exec.induct)
+  case 1
+  then show ?case by simp
+next
+  case (2 n pc st m f rs pcs)
+  then obtain instr where "prog pc = Some instr" by (cases "prog pc") auto
+  show ?case
+  proof (cases "instr = HALT")
+    case True
+    with "2.prems" \<open>prog pc = _\<close> show ?thesis by (auto 2 6)
+  next
+    case F: False
+    show ?thesis
+    proof (cases "pc' = pc")
+      case True
+      with \<open>prog pc = _\<close> show ?thesis by (auto 2 5)
+    next
+      case False
+      with \<open>prog pc = _\<close> 2(2,3) F show ?thesis
+        by - (
+          erule exec.elims;
+          auto 5 6
+            dest!: 2(1)[OF \<open>prog pc = _\<close> F, rotated, OF sym]
+            simp: option.split_asm if_split_asm
+          )
+    qed
+  qed
+qed
+
+lemma exec_visited:
+  assumes "exec prog n s pcs = Some ((pc, st, m, f, rs), pcs')"
+  obtains pcs'' where
+    "visited prog n s (pc, st, m, f, rs) pcs'' \<and> pcs' = pc # pcs'' @ pcs \<and> prog pc = Some HALT"
+  apply atomize_elim
+  using assms proof (induction P \<equiv> prog n s pcs arbitrary: pc st m f rs rule: exec.induct)
+  case 1
+  then show ?case by simp
+next
+  case (2 n pc' st' m' f' rs' pcs')
+  then obtain instr where "prog pc' = Some instr" by (cases "prog pc'") auto
+  show ?case
+  proof (cases "instr = HALT")
+    case True
+    with "2.prems" \<open>prog pc' = _\<close> show ?thesis by (auto elim: exec.elims intro: visited.intros)
+  next
+    case False
+    with 2(2) \<open>prog pc' = _\<close> show ?thesis
+      by - (
+          erule exec.elims;
+          auto split: option.split_asm if_split_asm intro: visited.intros
+          dest!: 2(1)[OF \<open>prog pc' = _\<close> False, rotated, OF sym]
+          )
+  qed
+qed
+
+lemma exec_reset':
+  "\<exists> x. \<exists> pc \<in> set pcs'. Some (STOREC c x) = P pc"
+  if "exec P n (pc, st, s, f, r) pcs = Some ((pc', st', s', f', r'), pcs')" "c \<in> set r' - set r"
+proof -
+  from exec_visited[OF that(1)] obtain pcs'' where *:
+    "visited P n (pc, st, s, f, r) (pc', st', s', f', r') pcs''"
+    "pcs' = pc' # pcs'' @ pcs \<and> P pc' = Some HALT"
+    by auto
+  from visited_reset[OF this(1) that(2)] obtain x pc where
+    "pc\<in>set pcs''" "Some (STOREC c x) = P pc"
+    by auto
+  with *(2) show ?thesis by auto
+qed
+
+lemma steps_approx_out_of_range:
+  "steps_approx n prog pc = {}" if "pc \<ge> length prog"
+  using that by (induction n) auto
+
+lemma steps_resets_mono:
+  "set r \<subseteq> set r'" if "steps P n (pc, st, s, f, r) (pc', st', s', f', r')"
+  using that
+  by (induction "(pc, st, s, f, r)" "(pc', st', s', f', r')" arbitrary: pc st s f r;
+      fastforce intro: visited.intros elim!: step.elims split: if_split_asm)
+
+lemma resets_start:
+  assumes
+    "\<forall> pc \<in> {pc..pc'}. \<exists> c x. prog ! pc = Some (INSTR (STOREC c x))"
+    "steps
+      (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None))
+      n (pc, st, s, f, r) (pc_t, st', s', f', r')"
+    "prog ! pc_t = Some (INSTR HALT)"
+  shows "{c. \<exists> x. \<exists> pc \<in> {pc .. pc'}. prog ! pc = Some (INSTR (STOREC c x))} \<subseteq> set r'"
+    using assms(2,3,1)
+  proof (induction
+    "(map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None))" n "(pc, st, s, f, r)"
+    "(pc_t, st', s', f', r')" arbitrary: pc st s f r
+    )
+      case prems: 1
+      show ?case
+      proof (cases "pc' \<ge> pc_t")
+        case True
+        with prems obtain c x where "prog ! pc_t = Some (INSTR (STOREC c x))"
+          by fastforce
+        with prems show ?thesis by simp
+      next
+        case False
+        with prems show ?thesis by auto
+      qed
+    next
+      case prems: (2 cmd pc st m f rs s n)
+      show ?case
+      proof (cases "pc' \<ge> pc")
+        case True
+        with prems(6) obtain c d where "prog ! pc = Some (INSTR (STOREC c d))"
+          by fastforce
+        moreover obtain pc1 st' s' f' r1 where "s = (pc1, st', s', f', r1)"
+          using prod.exhaust by metis
+        ultimately have "pc1 = pc + 1"
+          using prems(1,2) by (auto elim!: step.elims split: if_split_asm)
+        with prems(4)[OF \<open>s = _\<close> prems(5)] prems(6) have
+          "{c. \<exists>x. \<exists>pc\<in>{pc1..pc'}. prog ! pc = Some (INSTR (STOREC c x))} \<subseteq> set r'"
+          by auto
+        moreover have "c \<in> set r1"
+          using prems(1,2) \<open>s = _\<close> \<open>prog ! pc = _\<close> by (auto elim!: step.elims split: if_split_asm)
+        moreover then have "c \<in> set r'"
+          using prems(3) \<open>s = _\<close> by (auto dest: steps_resets_mono)
+        ultimately show ?thesis
+          using \<open>pc1 = _\<close> \<open>prog ! pc = _\<close> apply clarsimp
+          subgoal for _ _ pc''
+            by (cases "pc'' = pc"; force)
+          done
+      next
+        case False
+        then show ?thesis by auto
+      qed
+    qed
+
+unbundle lattice_syntax
+
+function find_resets_start where
+  "find_resets_start prog pc =
+    (
+    if pc < length prog
+    then
+      case prog ! pc of
+        Some (INSTR (STOREC c x)) \<Rightarrow> (Some pc \<squnion> find_resets_start prog (pc + 1)) |
+        _ \<Rightarrow> None
+    else None
+    )
+  "
+  by auto
+
+termination
+  by (relation "measure (\<lambda> (prog, pc). length prog - pc)") auto
+
+lemma find_resets_start:
+  "\<forall> pc \<in> {pc..pc'}. \<exists> c x. prog ! pc = Some (INSTR (STOREC c x))" if
+  "find_resets_start prog pc = Some pc'"
+  using that
+proof (induction arbitrary: pc' rule: find_resets_start.induct)
+  case prems: (1 prog pc)
+  from prems(2) show ?case
+    apply (simp split: if_split_asm)
+    apply (
+        auto 4 3 simp flip: not_less_eq_eq
+        simp del: find_resets_start.simps simp: le_max_iff_disj sup_nat_def sup_option_def
+        split: option.split_asm instr.split_asm instrc.split_asm dest!: prems(1)
+        )
+    done
+qed
+
+lemmas resets_start' = resets_start[OF find_resets_start]
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  definition
+    "collect_store'' pc \<equiv>
+    case find_resets_start prog pc of
+      None \<Rightarrow> {} |
+      Some pc' \<Rightarrow>
+        {(c, x). Some (INSTR (STOREC c x)) \<in> ((!) prog) ` {pc .. pc'}}"
+
+end
+
+lemma bexp_atLeastAtMost_iff:
+  "(\<forall> pc \<in> {pc_s..pc_t}. P pc) \<longleftrightarrow> (\<forall> pc. pc_s \<le> pc \<and> pc \<le> pc_t \<longrightarrow> P pc)"
+  by auto
+
+lemma bexp_atLeastLessThan_iff:
+  "(\<forall> pc \<in> {pc_s..<pc_t}. P pc) \<longleftrightarrow> (\<forall> pc. pc_s \<le> pc \<and> pc < pc_t \<longrightarrow> P pc)"
+  by auto
+
+lemma guaranteed_execution:
+  assumes
+    "\<forall> pc \<in> {pc..<pc_t}.
+      prog ! pc \<noteq> None
+      \<and> prog ! pc \<notin> Some ` INSTR `
+        {STORE, HALT, POP, CALL, RETURN, instr.AND, instr.NOT, instr.ADD, instr.LT, instr.LE, instr.EQ}
+      \<and> (\<forall> c d. prog ! pc = Some (INSTR (STOREC c d)) \<longrightarrow> d = 0)
+      "
+    "\<forall> pc \<in> {pc..<pc_t}. \<forall> pc'. prog ! pc = Some (INSTR (JMPZ pc')) \<longrightarrow> pc' > pc \<and> pc' \<le> pc_t"
+    "prog ! pc_t = Some (INSTR HALT)" "pc_t \<ge> pc" "n > pc_t - pc" "pc_t < length prog"
+  shows "\<exists> st' s' f' r'. steps
+      (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None))
+      n (pc, st, s, f, r) (pc_t, st', s', f', r')"
+  using assms
+proof (induction "pc_t - pc" arbitrary: pc st s f r n rule: less_induct)
+  case less
+  let ?prog = "(map_option stripf \<circ> (\<lambda>pc. if pc < length prog then prog ! pc else None))"
+  from \<open>pc_t - pc < n\<close> obtain n' where [simp]: "n = Suc n'" by (cases n) auto
+  from less.prems(3-) have [simp]: "pc < length prog" "pc_t < length prog" by auto
+  show ?case
+  proof (cases "pc_t = pc")
+    case True
+    then show ?thesis by force
+  next
+    case False
+    with less.prems(1,4) have valid_instr:
+      "prog ! pc \<notin> Some ` INSTR `
+       {STORE, HALT, POP, CALL, RETURN, instr.AND, instr.NOT, instr.ADD, instr.LT, instr.LE, instr.EQ}"
+      "prog ! pc \<noteq> None"
+      by auto
+    from \<open>pc \<le> _\<close> \<open>pc_t \<noteq> _\<close> less.prems(2) have jumps:
+      "pc' > pc \<and> pc' \<le> pc_t" if "prog ! pc = Some (INSTR (JMPZ pc'))" for pc'
+      using that by fastforce
+    from \<open>pc \<le> _\<close> \<open>pc_t \<noteq> _\<close> less.prems(1) have stores:
+      "d = 0" if "prog ! pc = Some (INSTR (STOREC c d))" for c d
+      using that by fastforce
+    show ?thesis
+    proof (cases "prog ! pc")
+      case None
+      from \<open>pc_t \<noteq> _\<close> less.prems(1,4) have "prog ! pc \<noteq> None"
+        by simp
+      with None show ?thesis by auto
+    next
+      case (Some a)
+      then show ?thesis
+      proof (cases a)
+        case (INSTR instr)
+        have *:
+          "\<exists>st' s' f' r'. steps ?prog n' (pc', st, s, f, r) (pc_t, st', s', f', r')"
+          if "pc < pc'" "pc' \<le> pc_t" for pc' st s f r
+          apply (rule less.hyps[of pc'])
+          subgoal
+            using that by simp
+          subgoal
+            using that less.prems(1) by force
+          subgoal
+            apply clarsimp
+            subgoal premises prems for pc_s pc'
+            proof -
+              from prems that have "pc_s \<in> {pc..<pc_t}" by simp
+              with prems(1) less.prems(2) show ?thesis by blast
+            qed
+            done
+          using \<open>prog ! pc_t = _\<close> \<open>pc_t - pc < n\<close> that by auto
+        from \<open>pc_t \<noteq> pc\<close> \<open>pc \<le> _\<close> have "Suc pc \<le> pc_t" by simp
+        then obtain pc' st' s' f' r' where
+          "step instr (pc, st, s, f, r) = Some (pc', st', s', f', r')" "pc < pc'" "pc' \<le> pc_t"
+          apply atomize_elim
+          apply (cases instr)
+          using valid_instr \<open>a = _\<close> \<open>_ = Some a\<close> by (auto simp: stores jumps)
+        with \<open>a = _\<close> \<open>_ = Some a\<close> show ?thesis by (force dest!: *)
+      next
+        case (CEXP x2)
+        have *:
+          "\<exists>st' s' f' r'. steps ?prog n' (Suc pc, st, s, f, r) (pc_t, st', s', f', r')"
+          if "Suc pc \<le> pc_t" for st s f r
+          apply (rule less.hyps[of "Suc pc"])
+          subgoal
+            using \<open>pc \<le> _\<close> \<open>pc_t \<noteq> pc\<close> by simp
+          subgoal
+            using less.prems(1) by force
+          subgoal
+            apply clarsimp
+            subgoal premises prems for pc_s pc'
+            proof -
+              from prems have "pc_s \<in> {pc..<pc_t}" by simp
+              with prems(1) less.prems(2) show ?thesis by blast
+            qed
+            done
+          using \<open>prog ! pc_t = _\<close> \<open>pc_t - pc < n\<close> that by auto
+        from \<open>pc_t \<noteq> pc\<close> \<open>pc \<le> _\<close> have "Suc pc \<le> pc_t" by simp
+        with \<open>a = _\<close> \<open>_ = Some a\<close> show ?thesis by (force dest!: *)
+      qed
+    qed
+  qed
+qed
+
+function find_next_halt where
+  "find_next_halt prog pc =
+    (
+    if pc < length prog
+    then
+      case prog ! pc of
+        Some (INSTR HALT) \<Rightarrow> Some pc |
+        _ \<Rightarrow> find_next_halt prog (pc + 1)
+    else None
+    )
+  "
+  by auto
+
+termination
+  by (relation "measure (\<lambda> (prog, pc). length prog - pc)") auto
+
+lemma find_next_halt_finds_halt:
+  "prog ! pc' = Some (INSTR HALT) \<and> pc \<le> pc' \<and> pc' < length prog"
+  if "find_next_halt prog pc = Some pc'"
+using that proof (induction prog pc rule: find_next_halt.induct)
+  case prems: (1 prog pc)
+  from prems(20) show ?case
+    by (
+        simp,
+        simp
+        split: if_split_asm option.split_asm instrc.split_asm instr.split_asm
+        del: find_next_halt.simps;
+        fastforce dest: prems(1-19) simp del: find_next_halt.simps)
+qed
+
+definition
+  "guaranteed_execution_cond prog pc_s n \<equiv>
+    case find_next_halt prog pc_s of
+      None \<Rightarrow> False |
+      Some pc_t \<Rightarrow>
+       (
+       \<forall> pc \<in> {pc_s..<pc_t}.
+          prog ! pc \<noteq> None
+          \<and> prog ! pc \<notin> Some ` INSTR `
+            {STORE, HALT, POP, CALL, RETURN, instr.AND, instr.NOT, instr.ADD, instr.LT, instr.LE, instr.EQ}
+          \<and> (\<forall> c d. prog ! pc = Some (INSTR (STOREC c d)) \<longrightarrow> d = 0)
+        ) \<and>
+        (\<forall> pc \<in> {pc_s..<pc_t}. \<forall> pc'. prog ! pc = Some (INSTR (JMPZ pc')) \<longrightarrow> pc' > pc \<and> pc' \<le> pc_t)
+       \<and> n > pc_t - pc_s
+  "
+
+lemma guaranteed_execution_cond_alt_def[code]:
+  "guaranteed_execution_cond prog pc_s n \<equiv>
+    case find_next_halt prog pc_s of
+      None \<Rightarrow> False |
+      Some pc_t \<Rightarrow>
+       (
+       \<forall> pc \<in> {pc_s..<pc_t}.
+          prog ! pc \<noteq> None
+          \<and> prog ! pc \<notin> Some ` INSTR `
+            {STORE, HALT, POP, CALL, RETURN, instr.AND, instr.NOT, instr.ADD, instr.LT, instr.LE, instr.EQ}
+          \<and> (case prog ! pc of Some (INSTR (STOREC c d)) \<Rightarrow> d = 0 | _ \<Rightarrow> True)
+        ) \<and>
+        (\<forall> pc \<in> {pc_s..<pc_t}.
+          case prog ! pc of Some (INSTR (JMPZ pc')) \<Rightarrow> pc' > pc \<and> pc' \<le> pc_t | _ \<Rightarrow> True)
+       \<and> n > pc_t - pc_s
+  "
+proof (rule eq_reflection, goal_cases)
+  case 1
+  have *:
+    "(\<forall> c d. prog ! pc = Some (INSTR (STOREC c d)) \<longrightarrow> d = 0) \<longleftrightarrow>
+     (case prog ! pc of Some (INSTR (STOREC c d)) \<Rightarrow> d = 0 | _ \<Rightarrow> True)" for pc
+    by (auto split: option.split instrc.split instr.split)
+  have **:
+    "(\<forall> pc'. prog ! pc = Some (INSTR (JMPZ pc')) \<longrightarrow> pc' > pc \<and> pc' \<le> pc_t) \<longleftrightarrow>
+     (case prog ! pc of Some (INSTR (JMPZ pc')) \<Rightarrow> pc' > pc \<and> pc' \<le> pc_t | _ \<Rightarrow> True)" for pc pc_t
+    by (auto split: option.split instrc.split instr.split)
+  show ?case unfolding guaranteed_execution_cond_def * ** ..
+qed
+
+lemma guaranteed_execution':
+  "\<exists> pc_t st' s' f' r' pcs'. exec
+      (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None))
+      n (pc, st, s, f, r) pcs = Some ((pc_t, st', s', f', r'), pcs')"
+  if "guaranteed_execution_cond prog pc n"
+proof -
+  from that obtain pc_t where "find_next_halt prog pc = Some pc_t"
+    unfolding guaranteed_execution_cond_def bexp_atLeastAtMost_iff
+    by (auto split: option.split_asm)
+  then have "prog ! pc_t = Some (INSTR HALT) \<and> pc \<le> pc_t \<and> pc_t < length prog"
+    by (rule find_next_halt_finds_halt)
+  moreover then have "\<exists> st' s' f' r'. steps
+      (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None))
+      n (pc, st, s, f, r) (pc_t, st', s', f', r')"
+    using \<open>_ = Some pc_t\<close> that
+    unfolding guaranteed_execution_cond_def bexp_atLeastAtMost_iff
+    by - (rule guaranteed_execution, auto)
+  ultimately show ?thesis by (force dest: steps_exec)
+  qed
+
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  lemma collect_cexp_alt_def:
+    "collect_cexp =
+      set (List.map_filter
+        (\<lambda> x. case x of Some (CEXP ac) \<Rightarrow> Some ac | _ \<Rightarrow> None)
+         prog)"
+    unfolding collect_cexp_def set_map_filter by (auto split: option.split_asm instrc.split_asm)
+
+  lemma clkp_set'_alt_def:
+    "clkp_set' =
+      \<Union> (collect_clock_pairs ` set (concat inv)) \<union> (constraint_pair ` collect_cexp)"
+    unfolding clkp_set'_def collect_cexp_def by auto
+
+  definition
+    "collect_store = {(c, x). Some (INSTR (STOREC c x)) \<in> set prog}"
+
+  lemma collect_store_alt_def:
+    "collect_store =
+      set (List.map_filter
+        (\<lambda> x. case x of Some (INSTR (STOREC c x)) \<Rightarrow> Some (c, x) | _ \<Rightarrow> None)
+         prog)"
+    unfolding collect_store_def set_map_filter
+    by (auto split: option.split_asm instrc.split_asm instr.split_asm)
+
+  lemma clk_set'_alt_def: "clk_set' = (fst ` clkp_set' \<union> fst ` collect_store)"
+    unfolding clk_set'_def collect_store_def by auto
+
+end
+
+
+fun conj_instr :: "'t instrc \<Rightarrow> addr \<Rightarrow> bool" where
+  "conj_instr (CEXP _) _ = True" |
+  "conj_instr (INSTR COPY) _ = True" |
+  "conj_instr (INSTR (JMPZ pc)) pc_t = (pc = pc_t)" |
+  "conj_instr (INSTR instr.AND) _ = True" |
+  "conj_instr _ _ = False"
+
+(*
+inductive is_conj_block :: "'t instrc option list \<Rightarrow> addr \<Rightarrow> addr \<Rightarrow> bool" where
+  (* "is_conj_block prog pc pc" if "prog ! pc = Some (INSTR HALT)" | *)
+  "is_conj_block prog pc (pc' + 1)" if
+  "prog ! (pc' + 1) = Some (INSTR HALT)" "is_conj_block prog pc pc'" |
+  "is_conj_block prog pc pc" if "prog ! pc = Some (CEXP ac)" |
+  "is_conj_block prog pc (pc' + 3)" if
+  "prog ! (pc' + 1) = Some (INSTR COPY)" "prog ! (pc' + 2) = Some (CEXP ac)"
+  "prog ! (pc' + 3) = Some (INSTR instr.AND)"
+  "is_conj_block prog pc pc'" |
+  "is_conj_block prog pc (pc' + 4)" if
+  "prog ! (pc' + 1) = Some (INSTR COPY)"
+  "prog ! (pc' + 2) = Some (INSTR (JMPZ pc_t))" "prog ! pc_t = Some (INSTR HALT)"
+  "prog ! (pc' + 3) = Some (CEXP ac)" "prog ! (pc' + 4) = Some (INSTR instr.AND)"
+  "is_conj_block prog pc pc'"
+*)
+
+inductive is_conj_block' :: "'t instrc option list \<Rightarrow> addr \<Rightarrow> addr \<Rightarrow> bool" where
+  "is_conj_block' prog pc pc" if
+  "pc < length prog"
+  "prog ! pc = Some (INSTR HALT)" |
+  (*
+  "is_conj_block' prog pc (pc + 2)" if
+  "pc + 2 < length prog"
+  "prog ! (pc) = Some (INSTR COPY)" "prog ! (pc + 1) = Some (CEXP ac)"
+  "prog ! (pc + 2) = Some (INSTR instr.AND)" |
+  *)
+  (*
+  "is_conj_block' prog pc (pc' + 2)" if
+  "pc' + 2 < length prog"
+  "prog ! (pc') = Some (INSTR COPY)" "prog ! (pc' + 1) = Some (CEXP ac)"
+  "prog ! (pc' + 2) = Some (INSTR instr.AND)"
+  "is_conj_block' prog pc pc'"
+  *)
+  "is_conj_block' prog pc pc'" if
+  "pc' < length prog"
+  "prog ! pc = Some (INSTR COPY)" "prog ! (pc + 1) = Some (CEXP ac)"
+  "prog ! (pc + 2) = Some (INSTR instr.AND)"
+  "is_conj_block' prog (pc + 3) pc'" |
+  "is_conj_block' prog pc pc'" if
+  "pc' < length prog"
+  "prog ! pc = Some (INSTR COPY)"
+  "prog ! (pc + 1) = Some (INSTR (JMPZ pc'))" (* "prog ! pc_t = Some (INSTR HALT)" *)
+  "prog ! (pc + 2) = Some (CEXP ac)"
+  "prog ! (pc + 3) = Some (INSTR instr.AND)"
+  "is_conj_block' prog (pc + 4) pc'"
+
+inductive_cases stepscE: "stepsc prog n u (pc, st, m, f, rs) (pc', st', m', f', rs')"
+
+function check_conj_block' :: "'t instrc option list \<Rightarrow> addr \<Rightarrow> addr option" where
+  "check_conj_block' prog pc = (
+    if pc \<ge> length prog then None
+    else if prog ! pc = Some (INSTR HALT) then Some pc
+    else if
+      prog ! pc = Some (INSTR COPY) \<and> (case prog ! (pc + 1) of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False)
+      \<and> prog ! (pc + 2) = Some (INSTR instr.AND)
+    then check_conj_block' prog (pc + 3)
+    else if
+      prog ! pc = Some (INSTR COPY) \<and> (case prog ! (pc + 2) of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False)
+      \<and> prog ! (pc + 3) = Some (INSTR instr.AND)
+      \<and> (case prog ! (pc + 1) of Some (INSTR (JMPZ pc')) \<Rightarrow> True | _ \<Rightarrow> False)
+    then
+      (case (prog ! (pc + 1), check_conj_block' prog (pc + 4)) of (Some (INSTR (JMPZ pc')), Some pc'')
+        \<Rightarrow> if pc' = pc'' then Some pc' else None | _ \<Rightarrow> None)
+    else None
+    )
+  "
+  by pat_completeness auto
+
+termination
+  by (relation "measure (\<lambda> (prog, pc). length prog - pc)") auto
+
+lemma is_conj_block'_len_prog:
+  "pc' < length prog" if "is_conj_block' prog pc pc'"
+  using that by induction auto
+
+lemma check_conj_block':
+  "check_conj_block' prog pc = Some pc' \<Longrightarrow> is_conj_block' prog pc pc'"
+  apply (induction prog pc rule: check_conj_block'.induct)
+    apply (erule check_conj_block'.elims)
+  by (auto
+        split: if_split_asm option.split_asm instrc.split_asm instr.split_asm
+        simp del: check_conj_block'.simps
+        intro: is_conj_block'.intros is_conj_block'_len_prog)
+
+lemma stepsc_reverseE':
+  assumes "stepsc prog (Suc n) u s s''" "s'' \<noteq> s"
+  obtains pc' st' m' f' rs' cmd where
+    "stepc cmd u (pc', st', m', f', rs') = Some s''"
+    "prog pc' = Some cmd"
+    "stepsc prog n u s (pc', st', m', f', rs')"
+  apply atomize_elim
+  using assms
+proof (induction prog "Suc n" u s s'' arbitrary: n rule: stepsc.induct)
+  case (1 prog n u start)
+  then show ?case by simp
+next
+  case prems: (2 cmd u pc st m f rs s prog n s')
+  then show ?case
+  proof (cases "s' = s")
+    case True
+    with prems show ?thesis
+      apply simp
+      apply solve_ex_triv+
+      by (auto elim: stepsc.cases)
+  next
+    case False
+    with prems(3) have "n > 0" by (auto elim!: stepsc.cases)
+    then obtain n' where "n = Suc n'" by (cases n) auto
+    from prems(4)[OF this False] obtain cmd pc' st' m' f' rs' where
+      "stepc cmd u (pc', st', m', f', rs') = Some s'" "prog pc' = Some cmd"
+      "stepsc prog n' u s (pc', st', m', f', rs')"
+      by atomize_elim
+    with prems show ?thesis unfolding \<open>n = _\<close> by blast
+  qed
+qed
+
+lemma stepsc_reverseE:
+  assumes "stepsc prog n u s s''" "s'' \<noteq> s"
+  obtains n' pc' st' m' f' rs' cmd where
+    "n = Suc n'"
+    "stepc cmd u (pc', st', m', f', rs') = Some s''"
+    "prog pc' = Some cmd"
+    "stepsc prog n' u s (pc', st', m', f', rs')"
+proof -
+  from assms have "n > 0" by (auto elim!: stepsc.cases)
+  then obtain n' where "n = Suc n'" by (cases n) auto
+  with assms show ?thesis by (auto elim!: stepsc_reverseE' intro!: that)
+qed
+
+lemma
+  "pc = pc' - 1" if
+  "stepc (CEXP cc) u (pc, st, m, f, rs) = Some (pc', st', m', f', rs')"
+  using that by auto
+
+lemma
+  "pc = pc' - 1" if
+  "stepc (INSTR instr) u (pc, st, m, f, rs) = Some (pc', st', m', f', rs')"
+  "\<not> (\<exists> x. instr = JMPZ pc')" "instr \<noteq> RETURN" "instr \<noteq> CALL"
+  using that by (auto split: option.split_asm if_split_asm elim!: step.elims)
+
+lemma stepc_pc_no_jump:
+  "pc = pc' - 1" if
+  "stepc cmd u (pc, st, m, f, rs) = Some (pc', st', m', f', rs')"
+  "cmd \<noteq> INSTR (JMPZ pc')" "cmd \<noteq> INSTR RETURN" "cmd \<noteq> INSTR CALL"
+  using that by (cases cmd) (auto split: option.split_asm if_split_asm elim!: step.elims)
+
+inductive stepsn :: "'t programc \<Rightarrow> nat \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state \<Rightarrow> bool"
+  for prog where
+  "stepsn prog 0 u start start" | (*
+  "stepsc prog (Suc n) u s s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s'"
+    "stepsc prog n u s (pc, st, m, f, rs)"
+    "prog pc = Some cmd" *)
+  "stepsn prog (Suc n) u (pc, st, m, f, rs) s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s"
+    "prog pc = Some cmd"
+    "stepsn prog n u s s'"
+
+declare stepsn.intros[intro]
+
+lemma stepsc_stepsn:
+  assumes "stepsc P n u s s'"
+  obtains n' where "stepsn P n' u s s'" "n' < n"
+  using assms by induction auto
+
+lemma stepsn_stepsc:
+  assumes "stepsn P n' u s s'" "n' < n"
+  shows "stepsc P n u s s'"
+  using assms
+  proof (induction arbitrary: n)
+    case (1 u start)
+    then obtain n' where "n = Suc n'" by (cases n) auto
+    then show ?case by auto
+  next
+    case (2 cmd u pc st m f rs s n s' n')
+    from \<open>_ < n'\<close> have "n < n' - 1" by simp
+    from 2(4)[OF this] 2(1,2,3,5) show ?case
+    proof -
+      have "(\<exists>fa n fb p. P = fa \<and> n' = Suc n \<and> u = fb \<and> (pc, st, m, f, rs) = p \<and> s' = p) \<or> (\<exists>i fa n is isa b ns p fb na pa. P = fb \<and> n' = Suc na \<and> u = fa \<and> (pc, st, m, f, rs) = (n, is, isa, b, ns) \<and> s' = pa \<and> stepc i fa (n, is, isa, b, ns) = Some p \<and> fb n = Some i \<and> stepsc fb na fa p pa)"
+        using "2.prems" Suc_pred' \<open>P pc = Some cmd\<close> \<open>stepc cmd u (pc, st, m, f, rs) = Some s\<close> \<open>stepsc P (n' - 1) u s s'\<close> gr_implies_not_zero by blast
+      then show ?thesis
+        by blast
+    qed
+  qed
+
+lemma stepsn_extend:
+  assumes "stepsn P n1 u s s1" "stepsn P n2 u s s2" "n1 \<le> n2"
+  shows "stepsn P (n2 - n1) u s1 s2"
+using assms
+proof (induction arbitrary: n2 s2)
+  case (1 u start)
+  then show ?case by simp
+next
+  case (2 cmd u pc st m f rs s n s')
+  from 2(1,2,5-) have "stepsn P (n2 - 1) u s s2" by (auto elim: stepsn.cases)
+  from 2(4)[OF this] \<open>Suc n <= _\<close> show ?case by simp
+qed
+
+lemma stepsc_halt:
+  "s' = (pc, s)" if "stepsc P n u (pc, s) s'" "P pc = Some (INSTR HALT)"
+  using that by (induction P n u "(pc, s)" s') auto
+
+lemma stepsn_halt:
+  "s' = (pc, s)" if "stepsn P n u (pc, s) s'" "P pc = Some (INSTR HALT)"
+  apply (rule stepsc_halt[OF stepsn_stepsc[where n = "Suc n"]])
+  using that by simp+
+
+lemma is_conj_block'_pc_mono:
+  "pc \<le> pc'" if "is_conj_block' prog pc pc'"
+  using that by induction auto
+
+lemma is_conj_block'_halt:
+  "prog ! pc' = Some (INSTR HALT)" if "is_conj_block' prog pc pc'"
+  using that by induction auto
+
+lemma numeral_4_eq_4:
+  "4 = Suc (Suc (Suc (Suc 0)))"
+  by simp
+
+lemma is_conj_block'_is_conj:
+  assumes "is_conj_block' P pc pc'"
+    and "stepsn (\<lambda> i. if i < length P then P ! i else None) n u (pc, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+    and "P ! pc_t = Some (INSTR HALT)"
+    (* and "pc < pc_t" (* "pc_t \<le> pc'" *) *)
+  shows "f \<and> pc_t = pc'"
+  using assms
+proof (induction arbitrary: n st s f rs)
+  case (1 pc prog)
+  with stepsn_halt[OF this(3)] show ?case by simp
+next
+  case prems: (2 pc' prog pc ac)
+  let ?P = "(\<lambda>i. if i < length prog then prog ! i else None)"
+  consider (0) "n = 0" | (1) "n = 1" | (2) "n = 2" | (3) "n \<ge> 3" by force
+  then show ?case
+  proof cases
+    case 0
+    with prems show ?thesis by (auto elim!: stepsn.cases)
+  next
+    case 1
+    with prems show ?thesis by (auto elim!: stepsn.cases simp: int_of_def split: if_split_asm)
+  next
+    case 2
+    with prems show ?thesis by (auto elim!: stepsn.cases simp: int_of_def numeral_2_eq_2 split: if_split_asm)
+  next
+    case 3
+    from prems have "pc + 2 < length prog" by (auto dest: is_conj_block'_pc_mono)
+    with prems(2-4) obtain st1 s1 rs1 where
+      "stepsn ?P 3 u (pc, st, s, f, rs) (pc + 3, st1, s1, f \<and> (u \<turnstile>a ac), rs1)"
+      by (force simp: int_of_def numeral_3_eq_3)
+    from stepsn_extend[OF this prems(7) 3] have
+      "stepsn ?P (n - 3) u (pc + 3, st1, s1, f \<and> (u \<turnstile>a ac), rs1) (pc_t, st_t, s_t, True, rs_t)" .
+    from prems(6)[OF this \<open>prog ! pc_t = _\<close>] have "pc_t = pc'" "u \<turnstile>a ac" f by auto
+    then show ?thesis by simp
+  qed
+next
+  case prems: (3 pc' prog pc ac)
+  let ?P = "(\<lambda>i. if i < length prog then prog ! i else None)"
+  consider (0) "n = 0" | (1) "n = 1" | (2) "n = 2" | (3) "n = 3" | (4) "n \<ge> 4" by force
+  then show ?case
+  proof cases
+    case 0
+    with prems show ?thesis by (auto elim!: stepsn.cases)
+  next
+    case 1
+    with prems show ?thesis by (auto elim!: stepsn.cases simp: int_of_def split: if_split_asm)
+  next
+    case 2
+    with prems show ?thesis by (auto elim!: stepsn.cases simp: int_of_def numeral_2_eq_2 split: if_split_asm)
+  next
+    case 3
+    with prems show ?thesis
+      by (auto elim!: stepsn.cases simp: int_of_def numeral_3_eq_3 split: if_split_asm dest!: is_conj_block'_halt)
+  next
+    case 4
+    from prems have "pc + 3 < length prog" by (auto dest: is_conj_block'_pc_mono)
+    show ?thesis
+    proof (cases f)
+      case True
+      with \<open>pc + 3 < _\<close> prems(2-6) obtain st1 s1 rs1 where
+        "stepsn ?P 4 u (pc, st, s, f, rs) (pc + 4, st1, s1, f \<and> (u \<turnstile>a ac), rs1)"
+      by (force simp: int_of_def numeral_3_eq_3 numeral_4_eq_4)
+      from stepsn_extend[OF this prems(8) 4] have
+        "stepsn ?P (n - 4) u (pc + 4, st1, s1, f \<and> (u \<turnstile>a ac), rs1) (pc_t, st_t, s_t, True, rs_t)" .
+      from prems(7)[OF this \<open>prog ! pc_t = _\<close>] have "pc_t = pc'" "u \<turnstile>a ac" f by auto
+      then show ?thesis by simp
+    next
+      case False
+      with \<open>pc + 3 < _\<close> prems(1-6) obtain st1 s1 rs1 where
+        "stepsn ?P 2 u (pc, st, s, f, rs) (pc', st1, s1, False, rs1)"
+        by (force simp: int_of_def numeral_2_eq_2)
+      from stepsn_extend[OF this prems(8)] 4 have
+        "stepsn ?P (n - 2) u (pc', st1, s1, False, rs1) (pc_t, st_t, s_t, True, rs_t)" by simp
+      from stepsn_halt[OF this] prems(1,6) show ?thesis by (auto dest: is_conj_block'_halt)
+    qed
+  qed
+qed
+
+lemma is_conj_block'_is_conj':
+  assumes "is_conj_block' P pc pc'"
+    and "stepst (\<lambda> i. if i < length P then P ! i else None) n u (pc, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+    (* and "pc < pc_t" (* "pc_t \<le> pc'" *) *)
+  shows "f \<and> pc_t = pc'"
+  using assms
+  unfolding stepst_def by (auto split: if_split_asm elim!: is_conj_block'_is_conj stepsc_stepsn)
+
+lemma is_conj_block'_is_conj2:
+  assumes "is_conj_block' P (pc + 1) pc'" "P ! pc = Some (CEXP ac)"
+    and "stepst (\<lambda> i. if i < length P then P ! i else None) n u (pc, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+  shows "(u \<turnstile>a ac) \<and> pc_t = pc'"
+proof -
+  let ?P = "(\<lambda> i. if i < length P then P ! i else None)"
+  from assms(1) have "pc < pc'" "pc' < length P"
+    by (auto dest: is_conj_block'_pc_mono is_conj_block'_len_prog)
+  with assms(2,3) obtain st' s' rs' where
+    "stepst ?P (n - 1) u (pc + 1, st', s', u \<turnstile>a ac, rs') (pc_t, st_t, s_t, True, rs_t)"
+    unfolding stepst_def
+    apply (clarsimp split: if_split_asm)
+    apply (erule stepsc.cases)
+    by auto
+  from is_conj_block'_is_conj'[OF assms(1) this] show ?thesis .
+qed
+
+lemma is_conj_block'_is_conj3:
+  assumes "is_conj_block' P (pc + 2) pc'" "P ! pc = Some (CEXP ac)" "P ! (pc + 1) = Some (INSTR instr.AND)"
+    and "stepst (\<lambda> i. if i < length P then P ! i else None) n u (pc, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+  shows "(u \<turnstile>a ac) \<and> pc_t = pc'"
+proof -
+  let ?P = "(\<lambda> i. if i < length P then P ! i else None)"
+  from assms(1) have "pc < pc'" "pc' < length P"
+    by (auto dest: is_conj_block'_pc_mono is_conj_block'_len_prog)
+  with assms(2,3,4) obtain st' s' rs' f where
+    "stepst ?P (n - 2) u (pc + 2, st', s', f \<and> (u \<turnstile>a ac), rs') (pc_t, st_t, s_t, True, rs_t)"
+    unfolding stepst_def
+    apply (clarsimp split: if_split_asm)
+    apply (erule stepsc.cases)
+     apply force
+    apply (erule stepsc.cases)
+    by (auto split: if_split_asm option.split_asm elim!: UPPAAL_Asm.step.elims)
+  from is_conj_block'_is_conj'[OF assms(1) this] show ?thesis by simp
+qed
+
+definition
+  "is_conj_block P pc pc' \<equiv>
+   (\<exists> ac. P ! pc = Some (CEXP ac)) \<and> is_conj_block' P (pc + 1) pc'
+   \<or> (\<exists> ac.
+      P ! pc = Some (CEXP ac)) \<and> P ! (pc + 1) = Some (INSTR instr.AND)
+      \<and> is_conj_block' P (pc + 2) pc'"
+
+lemma is_conj_block_alt_def[code]:
+"is_conj_block P pc pc' \<equiv>
+   (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False) \<and> is_conj_block' P (pc + 1) pc'
+   \<or> (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False) \<and> P ! (pc + 1) = Some (INSTR instr.AND)
+      \<and> is_conj_block' P (pc + 2) pc'"
+  unfolding is_conj_block_def
+  by (rule eq_reflection) (auto split: option.split_asm instrc.split_asm)
+
+lemma is_conj_block_is_conj:
+  assumes "is_conj_block P pc pc'" "P ! pc = Some (CEXP ac)"
+    and
+      "stepst
+        (\<lambda> i. if i < length P then P ! i else None) n u
+        (pc, st, s, f, rs)
+        (pc_t, st_t, s_t, True, rs_t)"
+  shows "(u \<turnstile>a ac) \<and> pc_t = pc'"
+  using assms
+  unfolding is_conj_block_def
+  apply safe
+  by ((drule is_conj_block'_is_conj2 is_conj_block'_is_conj3; simp), simp)+
+
+lemma is_conj_block'_decomp:
+  "is_conj_block P pc' pc''" if
+  "is_conj_block' P pc pc''" "P ! pc' = Some (CEXP ac)" "pc \<le> pc'" "pc' \<le> pc''"
+  using that
+proof (induction arbitrary: pc')
+  case (1 pc prog)
+  then show ?case by (simp add: is_conj_block_def)
+next
+  case prems: (2 pc'' prog pc ac)
+  with \<open>pc \<le> _\<close> consider "pc' = pc" | "pc' = Suc pc" | "pc' = Suc (Suc pc)" | "pc + 3 \<le> pc'"
+    by force
+  then show ?case using prems by cases (auto simp add: numeral_3_eq_3 is_conj_block_def)
+next
+  case prems: (3 pc'' prog pc ac)
+  with \<open>pc \<le> _\<close> consider
+    "pc' = pc" | "pc' = Suc pc" | "pc' = Suc (Suc pc)" | "pc' = Suc (Suc (Suc pc))" | "pc + 4 \<le> pc'"
+    by force
+  then show ?case using prems by cases (auto simp add: numeral_3_eq_3 numeral_4_eq_4 is_conj_block_def)+
+qed
+
+lemma is_conj_block_decomp:
+  "is_conj_block P pc' pc''" if
+  "is_conj_block P pc pc''" "P ! pc' = Some (CEXP ac)" "pc \<le> pc'" "pc' \<le> pc''"
+  using that
+  apply (subst (asm) is_conj_block_def)
+  apply safe
+  using One_nat_def add_Suc_right is_conj_block'_decomp le_simps(3) that(1)
+   apply (metis add.right_neutral le_neq_implies_less)
+  by (metis
+    One_nat_def Suc_1 add.right_neutral add_Suc_right instrc.simps(4) is_conj_block'_decomp
+    le_antisym not_less_eq_eq option.inject that(1))
+
+lemma steps_approx_finite[intro,simp]:
+  "finite (steps_approx n P pc_s)"
+  by (induction rule: steps_approx.induct; clarsimp split: option.split instrc.split instr.split)
+
+abbreviation "conv_P \<equiv> map (map_option (map_instrc real_of_int))"
+
+lemma stepst_stepc_extend:
+  "stepst P n u (pc', s') (pc'', s'')"
+  if "stepst P n u (pc, s) (pc'', s'')" "stepsc P n u (pc, s) (pc', s')"
+proof -
+  from that have "P pc'' = Some (INSTR HALT)" unfolding stepst_def by auto
+  from that obtain n1 n2 where *:
+    "stepsn P n1 u (pc, s) (pc'', s'')" "stepsn P n2 u (pc, s) (pc', s')" "n1 < n" "n2 < n"
+    unfolding stepst_def by (auto elim!: stepsc_stepsn)
+  show ?thesis
+  proof (cases "n1 \<ge> n2")
+    case True
+    from stepsn_extend[OF *(2,1) this] \<open>n1 < _\<close> \<open>n2 < _\<close> \<open>P pc'' = _\<close> show ?thesis
+      unfolding stepst_def by (auto intro: stepsn_stepsc)
+  next
+    case False
+    with stepsn_extend[OF *(1,2)] \<open>n1 < _\<close> \<open>n2 < _\<close> \<open>P pc'' = _\<close> show ?thesis
+      unfolding stepst_def by (auto intro: stepsn_stepsc dest!: stepsn_halt)
+  qed
+qed
+
+lemma conv_P_conj_block'[intro]:
+  "is_conj_block' (conv_P P) pc pc'" if "is_conj_block' P pc pc'"
+  using that
+  apply induction
+    apply (rule is_conj_block'.intros(1))
+     apply (simp; fail)
+    apply (simp; fail)
+   apply (rule is_conj_block'.intros(2))
+       apply (simp; fail)
+      apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+     apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+    apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+   apply (simp; fail)
+  apply (rule is_conj_block'.intros(3))
+       apply (simp; fail)
+      apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+     apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+    apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+   apply (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog)
+  apply (simp; fail)
+  done
+
+lemma conv_P_conj_block[intro]:
+  "is_conj_block (conv_P P) pc pc'" if "is_conj_block P pc pc'"
+  using that[unfolded is_conj_block_def]
+  apply safe
+  by (frule is_conj_block'_pc_mono; force dest: is_conj_block'_len_prog simp: is_conj_block_def)+
+
+context
+  fixes P :: "int instrc option list"
+    and pc_s :: addr
+    and n :: nat
+begin
+
+private abbreviation "prog i \<equiv> if i < length P then P ! i else None"
+
+lemma stepst_conv_P:
+  "stepst (\<lambda> i. if i < length (conv_P P) then conv_P P ! i else None) n u s s'" if
+  "stepst (conv_prog prog) n u s s'" using that unfolding stepst_def
+  apply safe
+  subgoal for pc a aa ab b z
+    apply rotate_tac
+    by (induction "conv_prog prog" n u s "(pc, a, aa, ab, b)" rule: stepsc.induct)
+      (auto split: if_split_asm)
+  by (auto split: if_split_asm)
+
+lemma is_conj:
+  fixes u :: "nat \<Rightarrow> real"
+  defines "S \<equiv> steps_approx n P pc_s"
+  defines "pc_c \<equiv> Min {pc. \<exists> ac. pc \<in> S \<and> P ! pc = Some (CEXP ac)}"
+  assumes "is_conj_block P pc_c (Max S)"
+    and "stepst (conv_prog prog) n u (pc_s, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+    and "stepsc (conv_prog prog) n u (pc_s, st, s, f, rs) (pc', st', s', f', rs')"
+    and "P ! pc' = Some (CEXP ac)" "pc' < length P"
+  shows "(u \<turnstile>a conv_ac ac) \<and> pc_t = Max S"
+proof -
+  from stepst_stepc_extend[OF assms(4,5)] have *:
+    "stepst (conv_prog prog) n u (pc', st', s', f', rs') (pc_t, st_t, s_t, True, rs_t)" .
+  from \<open>stepsc _ _ _ _ _\<close> \<open>P ! pc' = _\<close> \<open>pc' < _\<close> have "pc_c \<le> pc'" "pc' \<le> Max S"
+    unfolding pc_c_def S_def by (auto intro: stepsc_steps_approx Min_le Max_ge)
+  from is_conj_block_decomp[OF assms(3) \<open>P ! pc' = _\<close> this] have
+    "is_conj_block P pc' (Max S)" .
+  then have "is_conj_block (conv_P P) pc' (Max S)" by auto
+  from is_conj_block_is_conj[OF this _ stepst_conv_P[OF *]] \<open>P ! pc' = _\<close> \<open>pc' < _\<close>
+    show ?thesis
+    by auto
+qed
+
+lemma is_conj':
+  fixes u :: "nat \<Rightarrow> real"
+  defines "S \<equiv> steps_approx n P pc_s"
+  assumes "{pc. \<exists> ac. pc \<in> S \<and> P ! pc = Some (CEXP ac)} = {}"
+    and "stepsc (conv_prog prog) n u (pc_s, st, s, f, rs) (pc', st', s', f', rs')"
+    and "P ! pc' = Some (CEXP ac)" "pc' < length P"
+  shows False
+  using stepsc_steps_approx[OF assms(3,5)] assms(4) assms(2) unfolding S_def by auto
+
+definition
+"check_conj_block pc pc' \<equiv>
+   (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False) \<and> check_conj_block' P (pc + 1) = Some pc'
+   \<or> (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False) \<and> P ! (pc + 1) = Some (INSTR instr.AND)
+      \<and> check_conj_block' P (pc + 2) = Some pc'"
+
+lemma check_conj_block:
+  "check_conj_block pc pc' \<Longrightarrow> is_conj_block P pc pc'"
+  unfolding is_conj_block_alt_def check_conj_block_def
+  by (auto dest!: check_conj_block' simp del: check_conj_block'.simps)
+
+definition
+  "conjunction_check \<equiv>
+    let S = steps_approx n P pc_s; S' = {pc. \<exists> ac. pc \<in> S \<and> P ! pc = Some (CEXP ac)} in
+      S' = {} \<or> check_conj_block (Min S') (Max S)
+  "
+
+lemma conjunction_check_alt_def[code]:
+  "conjunction_check =
+    (
+     let
+        S = steps_approx n P pc_s;
+        S' = {pc. pc \<in> S \<and> (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False)}
+      in
+        S' = {} \<or> check_conj_block (Min S') (Max S)
+    )
+  "
+proof -
+  let ?S = "steps_approx n P pc_s"
+  have "
+    {pc. pc \<in> ?S \<and> (case P ! pc of Some (CEXP ac) \<Rightarrow> True | _ \<Rightarrow> False)}
+  = {pc. \<exists> ac. pc \<in> ?S \<and> P ! pc = Some (CEXP ac)}
+  " by safe (auto split: option.splits instrc.splits)
+  show ?thesis unfolding conjunction_check_def Let_def \<open>_ = _\<close> ..
+qed
+
+lemma conjunction_check:
+  fixes u :: "nat \<Rightarrow> real"
+  assumes "conjunction_check"
+    and "stepst (conv_prog prog) n u (pc_s, st, s, f, rs) (pc_t, st_t, s_t, True, rs_t)"
+    and "stepsc (conv_prog prog) n u (pc_s, st, s, f, rs) (pc', st', s', f', rs')"
+    and "P ! pc' = Some (CEXP ac)" "pc' < length P"
+  shows "u \<turnstile>a conv_ac ac"
+  using assms
+  unfolding conjunction_check_def Let_def
+  apply -
+  apply (erule disjE)
+   apply (drule is_conj'; simp)
+    apply (drule check_conj_block)
+  apply (subst is_conj; simp)
+  done
+
+end (* End of context for fixed program *)
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Asm.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Asm.thy
@@ -0,0 +1,172 @@
+theory UPPAAL_Asm
+  imports Main
+begin
+
+type_synonym addr = nat
+type_synonym val = int
+type_synonym reg = nat
+
+datatype instr =
+  JMPZ addr |
+  ADD |
+  NOT |
+  AND |
+  LT |
+  LE |
+  EQ |
+  PUSH int \<comment> \<open>Push value on stack\<close> |
+  POP |
+  LID reg \<comment> \<open>Push register value on stack\<close> |
+  STORE \<comment> \<open>Store stack value in register\<close> |
+  STOREI reg val \<comment> \<open>Store value in register\<close> |
+  COPY |
+  CALL |
+  RETURN |
+  HALT |
+  STOREC nat int \<comment> \<open>Special instruction, signals a clock reset\<close> |
+  SETF bool \<comment> \<open>Meta instruction\<close>
+
+type_synonym stack = "int list"
+type_synonym flag = bool
+type_synonym rstate = "int list" \<comment> \<open>Partial map from registers to values\<close>
+type_synonym state = "addr * stack * rstate * flag * nat list"
+\<comment> \<open>Instruction pointer, stack, register state, comparison flag, reset clocks\<close>
+
+definition int_of :: "bool \<Rightarrow> int" where
+  "int_of x \<equiv> if x then 1 else 0"
+
+fun step :: "instr \<Rightarrow> state \<Rightarrow> state option" where
+  "step (JMPZ q) (pc, st, m, f, rs) = Some (if f then (pc + 1) else q, st, m, f, rs)" |
+  "step ADD (pc, a # b # st, m, f, rs) = Some (pc + 1, (a + b) # st, m, f, rs)" |
+  "step NOT (pc, b # st, m , f, rs) = Some (pc + 1, st, m, \<not> f, rs)" |
+  "step AND (pc, b # st, m, f, rs) =
+    (if b = 0 \<or> b = 1
+     then Some (pc + 1, st, m, b = 1 \<and> f, rs)
+     else None)" |
+  "step LT (pc, a # b # st, m, f, rs) = Some (pc + 1, st, m, a < b, rs)" |
+  "step LE (pc, a # b # st, m, f, rs) = Some (pc + 1, st, m, a \<le> b, rs)" |
+  "step EQ (pc, a # b # st, m, f, rs) = Some (pc + 1, st, m, a = b, rs)" |
+  "step (PUSH v) (pc, st, m, f, rs) = Some (pc + 1, v # st, m, f, rs)" |
+  "step POP (pc, v # st, m, f, rs) = Some (pc + 1, st, m, f, rs)" |
+  "step (LID r) (pc, st, m, f, rs) = Some (pc + 1, m ! r # st, m, f, rs)" |
+  "step STORE (pc, v # r # st, m, f, rs) =
+    (if r \<ge> 0 then Some (pc + 1, st, m[nat r := v], f, rs) else None)" |
+  "step (STOREI r v) (pc, st, m, f, rs) = Some (pc + 1, st, m[r := v], f, rs)" |
+  "step COPY (pc, st, m, f, rs) = Some (pc + 1, int_of f # st, m, f, rs)" |
+  "step CALL (pc, q # st, m, f, rs) =
+    (if q \<ge> 0 then Some (nat q, int pc # st, m, f, rs) else None)" |
+  "step RETURN (pc, q # st, m, f, rs) =
+    (if q \<ge> 0 then Some (nat q + 1, st, m, f, rs) else None)" |
+  (*
+  "step HALT s = Some s" |
+  *)
+  "step (STOREC c d) (pc, st, m, f, rs) =
+    (if d = 0 then Some (pc + 1, st, m, f, c # rs) else None)" |
+  "step (SETF b) (pc, st, m, f, rs) = Some (pc + 1, st, m, b, rs)" |
+  "step _ _ = None"
+
+type_synonym program = "addr \<Rightarrow> instr option"
+type_synonym fuel = nat
+
+fun exec :: "program \<Rightarrow> fuel \<Rightarrow> state \<Rightarrow> addr list \<Rightarrow> (state * addr list) option" where
+  "exec _ 0 _ _ = None" |
+  "exec prog (Suc n) (pc, st, m, f, rs) pcs =
+    (case prog pc of
+       Some instr \<Rightarrow> (
+         if instr = HALT
+         then Some ((pc, st, m, f, rs), pc # pcs)
+         else
+           case step instr (pc, st, m, f, rs) of
+             Some s \<Rightarrow> exec prog n s (pc # pcs)
+           | None \<Rightarrow> None)
+     | None \<Rightarrow> None)"
+
+inductive steps :: "program \<Rightarrow> fuel \<Rightarrow> state \<Rightarrow> state \<Rightarrow> bool" where
+  "steps prog (Suc n) start start" | (*
+  "steps prog (Suc n) s s'" if
+    "step cmd (pc, st, m, f, rs) = Some s'"
+    "steps prog n s (pc, st, m, f, rs)"
+    "prog pc = Some cmd" | *)
+  "steps prog (Suc n) (pc, st, m, f, rs) s'" if
+    "step cmd (pc, st, m, f, rs) = Some s"
+    "prog pc = Some cmd"
+    "steps prog n s s'"
+
+inductive visited :: "program \<Rightarrow> fuel \<Rightarrow> state \<Rightarrow> state \<Rightarrow> addr list \<Rightarrow> bool" where
+  "visited prog (Suc n) s s []" |
+  "visited prog (Suc n) (pc, st, m, f, rs) s' (pcs @ [pc])" if
+    "step cmd (pc, st, m, f, rs) = Some s"
+    "prog pc = Some cmd"
+    "visited prog n s s' pcs"
+
+lemmas [intro] = steps.intros
+
+lemma exec_steps:
+  assumes "exec prog n s pcs = Some ((pc, st, m, f, rs), pcs')"
+  shows "steps prog n s (pc, st, m, f, rs) \<and> prog pc = Some HALT"
+using assms proof (induction P \<equiv> prog n s pcs arbitrary: pc st m f rs rule: exec.induct)
+  case 1
+  then show ?case by simp
+next
+  case (2 n pc' st' m' f' rs' pcs')
+  then obtain instr where "prog pc' = Some instr" by (cases "prog pc'") auto
+  show ?case
+  proof (cases "instr = HALT")
+    case True
+    with "2.prems" \<open>prog pc' = _\<close> show ?thesis by auto
+  next
+    case False
+    with 2 \<open>prog pc' = _\<close> show ?thesis by (auto split: option.split_asm)
+  qed
+qed
+
+lemma steps_halt:
+  assumes "steps prog n (pc, st, m, f, rs) s" "prog pc = Some HALT"
+  shows "s = (pc, st, m, f, rs)"
+  using assms by (induction prog n "(pc, st, m, f, rs)" s) auto
+
+lemma steps_exec:
+  assumes "steps prog n s (pc, st, m, f, rs)" "prog pc = Some HALT"
+  shows "\<exists> pcs'. exec prog n s pcs = Some ((pc, st, m, f, rs), pcs')"
+using assms proof (induction P \<equiv> prog n s "(pc, st, m, f, rs)" arbitrary: pcs rule: steps.induct)
+  case (1 n)
+  then show ?case by simp
+next
+  case (2 cmd pc' st' m' f' rs' s n)
+  then obtain pcs' where "exec prog n s (pc' # pcs) = Some ((pc, st, m, f, rs), pcs')" by auto
+  with 2(1-3) show ?case by (auto dest!: steps_halt)
+qed
+
+lemma visited_halt:
+  assumes "visited prog n (pc, st, m, f, rs) s pcs" "prog pc = Some HALT"
+  shows "s = (pc, st, m, f, rs)"
+  using assms by (induction prog n "(pc, st, m, f, rs)" s pcs) auto
+
+lemma exec_length_mono:
+  assumes "exec prog n s pcs = Some (s', pcs')"
+  shows "length pcs' > length pcs"
+  using assms
+    by (induction P \<equiv> prog n s pcs arbitrary: s' pcs' rule: exec.induct)
+       (force split: option.splits if_splits)+
+
+inductive_cases visitedE1: "visited prog n s (pc, st, m, f, rs) []"
+
+lemma visited_exec:
+  assumes "visited prog n s (pc, st, m, f, rs) pcs" "prog pc = Some HALT"
+  shows "exec prog n s pcs' = Some ((pc, st, m, f, rs), pc # pcs @ pcs')"
+using assms proof (induction P \<equiv> prog n s "(pc, st, m, f, rs)" pcs arbitrary: pcs' rule: visited.induct)
+  case (1 n)
+  then show ?case by simp
+next
+  case (2 cmd pc' st' m' f' rs' s n pcs pcs')
+  with 2 have "exec prog n s (pc' # pcs') = Some ((pc, st, m, f, rs), pc # pcs @ pc' # pcs')"
+    by auto
+  with 2(1-3) show ?case by (auto dest!: steps_halt)
+qed
+
+lemma visited_exec':
+  assumes "visited prog n s (pc, st, m, f, rs) pcs" "prog pc = Some HALT"
+  shows "exec prog n s [] = Some ((pc, st, m, f, rs), pc # pcs)"
+using visited_exec assms by auto
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Asm_Clocks.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Asm_Clocks.thy
@@ -0,0 +1,101 @@
+theory UPPAAL_Asm_Clocks
+  imports Timed_Automata.Timed_Automata Timed_Automata.Normalized_Zone_Semantics
+    UPPAAL_Asm Complex_Main
+begin
+
+(*
+abbreviation conv_ac :: "('a, int) acconstraint \<Rightarrow> ('a, real) acconstraint" where
+  "conv_ac \<equiv> map_acconstraint id real_of_int"
+*)
+
+datatype 't instrc =
+  INSTR instr |
+  CEXP "(nat, 't) acconstraint"
+
+fun stepc :: "'t instrc \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state option" where
+  "stepc (INSTR instr) u s =
+    (case step instr s of
+       Some s' \<Rightarrow> Some s'
+     | None \<Rightarrow> None)" |
+  "stepc (CEXP cc) u (pc, st, m, f, rs) = Some (pc + 1, st, m, u \<turnstile>a cc, rs)"
+
+type_synonym 't programc = "addr \<Rightarrow> 't instrc option"
+
+(*
+fun stepsc :: "'t programc \<Rightarrow> fuel \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state option" where
+  "stepsc _ 0 _ _ = None" |
+  "stepsc prog (Suc n) u (pc, st, m, f, rs) =
+    (case prog pc of
+       Some instr \<Rightarrow>
+         if instr = INSTR HALT
+         then Some (pc, st, m, f, rs)
+         else
+           case stepc instr u (pc, st, m, f, rs) of
+             Some s \<Rightarrow> stepsc prog n u s
+           | None \<Rightarrow> None
+     | None \<Rightarrow> None)"
+*)
+
+inductive stepsc :: "'t programc \<Rightarrow> fuel \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state \<Rightarrow> bool" where
+  "stepsc prog (Suc n) u start start" | (*
+  "stepsc prog (Suc n) u s s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s'"
+    "stepsc prog n u s (pc, st, m, f, rs)"
+    "prog pc = Some cmd" *)
+  "stepsc prog (Suc n) u (pc, st, m, f, rs) s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s"
+    "prog pc = Some cmd"
+    "stepsc prog n u s s'"
+
+inductive visitedc :: "'t programc \<Rightarrow> fuel \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state \<Rightarrow> addr list \<Rightarrow> bool" where
+  "visitedc prog (Suc n) u start start []" | (*
+  "stepsc prog (Suc n) u s s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s'"
+    "stepsc prog n u s (pc, st, m, f, rs)"
+    "prog pc = Some cmd" *)
+  "visitedc prog (Suc n) u (pc, st, m, f, rs) s' (pcs @ [pc])" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s"
+    "prog pc = Some cmd"
+    "visitedc prog n u s s' pcs"
+
+(*
+inductive visited :: "'t programc \<Rightarrow> fuel \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state \<Rightarrow> (nat, 't) cconstraint \<Rightarrow> bool" where
+  "visited prog (Suc n) u start start []" | (*
+  "stepsc prog (Suc n) u s s'" if
+    "stepc cmd u (pc, st, m, f, rs) = Some s'"
+    "stepsc prog n u s (pc, st, m, f, rs)"
+    "prog pc = Some cmd" *)
+  "visited prog (Suc n) u (pc, st, m, f, rs) s' (ac # cc)" if
+    "stepc (CEXP ac) u (pc, st, m, f, rs) = Some s"
+    "prog pc = Some (CEXP ac)"
+    "visited prog n u s s' S" |
+  "visited prog (Suc n) u (pc, st, m, f, rs) s' cc" if
+    "stepc (INSTR instr) u (pc, st, m, f, rs) = Some s"
+    "prog pc = Some (INSTR instr)"
+    "visited prog n u s s' cc"
+*)
+
+definition "stepst prog n u start \<equiv> \<lambda> (pc, st, m, f, rs).
+  stepsc prog n u start (pc, st, m, f, rs) \<and> prog pc = Some (INSTR HALT)"
+
+(*
+inductive stepsc :: "'t programc \<Rightarrow> fuel \<Rightarrow> (nat, 't :: time) cval \<Rightarrow> state \<Rightarrow> state \<Rightarrow> bool" where
+  "stepsc prog (Suc n) u (pc, st, m, f, rs) (pc, st, m, f, rs)" if "prog pc = Some (INSTR HALT)" |
+  "stepsc prog (Suc n) u start end" if
+    "stepsc prog n u start (pc, st, m, f, rs)"
+    "prog pc = Some cmd" "cmd \<noteq> INSTR HALT" "stepc cmd u (pc, st, m, f, rs) = Some end"
+*)
+
+fun strip :: "'t instrc \<Rightarrow> instr" where
+  "strip (INSTR instr) = instr" |
+  "strip _ = HALT"
+
+fun stripf :: "'t instrc \<Rightarrow> instr" where
+  "stripf (INSTR instr) = instr" |
+  "stripf _ = SETF False"
+
+fun stript :: "'t instrc \<Rightarrow> instr" where
+  "stript (INSTR instr) = instr" |
+  "stript _ = SETF True"
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Model_Checking.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_Model_Checking.thy
@@ -0,0 +1,1177 @@
+theory UPPAAL_Model_Checking
+  imports
+    UPPAAL_State_Networks_Impl_Refine
+    Munta_Base.TA_More
+    Munta_Base.Abstract_Term
+begin
+
+hide_const models
+
+(* Could be moved to UPPAAL_State_Networks *)
+inductive step_u' ::
+  "('a, 't :: time, 's) unta \<Rightarrow> nat \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval
+  \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>_ \<langle>_, _, _\<rangle> \<rightarrow> \<langle>_, _, _\<rangle>" [61,61,61,61,61,61] 61) where
+  "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L'', s'', u''\<rangle>" if
+  "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>" "a \<noteq> Del" "A \<turnstile>n \<langle>L', s', u'\<rangle> \<rightarrow>a \<langle>L'', s'', u''\<rangle>"
+
+inductive steps_un' ::
+  "('a, 't :: time, 's) unta \<Rightarrow> nat \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval
+  \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>_ \<langle>_, _, _\<rangle> \<rightarrow>* \<langle>_, _, _\<rangle>" [61,61,61,61,61,61] 61)
+where
+  refl: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L, s, u\<rangle>" |
+  step: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<Longrightarrow> A \<turnstile>n \<langle>L', s', u'\<rangle> \<rightarrow> \<langle>L'', s'', u''\<rangle>
+        \<Longrightarrow> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+
+declare steps_un'.intros[intro]
+
+lemma stepI2:
+  "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>" if
+  "A \<turnstile>n \<langle>L', s', u'\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>" "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+  using that by induction auto
+
+context Equiv_TA
+begin
+
+lemma prod_correct'_action:
+  "(\<exists> a. defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>) =
+   (\<exists> a. state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> a \<noteq> Del)"
+  apply standard    
+  subgoal
+    by (blast elim: prod.prod_sound'_action)
+   apply clarify
+   subgoal for a
+     apply (cases a; simp; erule prod.prod_complete_silent prod.prod_complete_sync, fast)
+     done
+  done
+
+lemma prod_correct'_delay:
+  "(\<exists> d. defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>) =
+   state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  by (blast dest: prod.prod_sound'_delay elim: prod.prod_complete_delay)
+
+lemma equiv_correct:
+  "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> = A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  by (blast intro!: equiv_sound equiv_complete)
+
+lemma prod_correct_action:
+  "(\<exists> a. defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>a \<langle>(L', s'), u'\<rangle>) =
+   (\<exists> a. A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> a \<noteq> State_Networks.label.Del)"
+  unfolding prod_correct'_action equiv_correct ..
+
+lemma prod_correct_delay:
+  "(\<exists> d. defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>d \<langle>(L', s'), u'\<rangle>) =
+  A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L', s', u'\<rangle>"
+  unfolding prod_correct'_delay equiv_correct ..
+
+lemma prod_correct:
+  "defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle> =
+  (\<exists> a. A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>)"
+  apply standard
+  subgoal
+    using prod_correct_action[of u L' s' u'] prod_correct_delay[of u L' s' u']
+      Timed_Automata.step.cases by metis
+  subgoal
+    apply clarify
+    subgoal for a
+      apply (cases a)
+      using prod_correct_action[of u L' s' u'] prod_correct_delay[of u L' s' u']
+        Timed_Automata.step.intros apply metis+
+      done
+    done
+  done
+
+context
+  assumes "0 < p"
+begin
+
+lemmas equiv_complete'' = equiv_complete''[OF _ \<open>0 < p\<close>]
+
+definition
+  "all_prop L' s' \<equiv>
+    (\<forall>q<p. \<exists>pc st s'' rs pcs.
+      exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+      Some ((pc, st, s'', True, rs), pcs)
+    ) \<and> bounded B s' \<and> L' \<in> defs.states' s' cancel\<open>\<and> (\<forall>q<p. (defs.P ! q) (L' ! q) s')\<close>
+  "
+
+lemma step_u_inv:
+  "all_prop L' s'" if "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  using equiv_complete''[OF that] equiv_complete'[OF that] unfolding all_prop_def by auto
+
+lemma step_inv:
+  "all_prop L' s'" if "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  using step_u_inv[OF equiv_sound[OF that]] .
+
+lemma Equiv_TA_I:
+  "Equiv_TA A n L' s'" if *[unfolded all_prop_def]: "all_prop L' s'"
+  using * by - (standard, auto intro!: pred_time_indep upd_time_indep clock_conj Len)
+
+lemma step_u'_inv:
+  "all_prop L'' s'' \<and> defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L'', s''), u''\<rangle>"
+  if "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L'', s'', u''\<rangle>"
+using that proof cases
+  case prems: (1 L' s' u')
+  from step_u_inv[OF prems(1)] have *[unfolded all_prop_def]: "all_prop L' s'" .
+  interpret equiv: Equiv_TA A n L' s'
+    using Equiv_TA_I[OF step_u_inv[OF prems(1)]] .
+  from equiv.step_u_inv[OF \<open>0 < p\<close> prems(3)] show ?thesis
+    using prems prod_correct_delay[of u L' s' u'] equiv.prod_correct_action[of u' L'' s'' u'']
+      Timed_Automata.step'.intros
+    by metis
+qed
+
+lemma step'_inv:
+  "all_prop L'' s'' \<and> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L'', s'', u''\<rangle>"
+  if "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L'', s''), u''\<rangle>"
+using that proof cases
+  case prems: (step' d l' u' a)
+  obtain L' s' where "l' = (L', s')"
+    by force
+  from step_inv prod_correct'_delay prems(1) have *:
+    "all_prop L' s'"
+    unfolding \<open>l' = _\<close> by fast
+  interpret equiv: Equiv_TA A n L' s'
+    by (rule Equiv_TA_I[OF *])
+  from equiv.step_inv[OF \<open>0 < p\<close>] equiv.prod_correct'_action prems(2)[unfolded \<open>l' = _\<close>] have
+    "all_prop L'' s''"
+    by metis
+  then show ?thesis
+    using prems prod_correct_delay[of u L' s' u'] equiv.prod_correct_action[of u' L'' s'' u'']
+      step_u'.intros
+    unfolding \<open>l' = _\<close> by metis
+qed
+
+lemma prod_correct_step':
+  "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle> =
+  A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+  using step'_inv step_u'_inv by blast
+
+lemma all_prop_start:
+  "all_prop L s"
+  using Equiv_TA_axioms unfolding Equiv_TA_def all_prop_def by auto
+
+lemma steps_u'_inv:
+  "all_prop L'' s'' \<and> defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L'', s''), u''\<rangle>"
+  if "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+  using that
+proof (induction A \<equiv> A n \<equiv> n L \<equiv> L s \<equiv> s u L'' s'' u'')
+  case (refl u)
+  show ?case using all_prop_start by auto
+next
+  case (step u L' s' u' L'' s'' u'')
+  then interpret equiv: Equiv_TA A n L' s'
+    by (blast intro: Equiv_TA_I)
+  from equiv.step_u'_inv[OF \<open>0 < p\<close> step.hyps(3)] step.hyps(1-2) show ?case
+    by (blast intro: steps'_altI)
+qed
+
+lemma steps'_inv:
+  "all_prop L'' s'' \<and> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+  if "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L'', s''), u''\<rangle>"
+  using that all_prop_start
+proof (induction defs.prod_ta "(L, s)" u "(L'', s'')" u'' arbitrary: L s)
+  case (refl' u)
+  then show ?case using all_prop_start by auto
+next
+  case (step' u l' u' u'' L s)
+  obtain L' s' where "l' = (L', s')" by force
+  from step' interpret equiv: Equiv_TA A n L s
+    by (blast intro: Equiv_TA_I)
+  from equiv.step'_inv[OF \<open>0 < p\<close> step'(1)[unfolded \<open>l' = _\<close>]] step'(3)[OF \<open>l' = _\<close>] show ?case
+    by (auto intro: stepI2)
+qed
+
+lemma steps_un'_complete:
+  "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L'', s''), u''\<rangle>"
+  if "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+  using steps_u'_inv[OF that] ..
+
+lemma steps'_sound:
+  "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+  if "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L'', s''), u''\<rangle>"
+  using steps'_inv[OF that] ..
+
+lemma prod_reachable_correct:
+  "defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<longleftrightarrow> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>"
+  using steps'_sound steps_un'_complete by fast
+
+lemma Bisimulation_Invariant_I:
+  "Bisimulation_Invariant
+  (\<lambda> (L, s, u) (L', s', u'). defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)
+  (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+  (=)
+  (\<lambda> (L, s, u). all_prop L s)
+  (\<lambda> (L, s, u). all_prop L s)"
+  apply (standard; clarsimp)
+  apply (simp add: Equiv_TA.prod_correct_step' Equiv_TA_I \<open>0 < p\<close>)+
+  using Equiv_TA.step_u'_inv Equiv_TA_I \<open>0 < p\<close> apply blast+
+  done
+
+interpretation Bisimulation_Invariant
+  "\<lambda> (L, s, u) (L', s', u'). defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>"
+  "\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>"
+  "(=)"
+  "\<lambda> (L, s, u). all_prop L s"
+  "\<lambda> (L, s, u). all_prop L s"
+  by (rule Bisimulation_Invariant_I)
+
+end (* p > 0 *)
+
+end (* Equiv TA *)
+
+definition models ("_,_ \<Turnstile>_ _" [61,61] 61) where
+  "A,a0 \<Turnstile>n \<Phi> \<equiv> (case \<Phi> of
+    formula.EX \<phi> \<Rightarrow>
+      Graph_Defs.Ex_ev
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_bexp \<phi> L s)
+  | formula.EG \<phi> \<Rightarrow>
+      Graph_Defs.Ex_alw
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_bexp \<phi> L s)
+  | formula.AX \<phi> \<Rightarrow>
+      Graph_Defs.Alw_ev
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_bexp \<phi> L s)
+  | formula.AG \<phi> \<Rightarrow>
+      Graph_Defs.Alw_alw
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_bexp \<phi> L s)
+  | formula.Leadsto \<phi> \<psi> \<Rightarrow>
+      Graph_Defs.leadsto
+        (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+        (\<lambda> (L, s, _). check_bexp \<phi> L s)
+        (\<lambda> (L, s, _). check_bexp \<psi> L s)
+  ) a0
+  "
+
+definition
+  "has_deadlock A n a0 \<equiv>
+    Graph_Defs.deadlock (\<lambda> (L, s, u) (L', s', u'). A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>) a0"
+
+lemmas models_iff = models_def[unfolded Graph_Defs.Ex_alw_iff Graph_Defs.Alw_alw_iff]
+
+context Reachability_Problem
+begin
+
+lemma reaches_steps':
+  "reaches (l, u) (l', u') \<longleftrightarrow> conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>"
+  apply standard
+  subgoal premises prems
+    using prems by (induction "(l, u)" "(l', u')" arbitrary: l' u') (auto intro: steps'_altI)
+  subgoal premises prems
+    using prems by induction (auto intro: converse_rtranclp_into_rtranclp)
+  done
+
+lemma clocks_I:
+  "(\<forall> c. c \<in> clk_set (conv_A A) \<longrightarrow> u c = u' c)" if "\<forall> c \<in> {1..n}. u c = u' c"
+  unfolding clk_set_conv_A using clocks_n using that by auto
+
+lemma init_dbm_reaches_iff:
+  "(\<exists> u \<in> [curry init_dbm]v,n. \<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>)
+  \<longleftrightarrow> ([curry (init_dbm :: real DBM')]v,n \<noteq> {} \<and>
+    (\<forall> u \<in> [curry init_dbm]v,n. \<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>))
+  "
+proof -
+  interpret ta_bisim: Bisimulation_Invariant
+    "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u'). l' = l \<and> (\<forall> c. c \<in> clk_set (conv_A A) \<longrightarrow> u c = u' c))"
+    "(\<lambda>_. True)" "(\<lambda>_. True)"
+    by (rule ta_bisimulation[of "conv_A A"])
+  show ?thesis
+    apply safe
+      apply force
+    subgoal for u1 u' u2
+      unfolding init_dbm_semantics reaches_steps'[symmetric]
+      apply (drule ta_bisim.A_B.simulation_reaches[of _ _ "(l0, u2)"])
+      subgoal
+        using clocks_I[of u1 u2] by fastforce
+      by auto
+    subgoal for u
+      by blast
+    done
+qed
+
+theorem reachable_decides_emptiness_new:
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+  \<longleftrightarrow> [curry (init_dbm :: real DBM')]v,n \<noteq> {} \<and>
+    (\<forall> u \<in> [curry init_dbm]v,n. \<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>)"
+  unfolding reachable_decides_emptiness init_dbm_reaches_iff ..
+
+lemma reachable_decides_emptiness'_new:
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {})
+  \<longleftrightarrow> (\<forall> u. (\<forall> c \<in> {1..n}. u c = 0) \<longrightarrow> (\<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle>))"
+  unfolding reachable_decides_emptiness_new
+  using init_dbm_semantics' init_dbm_semantics'' init_dbm_non_empty by blast
+
+lemma reachability_check_new_aux:
+  "(\<exists> D'. E** a0 (l', D') \<and> [curry (conv_M D')]v,n \<noteq> {} \<and> F l')
+  \<longleftrightarrow> (\<forall> u. (\<forall> c \<in> {1..n}. u c = 0) \<longrightarrow> (\<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l'))"
+  using reachable_decides_emptiness'_new[of l'] by fast
+
+theorem reachability_check_new:
+    "(\<exists> D'. E** a0 (l', D') \<and> F_rel (l', D'))
+    \<longleftrightarrow> (\<forall> u. (\<forall> c \<in> {1..n}. u c = 0) \<longrightarrow> (\<exists> u'. conv_A A \<turnstile>' \<langle>l0, u\<rangle> \<rightarrow>* \<langle>l', u'\<rangle> \<and> F l'))"
+  using reachability_check_new_aux[of l'] check_diag_empty_spec reachable_empty_check_diag
+  unfolding F_rel_def by auto
+
+lemma init_state_in_state_set:
+  "l0 \<in> state_set (trans_of A)" if "\<not> deadlock (l0, u0)"
+proof -
+  obtain l u where "conv_A A \<turnstile>' \<langle>l0, u0\<rangle> \<rightarrow> \<langle>l, u\<rangle>"
+    using \<open>\<not> deadlock _\<close> unfolding deadlock_def deadlocked_def by force
+  then have "l0 \<in> state_set (trans_of (conv_A A))"
+    unfolding state_set_def
+    by cases (auto elim!: step_a.cases step_t.cases)
+  then show ?thesis
+    unfolding state_set_def unfolding trans_of_def conv_A_def by (cases A) force
+qed
+
+lemma init_state_in_state_set':
+  "l0 \<in> state_set (trans_of A)" if "(\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0))"
+  using init_state_in_state_set that by auto
+
+end (* Reachability Problem *)
+
+context Reachability_Problem_Impl
+begin
+
+context
+    fixes Q :: "'s \<Rightarrow> bool" and Q_fun
+    assumes Q_fun: "(Q_fun, Q) \<in> inv_rel loc_rel states'"
+begin
+
+lemma leadsto_spec_refine:
+  "leadsto_spec_alt Q
+  \<le> SPEC (\<lambda> r. \<not> r \<longleftrightarrow>
+    (\<nexists>x. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** (l0, init_dbm) x \<and>
+       F (fst x) \<and> Q (fst x) \<and>
+       (\<exists>a. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))** x a \<and>
+            (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))++ a a)
+    ))"
+proof -
+  have *:"
+    (\<lambda>x y. (case y of (l', M') \<Rightarrow> E_op''.E_from_op x (l', M') \<and> \<not> check_diag n M') \<and>
+    \<not> (case y of (l, M) \<Rightarrow> check_diag n M))
+    = (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))"
+    by (intro ext) auto
+  have **:
+    "(\<lambda>x y. (case y of (l', M') \<Rightarrow> E_op''.E_from_op x (l', M') \<and> \<not> check_diag n M') \<and>
+     (case y of (l, M) \<Rightarrow> Q l) \<and> \<not> (case y of (l, M) \<Rightarrow> check_diag n M))
+     = (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))"
+    by (intro ext) auto
+  have ***: "\<not> check_diag n b"
+    if "(\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** a0 (a, b)" for a b
+    using that by cases (auto simp: a0_def)
+  show ?thesis
+    unfolding leadsto_spec_alt_def[OF Q_fun]
+    unfolding PR_CONST_def a0_def[symmetric] by (auto dest: *** simp: * **)
+  qed
+
+lemma leadsto_impl_hnr':
+  "(uncurry0
+    (leadsto_impl state_copy_impl
+      (succs_P_impl' Q_fun) a0_impl subsumes_impl (return \<circ> fst)
+      succs_impl' emptiness_check_impl F_impl (Q_impl Q_fun) tracei),
+   uncurry0
+    (SPEC
+      (\<lambda>r. (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0)) \<longrightarrow>
+           (\<not> r) =
+           (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow>
+                  leadsto (\<lambda>(l, u). F l) (\<lambda>(l, u). \<not> Q l) (l0, u0)))))
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+proof -
+  define p1 where "p1 \<equiv>
+    (\<nexists>x. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b))** (l0, init_dbm) x \<and>
+         F (fst x) \<and> Q (fst x) \<and>
+         (\<exists>a. (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))** x a \<and>
+              (\<lambda>a b. E_op''.E_from_op a b \<and> \<not> check_diag n (snd b) \<and> Q (fst b))++ a a))"
+  define p2 where "p2 \<equiv> (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> \<not> deadlock (l0, u0))"
+  define p3 where
+    "p3 \<equiv> (\<forall>u0. (\<forall>c\<in>{1..n}. u0 c = 0) \<longrightarrow> leadsto (\<lambda>(l, u). F l) (\<lambda>(l, u). \<not> Q l) (l0, u0))"
+  show ?thesis
+  unfolding state_set_eq[symmetric]
+  using Reachability_Problem_Impl_Op.leadsto_impl_hnr[OF Reachability_Problem_Impl_Op_axioms,
+    OF Q_fun precond_a0,
+    FCOMP leadsto_spec_refine[THEN Id_SPEC_refine, THEN nres_relI], to_hnr, unfolded hn_refine_def,
+    of show_state show_clock
+  ]
+  using init_state_in_state_set'
+  using leadsto_mc[of F Q]
+  unfolding p1_def[symmetric] p2_def[symmetric] p3_def[symmetric]
+  apply (simp del: state_set_eq)
+  apply sepref_to_hoare
+  apply sep_auto
+  apply (erule cons_post_rule)
+  apply sep_auto
+  done
+qed
+
+end (* Context for leadsto predicate *)
+
+end (* Reachability Problem Impl *)
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+lemma F_reachable_correct'_new:
+  "impl.op.F_reachable
+  \<longleftrightarrow> (\<exists> L' s'. \<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> (\<exists> u'.
+      conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>
+      \<and>  check_bexp \<phi> L' s')
+    )" if "formula = formula.EX \<phi>"
+  using
+    that E_op''.E_from_op_reachability_check[of F_rel "PR_CONST (\<lambda>(x, y). F x y)"]
+    reachability_check_new
+  unfolding impl.E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+  unfolding F_rel_def unfolding F_def by force
+
+lemma F_reachable_correct'_new':
+  "impl.op.F_reachable
+  \<longleftrightarrow> (\<exists> L' s'. \<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> (\<exists> u'.
+      conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>
+      \<and>  \<not> check_bexp \<phi> L' s')
+    )" if "formula = formula.AG \<phi>"
+  using
+    that E_op''.E_from_op_reachability_check[of F_rel "PR_CONST (\<lambda>(x, y). F x y)"]
+    reachability_check_new
+  unfolding impl.E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+  unfolding F_rel_def unfolding F_def by force
+
+lemma F_reachable_correct_new:
+  "impl.op.F_reachable
+  \<longleftrightarrow> (\<exists> L' s'. \<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> (\<exists> u'.
+      conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+       \<and> check_bexp \<phi> L' s')
+    )" if "formula = formula.EX \<phi>"
+    unfolding F_reachable_correct'_new[OF that]
+    apply (subst product'.prod_reachable_correct[symmetric])
+    using prod_conv p_p p_gt_0 by simp+
+
+lemma F_reachable_correct_new':
+  "impl.op.F_reachable
+  \<longleftrightarrow> (\<exists> L' s'. \<forall> u. (\<forall> c \<in> {1..m}. u c = 0) \<longrightarrow> (\<exists> u'.
+      conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+       \<and> \<not> check_bexp \<phi> L' s')
+    )" if "formula = formula.AG \<phi>"
+    unfolding F_reachable_correct'_new'[OF that]
+    apply (subst product'.prod_reachable_correct[symmetric])
+    using prod_conv p_p p_gt_0 by simp+
+
+definition
+  "Alw_ev_checker = dfs_map_impl'
+     (impl.succs_P_impl' final_fun) impl.a0_impl impl.subsumes_impl (return \<circ> fst)
+     impl.state_copy_impl"
+
+definition
+  "leadsto_checker \<psi> = do {
+      r \<leftarrow> leadsto_impl
+      impl.state_copy_impl (impl.succs_P_impl' (\<lambda> (L, s). \<not> check_bexp \<psi> L s))
+      
+      impl.a0_impl impl.subsumes_impl (return \<circ> fst)
+      impl.succs_impl' impl.emptiness_check_impl impl.F_impl
+      (impl.Q_impl (\<lambda> (L, s). \<not> check_bexp \<psi> L s))
+      impl.tracei;
+      return (\<not> r)
+    }"
+
+definition
+  "model_checker = (
+  case formula of
+    formula.EX _ \<Rightarrow> reachability_checker' |
+    formula.AG _ \<Rightarrow> do {
+      r \<leftarrow> reachability_checker';
+      return (\<not> r)
+    } |
+    formula.AX _ \<Rightarrow> do {
+      r \<leftarrow> if PR_CONST (\<lambda>(x, y). F x y) (init, s0)
+      then Alw_ev_checker
+      else return False;
+      return (\<not> r)
+    } |
+    formula.EG _ \<Rightarrow>
+      if PR_CONST (\<lambda>(x, y). F x y) (init, s0)
+      then Alw_ev_checker
+      else return False |
+    formula.Leadsto _ \<psi> \<Rightarrow> leadsto_checker \<psi>
+  )
+  "
+
+lemma p'_gt_0:
+  "0 < defs'.p"
+  unfolding p_p by (rule p_gt_0)
+
+interpretation Bisim_A: Bisimulation_Invariant
+   "(\<lambda>(L, s, u) (L', s', u').
+       defs'.defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+   "(\<lambda>(L, s, u) (L', s', u').
+       conv N \<turnstile>max_steps \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)"
+   "(=)" "(\<lambda>(L, s, u). product'.all_prop L s)"
+   "(\<lambda>(L, s, u). product'.all_prop L s)"
+  by (rule product'.Bisimulation_Invariant_I) (rule p'_gt_0)
+
+interpretation Bisim_B: Bisimulation_Invariant
+  "(\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+  "(\<lambda>(L, s, u) (L', s', u').
+       defs'.defs.prod_ta \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow> \<langle>(L', s'), u'\<rangle>)"
+   "(\<lambda> (l, u') (L, s, u). (l, u') = ((L, s), u))" "\<lambda> _. True"
+   "(\<lambda>(L, s, u). product'.all_prop L s)"
+  unfolding prod_conv[symmetric]
+  apply standard
+  apply (solves auto)+
+  by (rule Bisim_A.A_invariant)
+
+interpretation Bisimulation_Invariant
+   "(\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+   "(\<lambda>(L, s, u) (L', s', u').
+       conv N \<turnstile>max_steps \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)"
+   "(\<lambda> (l, u') (L, s, u). (l, u') = ((L, s), u))" "\<lambda> _. True"
+   "(\<lambda>(L, s, u). product'.all_prop L s)"
+proof -
+  interpret bisim: Bisimulation_Invariant
+    "(\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+   "(\<lambda>(L, s, u) (L', s', u').
+       conv N \<turnstile>max_steps \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)"
+   "(\<lambda>a c. \<exists>b. (case b of (L, s, u) \<Rightarrow> product'.all_prop L s) \<and>
+              (case a of (l, u') \<Rightarrow> \<lambda>(L, s, u). (l, u') = ((L, s), u)) b \<and> b = c)"
+   "\<lambda> _. True"
+   "(\<lambda>(L, s, u). product'.all_prop L s)"
+    using Bisimulation_Invariant_composition[OF
+        Bisim_B.Bisimulation_Invariant_axioms Bisim_A.Bisimulation_Invariant_axioms
+        ]
+    .
+  show "Bisimulation_Invariant
+     (\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+     (\<lambda>(L, s, u) (L', s', u').
+       conv N \<turnstile>max_steps \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>)
+     (\<lambda>(l, u') (L, s, u). (l, u') = ((L, s), u)) (\<lambda>_. True) (\<lambda>(L, s, u). product'.all_prop L s)"
+    apply standard
+    subgoal for a b a'
+      apply (drule bisim.A_B_step[of a b a'])
+         apply auto
+      done
+    subgoal for a b a'
+      apply (drule bisim.B_A_step[of b a' a])
+         apply auto
+      done
+     apply simp
+    apply (drule bisim.B_invariant)
+     apply auto
+    done
+qed
+
+lemma models_correct:
+  "conv N,(init, s0, u0) \<Turnstile>max_steps \<Phi> = (case \<Phi> of
+    formula.EX \<phi> \<Rightarrow>
+        (\<lambda> ((L, s), u). \<exists> L' s' u'. conv_A A \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<and> check_bexp \<phi> L' s')
+  | formula.EG \<phi> \<Rightarrow>
+      Not o Graph_Defs.Alw_ev
+        (\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). \<not> check_bexp \<phi> L s)
+  | formula.AX \<phi> \<Rightarrow>
+      Graph_Defs.Alw_ev
+        (\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). check_bexp \<phi> L s)
+  | formula.AG \<phi> \<Rightarrow>
+      Not o (\<lambda> ((L, s), u).
+        \<exists> L' s' u'. conv_A A \<turnstile>' \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<and> \<not> check_bexp \<phi> L' s'
+      )
+  | formula.Leadsto \<phi> \<psi> \<Rightarrow>
+      Graph_Defs.leadsto
+        (\<lambda> (l, u) (l', u'). conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)
+        (\<lambda> ((L, s), _). check_bexp \<phi> L s)
+        (\<lambda> ((L, s), _). check_bexp \<psi> L s)
+  ) ((init, s0), u0)" if "\<not> deadlock ((init, s0), u0)"
+proof -
+  have *: "((Not \<circ>\<circ> case_prod) (\<lambda>(L, s) _. check_bexp \<phi> L s)) =
+    (\<lambda>((L, s), _). \<not> check_bexp \<phi> L s)" for \<phi> by auto
+
+  show ?thesis
+    apply (subst models_def)
+    apply (cases \<Phi>)
+    subgoal for \<phi>
+      apply simp
+
+      apply (subst Ex_ev_iff[
+            of "(\<lambda>((L, s), _). check_bexp \<phi> L s)" _ "((init, s0), u0)", symmetric, simplified
+            ])
+        apply (drule equiv'_D[simplified], force)
+       using product'.all_prop_start[OF p'_gt_0] apply (simp add: A_B.equiv'_def[simplified]; fail)
+      apply (subst Ex_ev[OF that])
+      unfolding reaches_steps'[symmetric]
+      apply auto
+      done
+    subgoal for \<phi>
+      apply simp
+      apply (subst Ex_alw_iff[
+            of "(\<lambda>((L, s), _). check_bexp \<phi> L s)" _ "((init, s0), u0)", symmetric, simplified
+            ])
+        apply (drule equiv'_D[simplified]; force)
+       using product'.all_prop_start[OF p'_gt_0] apply (simp add: A_B.equiv'_def[simplified]; fail)
+      unfolding Graph_Defs.Ex_alw_iff * ..
+    subgoal for \<phi>
+      apply simp
+      apply (subst Alw_ev_iff[
+            of "(\<lambda>((L, s), _). check_bexp \<phi> L s)" _ "((init, s0), u0)", symmetric, simplified
+            ])
+        apply (drule equiv'_D[simplified]; force)
+       using product'.all_prop_start[OF p'_gt_0] apply (simp add: A_B.equiv'_def[simplified]; fail)
+      unfolding Graph_Defs.Ex_alw_iff * ..
+    subgoal for \<phi>
+      apply simp
+      unfolding Graph_Defs.Alw_alw_iff
+      apply (subst Ex_ev_iff[
+            of "(\<lambda>((L, s), _). \<not>check_bexp \<phi> L s)" _ "((init, s0), u0)", symmetric, simplified
+            ])
+        apply (drule equiv'_D[simplified], subst *[symmetric], force)
+       using product'.all_prop_start[OF p'_gt_0] apply (simp add: A_B.equiv'_def[simplified]; fail)
+      apply (subst Ex_ev[OF that])
+      unfolding reaches_steps'[symmetric]
+      apply auto
+      done
+    subgoal for \<phi> \<psi>
+      apply simp
+      apply (subst Leadsto_iff[
+            of "(\<lambda>((L, s), _). check_bexp \<phi> L s)" _
+               "(\<lambda>((L, s), _). check_bexp \<psi> L s)" _ "((init, s0), u0)", symmetric, simplified
+            ])
+         apply (drule equiv'_D[simplified]; force)
+        apply (drule equiv'_D[simplified]; force)
+       using  product'.all_prop_start[OF p'_gt_0] apply (simp add: A_B.equiv'_def[simplified]; fail)
+      ..
+    done
+qed
+
+lemma final_fun_final':
+  "((\<lambda> (L, s). P L s), (\<lambda> (L, s). P L s)) \<in> inv_rel Id states'" for P
+  unfolding F_def final_fun_def inv_rel_def in_set_member[symmetric] list_ex_iff
+  by (force dest!: states'_states')
+
+lemma final_fun_final[intro, simp]:
+  "((\<lambda> (L, s). P L s), (\<lambda> (L, s). P L s)) \<in> inv_rel Id states" for P
+  using final_fun_final' states_states' by (rule inv_rel_mono)
+
+abbreviation "u0 \<equiv> (\<lambda> _. 0 :: real)"
+
+lemma deadlock_start_iff:
+  "Bisim_A.B.deadlock (init, s0, \<lambda>_. 0) \<longleftrightarrow> deadlock ((init, s0), u0)"
+  using product'.all_prop_start[OF p'_gt_0]
+  by - (rule deadlock_iff[of _ "(init, s0, u0)", symmetric]; simp)
+
+theorem model_check':
+  "(uncurry0 model_checker,
+    uncurry0 (
+      SPEC (\<lambda> r.
+        \<not> Graph_Defs.deadlock
+          (\<lambda> (L, s, u) (L', s', u'). conv N \<turnstile>max_steps \<langle>L, s, u\<rangle> \<rightarrow> \<langle>L', s', u'\<rangle>) (init, s0, u0) \<longrightarrow>
+        r = (conv N,(init, s0, u0) \<Turnstile>max_steps formula)
+      )
+    )
+   )
+  \<in> unit_assnk \<rightarrow>a bool_assn"
+proof -
+  have *: "(\<lambda>(l, u). \<not> (case l of (L, s) \<Rightarrow> (Not \<circ>\<circ>\<circ> check_bexp) \<phi> L s))
+    = (\<lambda>((L, s), _). check_bexp \<phi> L s)" for \<phi>
+    by auto
+  have **:
+    "(\<lambda>(l, u). \<not> (case l of (L, s) \<Rightarrow> check_bexp \<phi> L s)) = (\<lambda>((L, s), _). \<not> check_bexp \<phi> L s)"
+    for \<phi> by auto
+  have ***:
+    "(\<lambda>(l, u). case l of (L, s) \<Rightarrow> \<phi> L s) = (\<lambda>((L, s), _). \<phi> L s)" for \<phi>
+    by auto
+  have ****: "(\<lambda>(L, y). (Not \<circ>\<circ>\<circ> check_bexp) \<psi> L y) = (\<lambda>(L, y). \<not>check_bexp \<psi> L y)" for \<psi>
+    by auto
+  have *****:
+    "return True = (return False \<bind> return o Not)"
+    by auto
+
+  interpret ta_bisim: Bisimulation_Invariant
+    "(\<lambda>(l, u) (l', u').
+       conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u').
+       conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>)"
+    "(\<lambda>(l, u) (l', u').
+       l' = l \<and>
+       (\<forall> c. c \<in> clk_set (conv_A A) \<longrightarrow>
+            u c = u' c))"
+    "(\<lambda>_. True)" "(\<lambda>_. True)"
+    by (rule ta_bisimulation[of "conv_A A"])
+
+  have bisim2:
+    "(\<exists>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<and>
+                  \<not> Alw_ev (\<lambda>(l, u). \<phi> l) ((init, s0), u0))
+    \<longleftrightarrow> (\<not> Alw_ev (\<lambda>(l, u). \<phi> l) ((init, s0), u0))
+    " for \<phi>
+    apply safe
+    subgoal for u
+      apply (subst (asm) ta_bisim.Alw_ev_iff[of _ "(\<lambda>(l, u). \<phi> l)" _ "((init, s0), \<lambda>_. 0)"])
+      using clocks_I[of u "\<lambda>_. 0"] unfolding ta_bisim.A_B.equiv'_def
+        apply auto
+      done
+    by force
+
+  have bisim1:
+    "(\<exists>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<and> \<not> Alw_ev (\<lambda>((L, s), _). \<not> check_bexp \<phi> L s) ((init, s0), u0)) =
+     (\<not> Alw_ev (\<lambda>((L, s), _). \<not> check_bexp \<phi> L s) ((init, s0), u0))" for \<phi>
+    using bisim2[of "\<lambda> (L, s). \<not> check_bexp \<phi> L s"] unfolding *** .
+
+  have bisim3:
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow>
+      leadsto (\<lambda>((L, s), _). check_bexp \<phi> L s) (\<lambda>((L, s), _). check_bexp \<psi> L s) ((init, s0), u0)) =
+      leadsto (\<lambda>((L, s), _). check_bexp \<phi> L s) (\<lambda>((L, s), _). check_bexp \<psi> L s) ((init, s0), u0)
+    " for \<phi> \<psi>
+    apply safe
+     apply force
+    subgoal for u
+      apply (subst (asm) ta_bisim.Leadsto_iff[of
+            _ "(\<lambda>((L, s), _). check_bexp \<phi> L s)" _ "(\<lambda>((L, s), _). check_bexp \<psi> L s)"
+            _ "((init, s0), u)"
+            ])
+      using clocks_I[of u "\<lambda>_. 0"] unfolding ta_bisim.A_B.equiv'_def by auto
+    done
+
+  have bisim4:
+    "(\<forall>u0. (\<forall>c\<in>{Suc 0..m}. u0 c = 0) \<longrightarrow> \<not> deadlock ((init, s0), u0))
+    \<longleftrightarrow> \<not> deadlock ((init, s0), u0)
+    "
+    apply safe
+     apply fast
+    subgoal for u
+      apply (subst (asm) ta_bisim.deadlock_iff[of _ "((init, s0), u)"])
+      using clocks_I[of u "\<lambda>_. 0"] unfolding ta_bisim.A_B.equiv'_def by auto
+    done
+
+  have bisim5:
+    "(\<forall>u. (\<forall>c\<in>{1..m}. u c = 0) \<longrightarrow> (\<exists>u'. conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<and> \<phi> L' s'))
+  \<longleftrightarrow> (\<exists>u'. conv_A A \<turnstile>' \<langle>(init, s0), u0\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<and> \<phi> L' s')
+  " for \<phi> L' s'
+    unfolding reaches_steps'[symmetric]
+    apply safe
+     apply fast
+    subgoal for u' u
+      apply (drule ta_bisim.bisim.A_B_reaches[of _ _ "((init, s0), u)"])
+      subgoal
+        using clocks_I[of u "\<lambda>_. 0"] unfolding ta_bisim.equiv'_def by auto
+      unfolding ta_bisim.equiv'_def by auto
+    done
+
+  define protect where
+    "protect = ((\<lambda>(l, u) (l', u').
+                              conv_A A \<turnstile>' \<langle>l, u\<rangle> \<rightarrow> \<langle>l', u'\<rangle>))"
+
+  show ?thesis
+    unfolding deadlock_start_iff
+    using models_correct
+    apply simp
+    unfolding protect_def[symmetric]
+    apply sepref_to_hoare
+    apply sep_auto
+    unfolding model_checker_def reachability_checker'_def Alw_ev_checker_def leadsto_checker_def
+    apply (cases formula; simp)
+
+      \<comment> \<open>\<open>EX\<close>\<close>
+    subgoal premises prems for \<phi>
+      using impl.pw_impl_hnr_F_reachable[to_hnr, unfolded hn_refine_def]
+      apply (subst (asm) (2) F_reachable_correct'_new)
+       apply (rule prems; fail)
+      apply (subst (asm) bisim5)
+      apply sep_auto
+      unfolding final_fun_def F_def prems
+      apply (erule cons_post_rule)
+      apply (sep_auto simp: pure_def)
+      done
+
+        \<comment> \<open>\<open>EG\<close>\<close>
+    subgoal premises prems for \<phi>
+      using impl.Alw_ev_impl_hnr[
+          to_hnr, unfolded hn_refine_def
+          ]
+      unfolding final_fun_def F_def prems(2)
+        UPPAAL_Reachability_Problem_precompiled'.final_fun_def[
+          OF UPPAAL_Reachability_Problem_precompiled'_axioms
+          ]
+        UPPAAL_Reachability_Problem_precompiled_defs.F_def
+      apply sep_auto
+      unfolding **
+      subgoal
+        apply (subst (asm) bisim1)
+        apply (erule cons_post_rule)
+        using init_state_in_state_set[of u0]
+        apply (sep_auto simp: pure_def protect_def ***)
+        done
+      subgoal
+        apply (subst (asm) bisim1)
+        apply simp
+        apply (erule cons_post_rule)
+        using init_state_in_state_set[of u0]
+        apply (sep_auto simp: pure_def protect_def ***)
+        done
+      done
+
+        \<comment> \<open>\<open>AX\<close>\<close>
+    subgoal premises prems for \<phi>
+      using impl.Alw_ev_impl_hnr[to_hnr, unfolded hn_refine_def]
+      unfolding final_fun_def F_def
+      unfolding UPPAAL_Reachability_Problem_precompiled_defs.F_def
+      apply (subst
+          UPPAAL_Reachability_Problem_precompiled'.final_fun_def[
+            OF UPPAAL_Reachability_Problem_precompiled'_axioms
+            ])
+      apply (safe; clarsimp simp: prems(2))
+      unfolding bisim2
+      unfolding * ***
+      subgoal
+        using init_state_in_state_set[of u0]
+        by (sep_auto simp: pure_def protect_def)
+      subgoal
+        unfolding protect_def *****
+        apply (erule bind_rule)
+        using init_state_in_state_set[of u0]
+        apply (sep_auto simp: pure_def)
+        done
+      done
+
+        \<comment> \<open>\<open>AG\<close>\<close>
+    subgoal premises prems for \<phi>
+      using impl.pw_impl_hnr_F_reachable[to_hnr, unfolded hn_refine_def]
+      apply (subst (asm) (2) F_reachable_correct'_new')
+       apply (rule prems; fail)
+      apply (subst (asm) bisim5)
+      unfolding final_fun_def F_def prems
+      apply (sep_auto simp: pure_def)
+      done
+
+        \<comment> \<open>\<open>Leadsto\<close>\<close>
+    subgoal premises prems for \<phi> \<psi>
+      using impl.leadsto_impl_hnr'[
+          OF final_fun_final, of "Not oo check_bexp \<psi>",
+          to_hnr, unfolded hn_refine_def
+          ]
+      unfolding * F_def
+      apply (simp add: prems(2))
+      unfolding *** **** bisim3 bisim4
+      apply (erule bind_rule)
+      apply (sep_auto simp: pure_def protect_def)
+      done
+    done
+qed
+
+theorem model_check'_hoare:
+  "<emp>
+    model_checker
+  <\<lambda>r. \<up> ((\<not> Bisim_A.B.deadlock (init, s0, \<lambda>_. 0)) \<longrightarrow> r = (
+    conv N,(init, s0, u0) \<Turnstile>max_steps formula
+  ))>t"
+  using model_check'[to_hnr, unfolded hn_refine_def]
+  by (sep_auto simp: pure_def elim!: cons_post_rule)
+
+lemma Alw_ev_checker_alt_def':
+  "Alw_ev_checker \<equiv>
+    do {
+      x \<leftarrow> let
+        key = return \<circ> fst;
+        sub = impl.subsumes_impl;
+        copy = impl.state_copy_impl;
+        start = impl.a0_impl;
+        succs =  impl.succs_P_impl' final_fun
+      in dfs_map_impl' succs start sub key copy;
+      _ \<leftarrow> return ();
+      return x
+    }"
+  unfolding Alw_ev_checker_def by simp
+
+lemma leadsto_checker_alt_def':
+  "leadsto_checker \<psi> \<equiv>
+    do {
+      r \<leftarrow> let
+        key = return \<circ> fst;
+        sub = impl.subsumes_impl;
+        copy = impl.state_copy_impl;
+        start = impl.a0_impl;
+        final = impl.F_impl;
+        final' = (impl.Q_impl (\<lambda>(L, s). \<not> check_bexp \<psi> L s));
+        succs =  impl.succs_P_impl' (\<lambda>(L, s). \<not> check_bexp \<psi> L s);
+        succs' =  impl.succs_impl';
+        empty = impl.emptiness_check_impl;
+        trace = impl.tracei
+      in
+        leadsto_impl copy succs start sub key succs' empty final final' trace;
+      return (\<not> r)
+    }"
+  unfolding leadsto_checker_def by simp
+
+schematic_goal succs_P_impl_alt_def:
+  "impl.succs_P_impl (\<lambda>(L, s). P L s) \<equiv> ?impl" for P
+  unfolding impl.succs_P_impl_def[OF final_fun_final]
+  unfolding k_impl_alt_def
+  apply (abstract_let "trans_fun" trans_fun)
+  unfolding inv_fun_def[abs_def] trans_fun_def[abs_def] trans_s_fun_def trans_i_fun_def trans_i_from_def
+  apply (abstract_let "IArray (map IArray inv)" inv_a)
+  apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+  apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+  apply (abstract_let "IArray (map IArray trans_i_map)" trans_i_map)
+  by (rule Pure.reflexive)
+
+(* These implementations contain unnecessary list reversals *)
+lemmas succs_P'_impl_alt_def =
+  impl.succs_P_impl'_def[OF final_fun_final, unfolded succs_P_impl_alt_def]
+
+lemmas succs_impl'_alt_def =
+  impl.succs_impl'_def[unfolded succs_impl_alt_def]
+
+lemma reachability_checker'_alt_def':
+  "reachability_checker' \<equiv>
+    do {
+      x \<leftarrow> do {
+        let key = return \<circ> fst;
+        let sub = impl.subsumes_impl;
+        let copy = impl.state_copy_impl;
+        let start = impl.a0_impl;
+        let final = impl.F_impl;
+        let succs =  impl.succs_impl;
+        let empty = impl.emptiness_check_impl;
+        let tracei = impl.tracei;
+        pw_impl key copy tracei sub start final succs empty
+      };
+      _ \<leftarrow> return ();
+      return x
+    }"
+  unfolding reachability_checker'_def by simp
+
+schematic_goal reachability_checker'_alt_def:
+  "reachability_checker' \<equiv> ?impl"
+  unfolding reachability_checker'_alt_def' impl.succs_impl_def
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i)
+  (* apply (abstract_let "inv_fun" inv_fun) *) (* This is not pulling anything *)
+  apply (abstract_let "trans_fun" trans_fun)
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding final_fun_def[abs_def]
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+schematic_goal Alw_ev_checker_alt_def:
+  "Alw_ev_checker \<equiv> ?impl"
+  unfolding Alw_ev_checker_alt_def' final_fun_def
+    impl.succs_P_impl_def[OF final_fun_final] impl.succs_P_impl'_def[OF final_fun_final]
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i)
+  (* apply (abstract_let "inv_fun" inv_fun) *) (* This is not pulling anything *)
+  apply (abstract_let "trans_fun" trans_fun)
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding impl.F_impl_def
+  unfolding final_fun_def[abs_def]
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+schematic_goal leadsto_checker_alt_def:
+  "leadsto_checker \<equiv> ?impl"
+  unfolding leadsto_checker_alt_def'
+  unfolding impl.F_impl_def impl.Q_impl_def[OF final_fun_final]
+  unfolding impl.succs_P_impl'_def[OF final_fun_final]
+  unfolding impl.succs_impl'_def impl.succs_impl_def impl.succs_P_impl_def[OF final_fun_final]
+  unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+  unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+  unfolding k_impl_alt_def
+  apply (abstract_let k_i k_i)
+  (* apply (abstract_let "inv_fun" inv_fun) *) (* This is not pulling anything *)
+  apply (abstract_let "trans_fun" trans_fun)
+  unfolding impl.init_dbm_impl_def impl.a0_impl_def
+  unfolding final_fun_def
+  unfolding impl.subsumes_impl_def
+  unfolding impl.emptiness_check_impl_def
+  unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+schematic_goal reachability_checker'_alt_def_refined:
+  "reachability_checker' \<equiv> ?impl"
+  unfolding reachability_checker'_alt_def
+  unfolding fw_impl'_int
+  unfolding inv_fun_def trans_fun_def trans_s_fun_def trans_i_fun_def
+  unfolding trans_i_from_impl
+  unfolding runf_impl runt_impl check_g_impl pairs_by_action_impl check_pred_impl
+  apply (abstract_let "IArray (map IArray inv)" inv_ia)
+  apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+  apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+  apply (abstract_let "IArray (map IArray trans_i_map)" trans_i_map)
+  apply (abstract_let "IArray bounds" bounds_ia)
+  apply (abstract_let PF PF)
+  apply (abstract_let PT PT)
+  unfolding PF_alt_def PT_alt_def
+  apply (abstract_let PROG' PROG')
+  unfolding PROG'_def
+  apply (abstract_let "length prog" len_prog)
+  apply (abstract_let "IArray (map (map_option stripf) prog)" progf_ia)
+  apply (abstract_let "IArray (map (map_option stript) prog)" progt_ia)
+  apply (abstract_let "IArray prog" prog_ia)
+  unfolding all_actions_by_state_impl
+  apply (abstract_let "[0..<p]")
+  apply (abstract_let "[0..<na]")
+  apply (abstract_let "{0..<p}")
+  apply (abstract_let "[0..<m+1]")
+  by (rule Pure.reflexive)
+
+schematic_goal Alw_ev_checker_alt_def_refined:
+  "Alw_ev_checker \<equiv> ?impl"
+  unfolding Alw_ev_checker_alt_def
+  unfolding fw_impl'_int
+  unfolding inv_fun_def trans_fun_def trans_s_fun_def trans_i_fun_def
+  unfolding trans_i_from_impl
+  unfolding runf_impl runt_impl check_g_impl pairs_by_action_impl check_pred_impl
+  apply (abstract_let "IArray (map IArray inv)" inv_ia)
+  apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+  apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+  apply (abstract_let "IArray (map IArray trans_i_map)" trans_i_map)
+  apply (abstract_let "IArray bounds" bounds)
+  apply (abstract_let PF PF)
+  apply (abstract_let PT PT)
+  unfolding PF_alt_def PT_alt_def
+  apply (abstract_let PROG' PROG')
+  unfolding PROG'_def
+  apply (abstract_let "length prog" len_prog)
+  apply (abstract_let "IArray (map (map_option stripf) prog)" progf_ia)
+  apply (abstract_let "IArray (map (map_option stript) prog)" progt_ia)
+  apply (abstract_let "IArray prog" prog)
+  unfolding all_actions_by_state_impl
+  apply (abstract_let "[0..<p]")
+  apply (abstract_let "[0..<na]")
+  apply (abstract_let "{0..<p}")
+  apply (abstract_let "[0..<m+1]")
+  by (rule Pure.reflexive)
+
+schematic_goal leadsto_checker_alt_def_refined:
+  "leadsto_checker \<equiv> ?impl"
+  unfolding leadsto_checker_alt_def
+  unfolding fw_impl'_int
+  unfolding inv_fun_def trans_fun_def trans_s_fun_def trans_i_fun_def
+  unfolding trans_i_from_impl
+  unfolding runf_impl runt_impl check_g_impl pairs_by_action_impl check_pred_impl
+  apply (abstract_let "IArray (map IArray inv)" inv_ia)
+  apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+  apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+  apply (abstract_let "IArray (map IArray trans_i_map)" trans_i_map)
+  apply (abstract_let "IArray bounds" bounds_ia)
+  apply (abstract_let PF PF)
+  apply (abstract_let PT PT)
+  unfolding PF_alt_def PT_alt_def
+  apply (abstract_let PROG' PROG')
+  unfolding PROG'_def
+  apply (abstract_let "length prog" len_prog)
+  apply (abstract_let "IArray (map (map_option stripf) prog)" progf_ia)
+  apply (abstract_let "IArray (map (map_option stript) prog)" progt_ia)
+  apply (abstract_let "IArray prog" prog_ia)
+  unfolding all_actions_by_state_impl
+  apply (abstract_let "[0..<p]")
+  apply (abstract_let "[0..<na]")
+  apply (abstract_let "{0..<p}")
+  apply (abstract_let "[0..<m+1]")
+  by (rule Pure.reflexive)
+
+end
+
+concrete_definition reachability_checker'
+  uses UPPAAL_Reachability_Problem_precompiled'.reachability_checker'_alt_def_refined
+
+concrete_definition Alw_ev_checker
+  uses UPPAAL_Reachability_Problem_precompiled'.Alw_ev_checker_alt_def_refined
+
+concrete_definition leadsto_checker
+  uses UPPAAL_Reachability_Problem_precompiled'.leadsto_checker_alt_def_refined
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+lemmas model_checker_def_refined = model_checker_def[unfolded
+    reachability_checker'.refine[OF UPPAAL_Reachability_Problem_precompiled'_axioms]
+    Alw_ev_checker.refine[OF UPPAAL_Reachability_Problem_precompiled'_axioms]
+    leadsto_checker.refine[OF UPPAAL_Reachability_Problem_precompiled'_axioms]
+  ]
+
+end
+
+concrete_definition model_checker uses
+  UPPAAL_Reachability_Problem_precompiled'.model_checker_def_refined
+
+definition [code]:
+  "precond_mc p m k max_steps I T prog final bounds P s0 na \<equiv>
+    if UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k
+    then
+      model_checker p m max_steps I T prog bounds P s0 na k final
+      \<bind> (\<lambda> x. return (Some x))
+    else return None"
+
+theorem model_check:
+  "<emp> precond_mc p m k max_steps I T prog formula bounds P s0 na
+    <\<lambda> Some r \<Rightarrow> \<up>(
+        UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k \<and>
+        (\<not> has_deadlock (conv (N p I P T prog bounds)) max_steps (repeat 0 p, s0, \<lambda>_ . 0) \<longrightarrow>
+          r = conv (N p I P T prog bounds),(repeat 0 p, s0, \<lambda>_ . 0) \<Turnstile>max_steps formula
+        ))
+     | None \<Rightarrow> \<up>(\<not> UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k)
+    >t"
+proof -
+  define A where "A \<equiv> conv (N p I P T prog bounds)"
+  define check where "check \<equiv> A,(repeat 0 p, s0, \<lambda>_ . 0) \<Turnstile>max_steps formula"
+  note [sep_heap_rules] =
+    UPPAAL_Reachability_Problem_precompiled'.model_check'_hoare[
+      of p m max_steps I T prog bounds P s0 na k formula,
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      folded A_def check_def has_deadlock_def
+      ]
+  show ?thesis
+    unfolding UPPAAL_Reachability_Problem_precompiled_defs.init_def
+    unfolding A_def[symmetric] check_def[symmetric]
+    unfolding precond_mc_def by (sep_auto simp: model_checker.refine[symmetric])
+qed
+
+theorem model_check_alt:
+  "<emp> precond_mc p m k max_steps I T prog formula bounds P s0 na
+    <\<lambda> r. \<up> (
+    if UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k
+    then r \<noteq> None \<and>
+      (\<not> has_deadlock (conv (N p I P T prog bounds)) max_steps (repeat 0 p, s0, \<lambda>_ . 0) \<longrightarrow>
+      r = Some (
+        conv (N p I P T prog bounds),(repeat 0 p, s0, \<lambda>_ . 0) \<Turnstile>max_steps formula
+      ))
+    else r = None
+    )>t"
+proof -
+  define A where "A \<equiv> conv (N p I P T prog bounds)"
+  define check where "check \<equiv> A,(repeat 0 p, s0, \<lambda>_ . 0) \<Turnstile>max_steps formula"
+  note [sep_heap_rules] =
+    UPPAAL_Reachability_Problem_precompiled'.model_check'_hoare[
+      of p m max_steps I T prog bounds P s0 na k formula,
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      folded A_def check_def has_deadlock_def
+      ]
+  show ?thesis
+    unfolding UPPAAL_Reachability_Problem_precompiled_defs.init_def
+    unfolding A_def[symmetric] check_def[symmetric]
+    unfolding precond_mc_def by (sep_auto simp: model_checker.refine[symmetric])
+qed
+
+prepare_code_thms dfs_map_impl'_def leadsto_impl_def
+
+lemmas [code] =
+  reachability_checker'_def
+  Alw_ev_checker_def
+  leadsto_checker_def
+  model_checker_def[unfolded UPPAAL_Reachability_Problem_precompiled_defs.F_def PR_CONST_def]
+
+export_code
+  precond_mc Pure.type init_pred_check time_indep_check1 time_indep_check1 conjunction_check2
+  checking SML
+
+end (* Theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks.thy
@@ -0,0 +1,1464 @@
+theory UPPAAL_State_Networks
+  imports Munta_Model_Checker.State_Networks Timed_Automata.Normalized_Zone_Semantics UPPAAL_Asm_Clocks
+    AutoCorres2.Subgoals
+begin
+
+section \<open>Networks of Timed Automata -- UPPAAL Style\<close>
+
+text \<open>Networks of Timed Automata with Shared State and UPPAAL-style Assembler guards and updates.\<close>
+
+no_notation Ref.update ("_ := _" 62)
+no_notation fun_rel_syn (infixr "\<rightarrow>" 60)
+
+lemma finite_lists_boundedI:
+  assumes "\<forall> i < r. finite (S i)"
+    shows "finite {s. length s = r \<and> (\<forall>i<r. s ! i \<in> S i)}" (is "finite ?R")
+proof -
+  let ?S = "\<Union> {S i | i. i < r}"
+  have "?R \<subseteq> {s. set s \<subseteq> ?S \<and> length s = r}"
+    by (auto dest!: mem_nth)
+  moreover have "finite \<dots>" by (rule finite_lists_length_eq) (use assms in auto)
+  ultimately show ?thesis by (rule finite_subset)
+qed
+
+subsection \<open>Syntax and Operational Semantics\<close>
+
+text \<open>
+  We formalize Networks of Timed Automata with integer variable state using UPPAAL-style
+  guards and updates. The specification language for guards and updates is our formalization of
+  the UPPAAL like Assembler language.
+  We extend Networks of Timed Automata with arbitrary shared (global) state.
+  Syntactically, this extension is very simple.
+  We can just use the free action label slot to annotate edges with a guard
+  and an update function on discrete states.
+  The slightly more clumsy part is adding invariants for discrete states
+  by directly specifying an invariant annotating function.
+\<close>
+
+type_synonym
+  ('c, 'time, 's) invassn = "'s \<Rightarrow> ('c, 'time) cconstraint"
+
+type_synonym
+  ('a, 's) transition = "'s * addr * 'a * addr * 's"
+
+type_synonym
+  ('a, 'c, 'time, 's) uta = "('a, 's) transition set * ('c, 'time, 's) invassn"
+
+type_synonym
+  ('a, 'time, 's) unta =
+  "'time programc \<times> ('a act, nat, 'time, 's) uta list \<times> ('s \<Rightarrow> addr) list \<times> (int * int) list"
+
+definition
+  "bounded bounds s \<equiv>
+   length s = length bounds \<and> (\<forall> i < length s. fst (bounds ! i) < s ! i \<and> s ! i < snd (bounds ! i))"
+
+inductive step_u ::
+  "('a, 't :: time, 's) unta \<Rightarrow> nat \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval \<Rightarrow> 'a label
+  \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>_ \<langle>_, _, _\<rangle> \<rightarrow>_ \<langle>_, _, _\<rangle>" [61,61,61,61,61,61] 61)
+where
+  step_u_t:
+    "\<lbrakk>
+      \<forall> p < length N. \<exists> pc st s' rs.
+        stepst P n (u \<oplus> d) ((I ! p) (L ! p), [], s, True, []) (pc, st, s', True, rs);
+      \<forall> p < length N. u \<oplus> d \<turnstile> snd (N ! p) (L ! p);
+      d \<ge> 0;
+      bounded B s
+     \<rbrakk>
+    \<Longrightarrow> (P, N, I, B) \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>Del \<langle>L, s, u \<oplus> d\<rangle>" |
+  step_u_i:
+    "\<lbrakk>
+      stepst P n u (pc_g, [], s, True, []) (_, _, _, True, _);
+      stepst P n u (pc_u, [], s, True, []) (_, _, s', _, r);
+      \<forall> p < length N. \<exists> pc st s rs.
+        stepst P n u' ((I ! p) (L' ! p), [], s', True, []) (pc, st, s, True, rs);
+      (l, pc_g, Sil a, pc_u, l') \<in> fst (N ! p);
+      \<forall> p < length N. u' \<turnstile> snd (N ! p) (L' ! p);
+      L!p = l; p < length L; L' = L[p := l']; u' = [r\<rightarrow>0]u;
+      bounded B s'
+     \<rbrakk>
+    \<Longrightarrow> (P, N, I, B) \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>Act a \<langle>L', s', u'\<rangle>" |
+  step_u_s:
+    "\<lbrakk>
+      stepst P n u (pc_g1, [], s, True, []) (_, _, _, True, _);
+      stepst P n u (pc_g2, [], s, True, []) (_, _, _, True, _);
+      stepst P n u (pc_u2, [], s, True, []) (_, _, s1, _, r2);
+      \<comment>\<open>UPPAAL semantics quirk\<close>
+      ((\<exists> pc st s' f. stepst P n u (pc_u1, [], s, True, []) (pc, st, s', f, r1))
+        \<or> (\<not> (\<exists> pc st s' f r'. stepst P n u (pc_u1, [], s, True, []) (pc, st, s', f, r')) \<and> r1 = []));
+      stepst P n u (pc_u1, [], s1, True, []) ( _, _, s', _, _);
+      cancel\<open>stepst P n u (pc_u2, [], s1, True, []) ( _, _, s2, _, r2);\<close>
+      \<forall> p < length N. \<exists> pc st s rs.
+        stepst P n u' ((I ! p) (L' ! p), [], s', True, []) (pc, st, s, True, rs);
+      (l1, pc_g1, In a, pc_u1, l1') \<in> fst (N ! p);
+      (l2, pc_g2, Out a, pc_u2, l2') \<in> fst (N ! q);
+      \<forall> p < length N. u' \<turnstile> snd (N ! p) (L' ! p);
+      L!p = l1; L!q = l2; p < length L; q < length L; p \<noteq> q;
+      L' = L[p := l1', q := l2']; u' = [(r1 @ r2)\<rightarrow>0]u;
+      bounded B s'
+     \<rbrakk> \<Longrightarrow> (P, N, I, B) \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>Syn a \<langle>L', s', u'\<rangle>"
+
+inductive_cases[elim!]: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+
+inductive steps_un ::
+  "('a, 't :: time, 's) unta \<Rightarrow> nat \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval
+  \<Rightarrow> 's list \<Rightarrow> int list \<Rightarrow> (nat, 't) cval \<Rightarrow> bool"
+("_ \<turnstile>_ \<langle>_, _, _\<rangle> \<rightarrow>* \<langle>_, _, _\<rangle>" [61,61,61,61,61,61] 61)
+where
+  refl: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L, s, u\<rangle>" |
+  step: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<Longrightarrow> A \<turnstile>n  \<langle>L', s', u'\<rangle> \<rightarrow>a \<langle>L'', s'', u''\<rangle>
+        \<Longrightarrow> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>"
+
+declare steps_un.intros[intro]
+
+lemma stepI2:
+  "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>" if
+  "A \<turnstile>n \<langle>L', s', u'\<rangle> \<rightarrow>* \<langle>L'', s'', u''\<rangle>" "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  using that
+  apply induction
+   apply (rule steps_un.step)
+    apply (rule refl)
+   apply assumption
+  apply simp
+  by (rule steps_un.step; assumption)
+
+subsection \<open>Equivalent State Network Automaton\<close>
+
+(*
+abbreviation state_set :: "('a, 'c, 'time, 's) transition set \<Rightarrow> 's set" where
+  "state_set T \<equiv> fst ` T \<union> (snd o snd o snd o snd) ` T"
+*)
+
+definition "stripp p \<equiv> map_option strip o p"
+definition "stripfp p \<equiv> map_option stripf o p"
+definition "striptp p \<equiv> map_option stript o p"
+
+locale Equiv_TA_Defs =
+  fixes A :: "('a, 't, 's) unta"
+    and n :: nat \<comment> \<open>Fuel\<close>
+begin
+
+abbreviation "N \<equiv> fst (snd A)"
+abbreviation "P \<equiv> fst A"
+abbreviation "I \<equiv> fst (snd (snd A))"
+abbreviation "B \<equiv> snd (snd (snd A))"
+abbreviation "P' \<equiv> stripfp P"
+abbreviation "PF \<equiv> stripfp P"
+abbreviation "PT \<equiv> striptp P"
+definition "p \<equiv> length N"
+
+definition "make_f pc_u \<equiv> \<lambda> s.
+  case (exec P' n (pc_u, [], s, True, []) []) of
+    None \<Rightarrow> []
+  | Some ((_, _, _, _, r), _) \<Rightarrow> r"
+
+definition "make_mt pc_u \<equiv> \<lambda> s.
+  case (exec PT n (pc_u, [], s, True, []) []) of
+    None \<Rightarrow> None
+  | Some ((_, _, s', _, r), _) \<Rightarrow> Some s'"
+
+definition "make_mf pc_u \<equiv> \<lambda> s.
+  case (exec PF n (pc_u, [], s, True, []) []) of
+    None \<Rightarrow> None
+  | Some ((_, _, s', _, r), _) \<Rightarrow> Some s'"
+
+definition "make_c pc_g \<equiv> \<lambda> s.
+  case (exec PT n (pc_g, [], s, True, []) []) of
+    None \<Rightarrow> False
+  | Some ((_, _, _, f, _), _) \<Rightarrow> f"
+
+definition "make_g pc_g \<equiv> \<lambda> s.
+  case (exec PT n (pc_g, [], s, True, []) []) of
+    None \<Rightarrow> []
+  | Some ((_, _, _, _, _), pcs) \<Rightarrow>
+      List.map_filter (\<lambda> pc.
+        case P pc of
+          Some (CEXP ac) \<Rightarrow> Some ac
+        | _ \<Rightarrow> None
+          )
+        pcs"
+
+definition "
+  state_trans_t i \<equiv>
+    {(l, make_g pc_g, (a, make_c pc_g, make_mf pc_u), make_f pc_u, l') | l a l' pc_g pc_u.
+      (l, pc_g, a, pc_u, l') \<in> fst (N ! i)
+    }
+  "
+
+  (*
+definition "
+  state_trans_f i \<equiv>
+    {(l, \<lambda> i. [], (a, make_c pc_g, make_mf pc_u), make_f pc_u, l') | l a l' pc_g pc_u.
+      (l, pc_g, a, pc_u, l') \<in> fst (N ! i)
+    }
+  "
+*)
+
+(* definition "state_trans i \<equiv> state_trans_f i \<union> state_trans_t i" *)
+
+abbreviation "state_trans \<equiv> state_trans_t"
+
+definition "
+  state_pred i \<equiv> \<lambda> l s.
+    case (exec P' n ((I ! i) l, [], s, True, []) []) of
+      None \<Rightarrow> False
+    | Some ((_, _, _, f, _), _) \<Rightarrow> f \<and> bounded B s
+  "
+
+definition "
+  state_inv i \<equiv> snd (N ! i)
+"
+
+definition "
+  state_ta \<equiv> (map (\<lambda> p. (state_trans p, state_inv p)) [0..<p], map state_pred [0..<p])
+"
+
+sublocale defs: Prod_TA_Defs state_ta .
+
+lemma bounded_finite:
+    "finite {s. bounded B s}" (is "finite ?S")
+proof -
+  have
+    "?S \<subseteq> {s. length s = length B \<and> (\<forall>i<length B. fst (B ! i) < s ! i \<and> s ! i < snd (B ! i))}"
+    unfolding bounded_def by auto
+  moreover have "finite \<dots>" unfolding bounded_def using finite_lists_boundedI by force
+  ultimately show "finite ?S" by (rule finite_subset)
+qed
+
+(* Unused *)
+lemma finite_state:
+    "\<forall> q < p. \<forall> l. finite {s. (defs.P ! q) l s}"
+proof safe
+  fix q l assume \<open>q < p\<close>
+  let ?S = "{s. (defs.P ! q) l s}"
+  from \<open>q < p\<close> have "?S \<subseteq> {s. bounded B s}"
+    unfolding state_ta_def state_pred_def by (auto split: option.splits)
+  moreover have "finite \<dots>" by (rule bounded_finite)
+  ultimately show "finite ?S" by (rule finite_subset)
+qed
+
+end (* End of definitions locale *)
+
+fun is_instr :: "'t instrc \<Rightarrow> bool" where
+  "is_instr (INSTR _) = True" |
+  "is_instr _ = False"
+
+lemma step_stripf:
+  assumes
+    "is_instr cmd"
+  shows
+    "stepc cmd u (pc, st, s, f, rs) = step (stripf cmd) (pc, st, s, f, rs)"
+proof (cases cmd)
+  case (INSTR instr)
+  with assms(1) show ?thesis
+    by (cases instr) (auto split: option.split)
+next
+  case (CEXP x2)
+  with assms show ?thesis by auto
+qed
+
+lemma step_stript:
+  assumes
+    "is_instr cmd"
+  shows
+    "stepc cmd u (pc, st, s, f, rs) = step (stript cmd) (pc, st, s, f, rs)"
+proof (cases cmd)
+  case (INSTR instr)
+  with assms(1) show ?thesis
+    by (cases instr) (auto split: option.split)
+next
+  case (CEXP x2)
+  with assms show ?thesis by auto
+qed
+
+(* Move? *)
+lemmas [intro] = stepsc.intros
+
+lemma stepsc_f_complete:
+  assumes
+    "stepsc P n' u start end"
+    "\<And> pc' st s' f' rs cmd.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some cmd
+\<Longrightarrow> is_instr cmd"
+  shows
+    "steps (stripfp P) n' start end"
+  using assms proof (induction P \<equiv> P n' u \<equiv> u x4 \<equiv> start "end" arbitrary: start rule: stepsc.induct)
+    case 1
+    then show ?case unfolding stripfp_def by auto
+  next
+    case (2 cmd pc st m f rs s n' s')
+    have "is_instr cmd" if
+      "stepsc P n' u s (pc', st, s', f', rs)" "P pc' = Some cmd" for pc' st s' f' rs cmd
+      using 2(1,2) that by (auto intro: 2(5))
+    with 2(4) have *: "steps (stripfp P) n' s s'" by auto
+    show ?case
+    proof (cases cmd)
+      case (INSTR instr)
+      with 2(1) step_stripf have
+        "step (stripf cmd) (pc, st, m, f, rs) = Some s"
+        by (auto split: option.split_asm)
+      with 2(1-3) 2(5-) * show ?thesis unfolding stripfp_def by auto
+    next
+      case (CEXP ac)
+      with 2 show ?thesis by fastforce
+    qed
+  qed
+
+lemma stepsc_f_sound:
+  assumes
+    "steps (stripfp P) n' start end"
+    "\<And> pc' st s' f' rs cmd.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some cmd
+    \<Longrightarrow> is_instr cmd"
+  shows
+    "stepsc P n' u start end"
+using assms proof (induction "stripfp P" n' start "end")
+  case (1 n start)
+  then show ?case by auto
+next
+  case (2 instr pc st m f rs s n s')
+  from 2(2) obtain cmd where "P pc = Some cmd" unfolding stripfp_def by auto
+  show ?case
+  proof (cases cmd)
+    case prems: (INSTR instr)
+    with \<open>P pc = _\<close> 2(2) step_stripf[of cmd, symmetric] 2(1) have step:
+      "stepc cmd u (pc, st, m, f, rs) = Some s"
+      unfolding stripfp_def by auto
+    with \<open>P pc = _\<close> have "is_instr cmd" if
+      "stepsc P n u s (pc', st, s', f', rs)" "P pc' = Some cmd" for pc' st s' f' rs cmd
+      using that unfolding stripfp_def by (force intro: 2(5))
+    with 2(4) have "stepsc P n u s s'" by auto
+    with step \<open>P pc = _\<close> show ?thesis unfolding stripfp_def by auto
+  next
+    case prems: (CEXP ac)
+    from \<open>P pc = _\<close> 2(5) have "is_instr cmd" by blast
+    with prems show ?thesis by auto
+  qed
+qed
+
+definition
+  "time_indep P n start \<equiv>
+    \<forall> pc' st s' f' rs cmd u.
+      stepsc P n u start (pc', st, s', f', rs) \<and> P pc' = Some cmd
+  \<longrightarrow> is_instr cmd"
+
+lemma stepsc_t_complete:
+  assumes
+    "stepsc P n' u start end"
+    "\<And> pc' st s' f' rs ac.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some (CEXP ac) \<Longrightarrow> u \<turnstile>a ac"
+  shows
+    "steps (striptp P) n' start end"
+  using assms proof (induction P \<equiv> P n' u \<equiv> u x4 \<equiv> start "end" arbitrary: start rule: stepsc.induct)
+    case 1
+    then show ?case unfolding stripfp_def by auto
+  next
+    case (2 cmd pc st m f rs s n' s')
+    have "u \<turnstile>a ac" if
+      "stepsc P n' u s (pc', st, s', f', rs)" "P pc' = Some (CEXP ac)" for pc' st s' f' rs ac
+      using 2(1,2) that by (auto intro: 2(5))
+    with 2(4) have *: "steps (striptp P) n' s s'" by auto
+    show ?case
+    proof (cases cmd)
+      case (INSTR instr)
+      with 2(1) step_stript have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        by (auto split: option.split_asm)
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by auto
+    next
+      case (CEXP ac)
+      with 2(1-3) have "u \<turnstile>a ac" by (auto intro: 2(5))
+      with 2(1) have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        using \<open>cmd = _\<close> by auto
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by auto
+    qed
+  qed
+
+lemma stepsc_t_complete2:
+  assumes
+    "stepsc P n' u start (pc', st', s', f', rs')"
+    "\<And> pc' st s' f' rs ac.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some (CEXP ac) \<Longrightarrow> u \<turnstile>a ac"
+  shows
+    "steps (striptp P) n' start (pc', st', s', f', rs') \<and> (\<forall> ac. P pc' = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac)"
+  using assms
+  proof (induction P \<equiv> P n' u \<equiv> u x4 \<equiv> start "(pc', st', s', f', rs')" arbitrary: start rule: stepsc.induct)
+    case 1
+    then show ?case unfolding stripfp_def by blast
+  next
+    case (2 cmd pc st m f rs s n')
+    have "u \<turnstile>a ac" if
+      "stepsc P n' u s (pc', st, s', f', rs)" "P pc' = Some (CEXP ac)" for pc' st s' f' rs ac
+      using 2(1,2) that by (auto intro: 2(5))
+    with 2(4) have *:
+      "steps (striptp P) n' s (pc', st', s', f', rs')" "\<forall>ac. P pc' = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac"
+      by auto
+    show ?case
+    proof (cases cmd)
+      case (INSTR instr)
+      with 2(1) step_stript have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        by (auto split: option.split_asm)
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by auto
+    next
+      case (CEXP ac)
+      with 2(1-3) have "u \<turnstile>a ac" by (auto intro: 2(5))
+      with 2(1) have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        using \<open>cmd = _\<close> by auto
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by auto
+    qed
+  qed
+
+lemma stepsc_t_visitedc:
+  assumes
+    "stepsc P n' u start end"
+    "\<And> pc' st s' f' rs.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> Q pc'"
+  shows "\<exists> pcs. visitedc P n' u start end pcs
+  \<and> (\<forall> pc \<in> set pcs. Q pc)"
+  using assms by (induction) (fastforce intro!: visitedc.intros)+
+
+lemma visitedc_t_visited:
+  assumes
+    "visitedc P n' u start end pcs"
+    "\<And> pc' ac. pc' \<in> set pcs \<Longrightarrow> P pc' = Some (CEXP ac) \<Longrightarrow> u \<turnstile>a ac"
+  shows
+    "visited (striptp P) n' start end pcs
+  \<and> (\<forall> pc ac. pc' \<in> set pcs \<and> P pc' = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac)"
+  using assms
+  proof (induction P \<equiv> P n' u \<equiv> u x4 \<equiv> start "end" pcs arbitrary: start rule: visitedc.induct)
+    case 1
+    then show ?case by (auto intro: visited.intros)
+  next
+    case (2 cmd pc st m f rs s n' s' pcs)
+    have "u \<turnstile>a ac" if
+      "pc' \<in> set pcs" "P pc' = Some (CEXP ac)" for pc' ac
+      using 2(1,2) that by (auto intro: 2(5))
+    with 2(4) have *:
+      "visited (striptp P) n' s s' pcs" "\<forall>pc ac. pc' \<in> set pcs \<and> P pc' = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac"
+      by auto
+    show ?case
+    proof (cases cmd)
+      case (INSTR instr)
+      with 2(1) step_stript have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        by (auto split: option.split_asm)
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by (auto intro: visited.intros)
+    next
+      case (CEXP ac)
+      with 2(1-3) have "u \<turnstile>a ac" by (auto intro: 2(5))
+      with 2(1) have
+        "step (stript cmd) (pc, st, m, f, rs) = Some s"
+        using \<open>cmd = _\<close> by auto
+      with 2(1-3) 2(5-) * show ?thesis unfolding striptp_def by (auto intro: visited.intros)
+    qed
+  qed
+
+lemma stepsc_t_sound:
+  assumes
+    "steps (striptp P) n' start end"
+    "\<And> pc' st s' f' rs ac.
+    stepsc P n' u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some (CEXP ac) \<Longrightarrow> u \<turnstile>a ac"
+  shows
+    "stepsc P n' u start end"
+using assms proof (induction "striptp P" n' start "end")
+  case (1 n start)
+  then show ?case by auto
+next
+  case (2 instr pc st m f rs s n s')
+  from 2(2) obtain cmd where "P pc = Some cmd" unfolding striptp_def by auto
+  show ?case
+  proof (cases cmd)
+    case prems: (INSTR instr)
+    with \<open>P pc = _\<close> 2(2) step_stript[of cmd, symmetric] 2(1) have step:
+      "stepc cmd u (pc, st, m, f, rs) = Some s"
+      unfolding striptp_def by auto
+    with \<open>P pc = _\<close> have "u \<turnstile>a ac" if
+      "stepsc P n u s (pc', st, s', f', rs)" "P pc' = Some (CEXP ac)" for pc' st s' f' rs ac
+      using that unfolding striptp_def by (force intro: 2(5))
+    with 2(4) have "stepsc P n u s s'" by auto
+    with step \<open>P pc = _\<close> show ?thesis unfolding striptp_def by auto
+  next
+    case prems: (CEXP ac)
+    with \<open>P pc = _\<close> 2(2) have [simp]: "P pc = Some (CEXP ac)" by auto
+    then have "u \<turnstile>a ac" by (auto intro: 2(5))
+    with 2(2,1) \<open>cmd = _\<close> have step:
+      "stepc cmd u (pc, st, m, f, rs) = Some s"
+      unfolding striptp_def by auto
+    with \<open>P pc = Some cmd\<close> have "u \<turnstile>a ac" if
+      "stepsc P n u s (pc', st, s', f', rs)" "P pc' = Some (CEXP ac)" for pc' st s' f' rs ac
+      using that unfolding striptp_def by (force intro: 2(5))
+    with 2(4) have "stepsc P n u s s'" by auto
+    with step \<open>P pc = Some cmd\<close> show ?thesis unfolding striptp_def by auto
+  qed
+qed
+
+lemma stepsc_visitedc:
+  "\<exists> cc. visitedc P n u start end cc" if "stepsc P n u start end"
+  using that by induction (auto intro: visitedc.intros)
+
+lemma visitedc_stepsc:
+  "stepsc P n u start end" if "visitedc P n u start end cc"
+  using that by (induction; blast)
+
+lemma steps_visited:
+  "\<exists> cc. visited P n start end cc" if "steps P n start end"
+  using that by induction (auto intro: visited.intros)
+
+lemma visited_steps:
+  "steps P n start end" if "visited P n start end cc"
+  using that by (induction; blast)
+
+context
+  fixes P n u start
+  assumes constraints_conj: "\<forall> pc' st s' f' rs ac.
+    stepsc P n u start (pc', st, s', f', rs) \<and> P pc' = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac"
+begin
+
+  lemma stepsc_t_sound':
+    assumes
+      "steps (striptp P) n start end"
+    shows
+      "stepsc P n u start end"
+    using assms constraints_conj by (auto intro: stepsc_t_sound)
+
+  lemma stepsc_t_complete':
+    assumes
+      "stepsc P n u start end"
+    shows
+      "steps (striptp P) n start end"
+    using assms constraints_conj by (auto intro: stepsc_t_complete)
+
+  lemma stepsc_t_complete'':
+    assumes
+      "stepsc P n u start end"
+    shows
+      "\<exists> pcs. visitedc P n u start end pcs \<and> (\<forall> pc \<in> set pcs. \<forall> ac. P pc = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac)"
+    using assms constraints_conj by (auto intro: stepsc_t_complete stepsc_t_visitedc)
+
+  lemma stepsc_t_visited:
+    assumes
+      "stepsc P n u start end"
+    shows
+      "\<exists> pcs. visited (striptp P) n start end pcs \<and> (\<forall> pc \<in> set pcs. \<forall> ac. P pc = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac)"
+    using stepsc_t_complete''[OF assms] visitedc_t_visited by blast
+
+  lemma stepst_t_complete:
+    assumes
+      "stepst P n u start end"
+    shows
+      "\<exists> pcs. exec (striptp P) n start [] = Some (end, pcs) \<and> (\<forall> pc \<in> set pcs. \<forall> ac. P pc = Some (CEXP ac) \<longrightarrow> u \<turnstile>a ac)"
+      using assms by (auto dest!: stepsc_t_visited visited_exec' simp: striptp_def stepst_def)
+
+  lemma stepst_t_equiv:
+    "(\<exists> pcs'. exec (striptp P) n start pcs = Some ((pc, st, m, f, rs), pcs'))
+    \<longleftrightarrow> stepst P n u start (pc, st, m, f, rs)"
+    apply safe
+     apply (drule exec_steps)
+    unfolding stepst_def
+     apply safe
+      apply (rule stepsc_t_sound'; assumption)
+     apply (simp add: striptp_def)
+     apply safe
+    subgoal for z
+      by (cases z) auto
+    by (auto dest!: stepsc_t_complete' steps_exec simp: striptp_def)
+
+end
+
+context
+  fixes P n start
+  assumes time_indep: "time_indep P n start"
+begin
+
+  lemma time_indep':
+    "\<And> pc' st s' f' rs cmd.
+      stepsc P n u start (pc', st, s', f', rs) \<Longrightarrow> P pc' = Some cmd
+  \<Longrightarrow> is_instr cmd"
+    using time_indep unfolding time_indep_def by blast
+
+  lemma stepsc_f_complete':
+    assumes
+      "stepsc P n u start end"
+    shows
+      "steps (stripfp P) n start end"
+    using assms time_indep' by (auto intro: stepsc_f_complete[where P = P])
+
+  lemma stepsc_f_sound':
+    assumes
+      "steps (stripfp P) n start end"
+    shows
+      "stepsc P n u start end"
+    using assms time_indep' by (auto intro: stepsc_f_sound[where P = P])
+
+  lemma stepsc_f_equiv:
+    "steps (stripfp P) n start end \<longleftrightarrow> stepsc P n u start end"
+    using stepsc_f_sound' stepsc_f_complete' by fast
+
+  lemma stepst_f_equiv:
+    "(\<exists> pcs'. exec (stripfp P) n start pcs = Some ((pc, st, m, f, rs), pcs'))
+    \<longleftrightarrow> stepst P n u start (pc, st, m, f, rs)"
+    apply safe
+     apply (drule exec_steps)
+    unfolding stepst_def
+     apply safe
+      apply (rule stepsc_f_sound'; assumption)
+     apply (simp add: stripfp_def)
+     apply safe
+    subgoal for z
+      by (cases z) auto
+    by (auto dest!: stepsc_f_complete' steps_exec simp: stripfp_def)
+
+end (* End of context for equivalence *)
+
+lemma exec_acc:
+  assumes "exec P n s pcs = Some (s', pcs')"
+  shows "\<exists> pcs''. pcs' = pcs'' @ pcs"
+  using assms by (induction P n s pcs rule: exec.induct; force split: option.split_asm if_split_asm)
+
+lemma exec_acc':
+  assumes "Some (s', pcs') = exec P n s pcs"
+  shows "\<exists> pcs''. pcs' = pcs'' @ pcs"
+  using assms
+  using assms exec_acc by metis
+
+lemma exec_min_steps:
+  assumes "exec P n s pcs = Some (s', pcs' @ pcs)"
+  shows "exec P (length pcs') s pcs = Some (s', pcs' @ pcs)"
+using assms proof (induction n arbitrary: s pcs s' pcs')
+  case 0
+  then show ?case by auto
+next
+  case (Suc n)
+  obtain pc st m f rs pc' st' m' f' rs' where [simp]:
+    "s = (pc, st, m, f, rs)" "s' = (pc', st', m', f', rs')"
+    using prod.exhaust by metis
+  from Suc obtain instr where "P pc = Some instr" by (auto split: option.splits if_splits)
+  show ?case
+  proof (cases "instr = HALT")
+    case True
+    with \<open>P pc = _\<close> Suc show ?thesis by auto
+  next
+    case False
+    with Suc.prems \<open>P pc = _\<close> obtain s'' where s'':
+      "step instr s = Some s''" "exec P n s'' (pc # pcs) = Some (s', pcs' @ pcs)"
+      by (auto split: option.splits)
+    with exec_acc[OF this(2)] obtain pcs'' where "pcs' = pcs'' @ [pc]" by auto
+    with Suc.IH[of s'' "pc # pcs" s' "pcs''"] \<open>P pc = _\<close> False s'' show ?thesis by auto
+  qed
+qed
+
+lemma exec_steps_visited:
+  assumes
+    "exec P (length pcs') s pcs = Some (s', pcs' @ pcs)"
+    "steps P (length pcs') s (pc, st, m, f, rs)"
+  shows "pc \<in> set pcs'"
+  using assms proof (induction P \<equiv> P "length pcs'" s pcs arbitrary: pc st m f rs pcs' rule: exec.induct)
+  case 1
+  then show ?case by simp
+next
+  case (2 n pc' st' m' f' rs' pcs pcs')
+  from this(2)[symmetric] this(3) obtain instr where "P pc' = Some instr" by (cases "P pc'") auto
+  show ?case
+  proof (cases "instr = HALT")
+    case True
+    with "2.prems" \<open>P pc' = _\<close> \<open>Suc n = _\<close>[symmetric] show ?thesis by (force elim: steps.cases)
+  next
+    case False
+    with 2 obtain pcs'' where "pcs' = pcs'' @ [pc']"
+      apply atomize_elim
+      by (erule exec.elims) (auto dest: exec_acc' split: option.split_asm if_split_asm)
+    with False 2 \<open>P pc' = _\<close> \<open>Suc n = _\<close>[symmetric] show ?thesis
+      by (auto split: option.split_asm elim: steps.cases)
+  qed
+qed
+
+lemma stepsc_mono:
+  assumes "stepsc P n u start end" "n' \<ge> n"
+  shows "stepsc P n' u start end"
+  using assms proof (induction arbitrary: n')
+  case (1 prog n u start)
+  then show ?case by (cases n') auto
+next
+  case (2 cmd u pc st m f rs s prog n s')
+  then show ?case by (cases n') auto
+qed
+
+lemma stepst_mono:
+  assumes "stepst P n u start end" "n' \<ge> n"
+  shows "stepst P n' u start end"
+    using assms stepsc_mono unfolding stepst_def by blast
+
+definition
+  "state_indep P n \<equiv>
+    \<forall> pc f pc' st f' rs u s1 s1' s2 s2' rs1 rs2.
+      stepsc P n u (pc, st, s1, f, rs) (pc', st, s1', f', rs1) \<and>
+      stepsc P n u (pc, st, s2, f, rs) (pc', st, s2', f', rs2)
+  \<longrightarrow> rs1 = rs2"
+
+lemma exec_len:
+  "n \<ge> (length pcs' - length pcs)" if "exec P n s pcs = Some (s', pcs')"
+  using that
+  by (induction P n s pcs arbitrary: rule: exec.induct)
+     (force split: option.split_asm if_split_asm)+
+
+lemma steps_striptp_stepsc:
+  assumes "
+    \<And> pc st m f rs ac.
+      steps (striptp P) n s' (pc, st, m, f, rs) \<Longrightarrow> P pc = Some (CEXP ac)
+      \<Longrightarrow> u' \<turnstile>a ac
+    "
+    and "steps (striptp P) n s' s''"
+  shows "stepsc P n u' s' s''"
+  using assms(2,1)
+proof (induction "striptp P" n s' s'')
+  case (1 n start)
+  show ?case by (rule stepsc.intros)
+next
+  case (2 cmd pc st m f rs s n s')
+  have "u' \<turnstile>a ac"
+    if "P pc = Some (CEXP ac)" "UPPAAL_Asm.steps (striptp P) n s (pc, st, m, f, rs)"
+    for pc st m f rs ac
+    using that 2(1-3) by - (rule 2(5); force)
+  with 2(4) have "stepsc P n u' s s'" by auto
+  with 2(1-3) show ?case
+    apply (cases "P pc")
+     apply (simp add: striptp_def)
+    subgoal for cmd'
+      apply (cases cmd')
+      subgoal
+        by (force split: option.split simp: striptp_def)
+      subgoal
+        by (rule stepsc.intros) (auto intro!: 2(5) simp: striptp_def)
+      done
+    done
+qed
+
+locale Equiv_TA =
+  Equiv_TA_Defs A n for A :: "('a, 't :: time, 's) unta" and n :: nat +
+  fixes L :: "'s list" and s :: "int list"
+  assumes states[intro]: "L \<in> defs.states' s"
+      (*
+      and pred_time_indep:
+      "\<forall> L' s' u' s''. \<forall> p' < p. \<forall> L'' \<in> defs.states' s''. A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+    \<longrightarrow> time_indep P n ((I ! p') (L'' ! p'), [], s'', True, [])" *)
+    and pred_time_indep:
+      "\<forall> s. \<forall> L \<in> defs.states' s. \<forall> q < p. time_indep P n ((I ! q) (L ! q), [], s, True, [])"
+    and upd_time_indep:
+      "\<forall> l pc_g a l' pc_u s. \<forall> q < p. (l, pc_g, a, pc_u, l') \<in> fst (N ! q)
+    \<longrightarrow> time_indep P n (pc_u, [], s, True, [])"
+    and clock_conj:
+      "\<forall> l pc_g a l' pc_u s u. \<forall> q < p. (l, pc_g, a, pc_u, l') \<in> fst (N ! q) \<and>
+        (\<exists> pc' st s' rs. stepst P n u (pc_g, [], s, True, []) (pc', st, s', True, rs)) \<longrightarrow>
+        (\<forall> pc' st s' f' rs ac.
+        stepsc P n u (pc_g, [], s, True, []) (pc', st, s', f', rs) \<and> P pc' = Some (CEXP ac) \<longrightarrow>
+        u \<turnstile>a ac)"
+    (* Reset clocks may not depend on state. *)
+    (*
+    and "\<forall> l pc_g a pc_u l'. (l, pc_g, In a, pc_u, l') \<in> fst (N ! q) \<longrightarrow> state_indep P pc_u"
+    and "\<forall> l pc_g a pc_u l'. (l, pc_g, Out a, pc_u, l') \<in> fst (N ! q) \<longrightarrow> state_indep P pc_u"
+    *)
+  assumes Len: "length N = length I"
+      and inv: "\<forall> q < p. \<exists> pc st s' rs pcs.
+        exec P' n ((I ! q) (L ! q), [], s, True, []) [] = Some ((pc, st, s', True, rs), pcs)"
+      and bounded: "bounded B s"
+begin
+
+lemma length_defs_N[simp]:
+  "length defs.N = p"
+  unfolding p_def state_ta_def by simp
+
+lemma length_defs_P[simp]:
+  "length defs.P = p"
+  unfolding p_def state_ta_def by simp
+
+lemma inv':
+  "\<forall>p<length defs.P. (defs.P ! p) (L ! p) s"
+  using inv bounded unfolding state_ta_def state_pred_def by (force simp: p_def)
+
+lemma inv'':
+  "\<exists> pc st s' rs. stepst P n u' ((I ! q) (L ! q), [], s, True, []) (pc, st, s', True, rs)"
+  if "q < p" for u'
+proof -
+  from pred_time_indep that have "time_indep P n ((I ! q) (L ! q), [], s, True, [])" by blast
+  with inv that stepst_f_equiv[symmetric] show ?thesis by blast
+qed
+
+lemma A_simp[simp]:
+  "PP = P" "N' = N" "I' = I" "B' = B" if "A = (PP, N', I', B')"
+using that by auto
+
+lemma A_unfold:
+  "A = (P, N, I, B)"
+  by simp
+
+sublocale prod: Prod_TA state_ta by standard (auto simp: inv')
+
+lemma inv_simp:
+  "snd (defs.N ! q) (L' ! q) = snd (N ! q) (L' ! q)" if "q < p" for L'
+  using that unfolding state_ta_def state_inv_def by simp
+
+lemma defs_p_eq[simp]:
+  "defs.p = p"
+  by (simp add: defs.p_def p_def)
+
+lemma ball_lessThan[simp]:
+  "(\<forall> x \<in> {..<m}. Q x) \<longleftrightarrow> (\<forall> x < m. Q x)"
+  by auto
+
+lemma trans_state_taD:
+  assumes "(l, g, (a, c, m), f, l') \<in> fst (defs.N ! q)" "q < p"
+  shows
+    "(l, g, (a, c, m), f, l') \<in> state_trans_t q"
+  using assms unfolding state_ta_def by simp
+
+lemma N_transD:
+  assumes "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)" "q < p"
+  shows "(l, make_g pc_g, (a, make_c pc_g, make_mf pc_u), make_f pc_u, l') \<in> fst (defs.N ! q)"
+    using assms unfolding state_ta_def state_trans_t_def by auto
+
+lemma pred_time_indep':
+  "\<forall> L' s' u'. \<forall> p' < p. A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+\<longrightarrow> time_indep P n ((I ! p') (L' ! p'), [], s', True, [])"
+  using pred_time_indep oops
+
+lemma P_steps_upd:
+  assumes
+    "Some s'' = make_mf pc_u s'" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)" "q < p"
+  shows
+    "\<exists> pc st f. stepst P n u' (pc_u, [], s', True, []) (pc, st, s'', f, make_f pc_u s')"
+proof -
+  from assms upd_time_indep have *: "time_indep P n (pc_u, [], s', True, [])" by auto
+  from assms(1) \<open>q < p\<close> obtain pc st f rs pcs where
+    "exec PF n (pc_u, [], s', True, []) [] = Some ((pc, st, s'', f, rs), pcs)"
+    unfolding make_mf_def state_ta_def by (fastforce split: option.splits)
+  with stepst_f_equiv[OF *] show ?thesis unfolding make_f_def by fastforce
+qed
+
+lemma P_steps_reset:
+  assumes
+    "q < p" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)"
+  shows
+    "(\<exists>pc st s'' f. stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, make_f pc_u s')) \<or>
+       (\<nexists>pc st s'' f r'. stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, r')) \<and>
+       make_f pc_u s' = []"
+proof (cases "\<exists>pc st s'' f r'. stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, r')")
+  case True
+  then obtain pc st s'' f r' where *:
+    "stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, r')"
+    by blast
+  from assms upd_time_indep have "time_indep P n (pc_u, [], s', True, [])" by auto
+  from * stepst_f_equiv[OF this, symmetric, where pcs = "[]"] show ?thesis
+    unfolding make_f_def by (force split: option.split)
+next
+  case False
+  from assms upd_time_indep have "time_indep P n (pc_u, [], s', True, [])" by auto
+  from False stepst_f_equiv[OF this, where pcs = "[]"] show ?thesis
+    unfolding make_f_def by (auto split: option.split)
+qed
+
+lemma steps_P_reset:
+  assumes
+    "(\<exists>pc st s'' f. stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, r)) \<or>
+     (\<nexists>pc st s'' f r'. stepst P n u (pc_u, [], s', True, []) (pc, st, s'', f, r')) \<and> r = []"
+    "q < p" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)"
+  shows "make_f pc_u s' = r"
+  using assms(1)
+proof (safe, goal_cases)
+  case prems: (1 pc st s'' f)
+  from assms upd_time_indep have "time_indep P n (pc_u, [], s', True, [])" by auto
+  from stepst_f_equiv[OF this, symmetric] prems \<open>q < p\<close> obtain pc st f pcs where
+    "exec PF n (pc_u, [], s', True, []) [] = Some ((pc, st, s'', f, r), pcs)"
+    unfolding make_f_def state_ta_def by (fastforce split: option.splits)
+  then show ?case unfolding make_f_def by (auto split: option.split)
+next
+  case prems: 2
+  have "exec PF n (pc_u, [], s', True, []) [] = None"
+  proof (cases "exec PF n (pc_u, [], s', True, []) []")
+    case None
+    then show ?thesis .
+  next
+    case (Some a)
+    obtain pc'' st'' s'' f'' rs'' pcs'' where "a = ((pc'', st'', s'', f'', rs''), pcs'')"
+      by (cases a) auto
+    from assms upd_time_indep have "time_indep P n (pc_u, [], s', True, [])" by auto
+    from stepst_f_equiv[OF this] \<open>_ = Some a\<close> prems \<open>a = _\<close> show ?thesis by auto metis
+  qed
+  then show ?case unfolding make_f_def by simp
+qed
+
+
+lemma steps_P_upd:
+  assumes
+    "stepst P n u' (pc_u, [], s', True, []) (pc, st, s'', f, r)"
+    "q < p" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)"
+  shows
+    "Some s'' = make_mf pc_u s'" (is "?A") "r = make_f pc_u s'" (is "?B")
+proof -
+  from assms upd_time_indep have "time_indep P n (pc_u, [], s', True, [])" by auto
+  from stepst_f_equiv[OF this, symmetric] assms(1) \<open>q < p\<close> obtain pc st f pcs where
+    "exec PF n (pc_u, [], s', True, []) [] = Some ((pc, st, s'', f, r), pcs)"
+    unfolding make_mf_def state_ta_def by (fastforce split: option.splits)
+  then show ?A ?B unfolding make_f_def make_mf_def by (auto split: option.split)
+qed
+
+lemma steps_P_guard:
+  assumes
+    "stepst P n u' (pc_g, [], s', True, []) (pc, st, s'', True, rs)"
+    "q < p" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)"
+  shows
+    "make_c pc_g s'" (is "?A") "u' \<turnstile> make_g pc_g s'" (is "?B")
+proof -
+  from stepst_t_complete[OF _ assms(1)] clock_conj assms obtain pcs where
+    "exec PT n (pc_g, [], s', True, []) [] = Some ((pc, st, s'', True, rs), pcs)"
+    "\<forall> pc\<in>set pcs. \<forall>ac. P pc = Some (CEXP ac) \<longrightarrow> u' \<turnstile>a ac"
+    by fastforce
+  then show ?A ?B unfolding make_c_def make_g_def
+    by (auto split: option.split instrc.split_asm simp: list_all_iff set_map_filter clock_val_def)
+qed
+
+lemma P_steps_guard:
+  assumes
+    "make_c pc_g s'" "u' \<turnstile> make_g pc_g s'"
+    "q < p" "(l, pc_g, a, pc_u, l') \<in> fst (N ! q)"
+  shows
+    "\<exists> pc s'' st rs. stepst P n u' (pc_g, [], s', True, []) (pc, st, s'', True, rs)"
+proof -
+  from assms(1) \<open>q < p\<close> obtain pc st s'' rs pcs where *:
+    "exec PT n (pc_g, [], s', True, []) [] = Some ((pc, st, s'', True, rs), pcs)"
+    unfolding make_c_def state_ta_def by (fastforce split: option.splits)
+  with exec_min_steps[of PT n _ "[]" _ pcs] have **:
+    "exec PT (length pcs) (pc_g, [], s', True, []) [] = Some ((pc, st, s'', True, rs), pcs @ [])"
+    by auto
+  from * assms(2) have
+    "u' \<turnstile> List.map_filter (\<lambda>pc. case P pc of
+        None \<Rightarrow> None
+      | Some (INSTR xa) \<Rightarrow> Map.empty xa
+      | Some (CEXP xa) \<Rightarrow> Some xa) pcs" unfolding make_g_def
+    by auto
+  moreover from ** exec_steps_visited[of PT pcs _ "[]"] have "pc \<in> set pcs"
+    if "steps PT (length pcs) (pc_g, [], s', True, []) (pc, st, m, f, rs)" for pc st m f rs
+    using that by fastforce
+  ultimately have "u' \<turnstile>a ac"
+    if "steps PT (length pcs) (pc_g, [], s', True, []) (pc, st, m, f, rs)" "P pc = Some (CEXP ac)"
+    for pc st m f rs ac
+    using that by (auto 4 3 split: option.splits simp: list_all_iff set_map_filter clock_val_def)
+  moreover from ** have
+    "steps PT (length pcs) (pc_g, [], s', True, []) (pc, st, s'', True, rs)" "PT pc = Some HALT"
+    by (auto dest: exec_steps)
+  ultimately have "stepst P (length pcs) u' (pc_g, [], s', True, []) (pc, st, s'', True, rs)"
+    by (auto intro: steps_striptp_stepsc simp: stepst_def striptp_def elim: stript.elims)
+  moreover from exec_len[OF *] have "n \<ge> length pcs" by simp
+  ultimately show ?thesis by (blast intro: stepst_mono)
+qed
+
+lemma P_bounded:
+  assumes
+    "(defs.P ! q) (L' ! q) s'" "q < p"
+  shows "bounded B s'"
+  using assms unfolding state_pred_def state_ta_def by (auto split: option.splits)
+
+lemma P_steps:
+  assumes
+    "(defs.P ! q) (L' ! q) s'"
+    "q < p" "L' \<in> defs.states' s'"
+  shows
+    "\<exists> pc st s'' rs. stepst P n u' ((I ! q) (L' ! q), [], s', True, []) (pc, st, s'', True, rs)"
+proof -
+  from assms pred_time_indep have *: "time_indep P n ((I ! q) (L' ! q), [], s', True, [])" by auto
+  from assms(1) \<open>q < p\<close> obtain pc st s'' rs pcs where
+    "exec PF n ((I ! q) (L' ! q), [], s', True, []) [] = Some ((pc, st, s'', True, rs), pcs)"
+    unfolding state_pred_def state_ta_def by (auto split: option.splits)
+  with stepst_f_equiv[OF *] show ?thesis by blast
+qed
+
+lemma steps_P:
+  assumes
+    "stepst P n u' ((I ! q) (L' ! q), [], s', True, []) (pc, st, s'', True, rs)"
+    "q < p" "L' \<in> defs.states' s'"
+    "bounded B s'"
+  shows
+    "(defs.P ! q) (L' ! q) s'"
+proof -
+  from assms pred_time_indep have "time_indep P n ((I ! q) (L' ! q), [], s', True, [])" by auto
+  from stepst_f_equiv[OF this] assms(1) obtain pcs' where
+    "exec PF n ((I ! q) (L' ! q), [], s', True, []) [] = Some ((pc, st, s'', True, rs), pcs')"
+    by blast
+  with \<open>q < p\<close> \<open>bounded B s'\<close> show ?thesis unfolding state_pred_def state_ta_def by simp
+qed
+
+lemma P_iff:
+  "(\<exists> pc st rs s''. stepst P n u' ((I ! q) (L' ! q), [], s', True, []) (pc, st, s'', True, rs)
+  \<and> bounded B s')
+  \<longleftrightarrow> (defs.P ! q) (L' ! q) s'" if "q < p" "L' \<in> defs.states' s'"
+  using that by (metis steps_P P_steps P_bounded)
+
+lemma states'_updI':
+  assumes "(L' ! q, g, (a, c, m), f, l') \<in> fst (defs.N ! q)" "L' \<in> defs.states' s''"
+  shows "L'[q := l'] \<in> defs.states' s'"
+  using assms
+  unfolding prod.states'_simp[of s' s'']
+  unfolding Product_TA_Defs.states_def
+  apply clarsimp
+  subgoal for p
+    by (cases "p = q"; force simp: prod.trans_of_N_s_2[simplified] Prod_TA_Defs.N_s_length)
+  done
+
+lemma states'_updI:
+  assumes "(L ! q, g, (a, c, m), f, l') \<in> fst (defs.N ! q)"
+  shows "L[q := l'] \<in> defs.states' s'"
+using assms by (auto intro: states'_updI')
+
+lemma states'_updI'':
+  assumes
+    "(L' ! q, g, (a, c, m), f, l') \<in> fst (defs.N ! q)"
+    "(L' ! q', g', (a', c', m'), f', l'') \<in> fst (defs.N ! q')"
+    "L' \<in> defs.states' s''" "q \<noteq> q'"
+  shows "L'[q := l', q' := l''] \<in> defs.states' s'"
+  using assms by (auto intro: states'_updI')
+
+lemma equiv_sound:
+  assumes step: "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+    shows "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  using step proof cases
+  case (step_sn_t N d I)
+  then show ?thesis
+    apply simp
+    apply (subst A_unfold)
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (rule step_u_t)
+    using inv'' bounded by (auto simp: inv_simp p_def)
+next
+  case (step_sn_i l g a c m f l' N p r I)
+  then show ?thesis
+    apply (simp)
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (simp add: Prod_TA_Defs.N_s_length)
+    apply (subst A_unfold)
+    apply (drule trans_state_taD)
+     apply assumption
+    unfolding state_trans_t_def
+    apply safe
+    apply (drule (2) P_steps_upd)
+    apply (drule (3) P_steps_guard)
+    apply safe
+    apply (rule step_u_i)
+    subgoals \<open>\<forall>p<length local.N. \<exists>pc. _\<close>
+      apply safe
+      apply (rule P_steps)
+        apply (fastforce simp: p_def)
+       apply (fastforce simp: p_def)
+      by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+    by (auto simp: inv_simp p_def Prod_TA_Defs.N_s_length intro!: P_bounded)
+next
+  case (step_sn_s l1 g1 a ci mi f1 l1' N p l2 g2 co mo f2 l2' q r1 r2 I)
+  then show ?thesis
+    apply (simp)
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (simp add: Prod_TA_Defs.N_s_length)
+    apply (subst A_unfold)
+    apply (drule trans_state_taD)
+     apply assumption
+    apply (drule trans_state_taD)
+     apply assumption
+    subgoal
+      unfolding state_trans_t_def
+      apply safe
+      apply (drule (2) P_steps_upd)
+      apply (drule (2) P_steps_upd)
+      apply (drule (3) P_steps_guard)
+      apply (drule (3) P_steps_guard)
+      apply safe
+      apply (rule step_u_s)
+      prefers \<open>_ \<or> _\<close>
+        apply (rule P_steps_reset; force)
+      subgoals \<open>\<forall>pa<length local.N. \<exists>pc. _\<close>
+        apply safe
+        apply (rule P_steps)
+          apply (fastforce simp: p_def)
+         apply (fastforce simp: p_def)
+        by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+      by (auto simp: inv_simp p_def Prod_TA_Defs.N_s_length intro!: P_bounded)
+    done
+qed
+
+lemma state_ta_unfold:
+  "state_ta = (defs.N, defs.P)"
+  by simp
+
+lemma equiv_complete:
+  assumes step: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+    shows "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+    using step proof cases
+    case (step_u_t N P d I)
+    note [simp] = A_simp[OF this(1)]
+    from step_u_t(2-) show ?thesis
+      by (auto simp: state_ta_def p_def state_inv_def intro: step_sn_t)
+  next
+    case (step_u_i P pc_g uu uv uw ux pc_u uy uz va r N I l a l' p)
+    note [simp] = A_simp[OF this(1)]
+    from step_u_i(2-) show ?thesis
+      apply -
+      apply (simp add: Prod_TA_Defs.N_s_length)
+      apply (subst state_ta_unfold)
+      apply (frule steps_P_guard(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply assumption
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply assumption
+        apply (simp; fail)
+      apply (drule N_transD)
+       apply assumption
+      apply (rule step_sn_i)
+                 apply assumption
+                apply (simp add: state_ta_def p_def state_inv_def state_pred_def; fail)+
+        apply (simp add: Prod_TA_Defs.N_s_length; fail)
+       apply (simp add: state_ta_def p_def state_inv_def state_pred_def; fail)+
+      by (auto 4 3 simp: p_def intro: steps_P intro!: states'_updI)
+  next
+    case (step_u_s P pc_g1 vb vc vd ve pc_g2 vf vg vh vi pc_u2 vj vk s1 vl r2 pc_u1 r1 vm vn vo vp
+          N I l1 a l1' p' l2 l2' q
+         )
+    note [simp] = A_simp[OF this(1)]
+    from \<open>q < length L\<close> have "q < p" by (simp add: Prod_TA_Defs.N_s_length)
+    from step_u_s(2-) show ?thesis
+      apply -
+      apply (simp add: Prod_TA_Defs.N_s_length)
+      apply (subst state_ta_unfold)
+      apply (frule steps_P_guard(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply assumption
+       apply (simp; fail)
+      apply (frule steps_P_guard(1))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply (rule \<open>q < p\<close>)
+       apply (simp; fail)
+      apply (drule steps_P_reset[simplified])
+        apply assumption
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply assumption
+        apply (simp; fail)
+      apply (drule N_transD)
+       apply assumption
+      apply (drule N_transD)
+       apply assumption
+      apply (rule step_sn_s)
+                         apply assumption
+                        apply assumption
+                apply (all \<open>(auto; fail)?\<close>)
+        apply (simp add: state_ta_def p_def state_inv_def state_pred_def; fail)
+         apply (simp add: Prod_TA_Defs.N_s_length; fail)
+       apply (simp add: Prod_TA_Defs.N_s_length; fail)
+      apply (clarsimp simp: p_def)
+      subgoal premises prems for p
+        using prems(2, 6-)
+        apply -
+        apply (erule allE[where x = p], erule impE, rule prems)
+        by (fastforce simp: p_def intro: steps_P intro!: states'_updI'')
+      done
+   qed
+
+lemma equiv_sound':
+  assumes step: "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+    shows "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> L' \<in> defs.states' s' \<and> (\<forall>q<p. \<exists>pc st s'' rs pcs.
+             exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+             Some ((pc, st, s'', True, rs), pcs))"
+  using step proof cases
+  case (step_sn_t N d I)
+  then show ?thesis
+    apply simp
+    apply (subst A_unfold)
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (rule conjI)
+    apply (rule step_u_t)
+    using inv inv'' bounded by (auto simp: inv_simp p_def)
+next
+  case (step_sn_i l g a c m f l' N p r I)
+  then show ?thesis
+    apply (simp)
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (simp add: Prod_TA_Defs.N_s_length)
+    apply (subst A_unfold)
+    apply (drule trans_state_taD)
+     apply assumption
+    subgoal
+      unfolding state_trans_t_def
+      apply safe
+      apply (drule (2) P_steps_upd)
+      apply (drule (3) P_steps_guard)
+      apply safe
+      apply (rule step_u_i)
+      subgoals \<open>\<forall>p<length local.N. \<exists>pc. _\<close>
+        apply safe
+        apply (rule P_steps)
+          apply (fastforce simp: p_def)
+         apply (fastforce simp: p_def)
+        by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+      apply (solves \<open>auto simp: inv_simp p_def Prod_TA_Defs.N_s_length intro!: P_bounded\<close>)+
+      subgoal
+        by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+      apply simp
+      subgoal premises prems for pc_g pc_u q
+        using prems(9) \<open>q < _\<close> unfolding state_ta_def state_pred_def
+        by (auto 4 3 simp: p_def split: option.splits)
+      done
+    done
+next
+  case (step_sn_s l1 g1 a ci mi f1 l1' N p l2 g2 co mo f2 l2' q r1 r2 I)
+  then show ?thesis
+    apply simp
+    apply (frule prod.A_simp(1))
+    apply (frule prod.A_simp(2))
+    apply (simp add: Prod_TA_Defs.N_s_length)
+    apply (subst A_unfold)
+    apply (drule trans_state_taD)
+     apply assumption
+    apply (drule trans_state_taD)
+     apply assumption
+    subgoal
+      unfolding state_trans_t_def
+      apply safe
+        apply (drule (2) P_steps_upd)
+        apply (drule (2) P_steps_upd)
+        apply (drule (3) P_steps_guard)
+        apply (drule (3) P_steps_guard)
+        apply safe
+        apply (rule step_u_s)
+      prefers disj
+        apply (rule P_steps_reset; force)
+      subgoals \<open>\<forall>pa<length local.N. \<exists>pc. _\<close>
+        apply safe
+        apply (rule P_steps)
+          apply (fastforce simp: p_def)
+         apply (fastforce simp: p_def)
+        by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+      subgoals \<open>_ \<in> defs.states' s'\<close>
+        by (metis Prod_TA_Defs'.states'_simp Prod_TA_Defs'.states_step local.step states)
+      subgoals \<open>\<exists>pc. _\<close>
+        apply simp
+        subgoal premises prems for pc_g pc_ga pc_u pc_ua q'
+          using prems(14) \<open>q' < _\<close> unfolding state_ta_def state_pred_def
+          by (auto 4 3 simp: p_def split: option.splits)
+        done
+      apply (auto simp: inv_simp p_def Prod_TA_Defs.N_s_length intro!: P_bounded)
+      done
+    done
+qed
+
+lemma equiv_complete':
+  assumes step: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>"
+  shows "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle> \<and> L' \<in> defs.states' s'
+      \<and> (\<forall> q < p. (defs.P ! q) (L' ! q) s')"
+    using step proof cases
+    case (step_u_t N P d I)
+    note [simp] = A_simp[OF this(1)]
+    from step_u_t(2-) show ?thesis
+      apply safe
+      subgoal
+        by (auto simp: state_ta_def p_def state_inv_def intro: step_sn_t)
+      by (fastforce simp: p_def intro: steps_P intro!: states'_updI)+
+  next
+    case (step_u_i P pc_g uu uv uw ux pc_u uy uz va r N I l a l' p)
+    note [simp] = A_simp[OF this(1)]
+    from step_u_i(2-) show ?thesis
+      apply -
+      apply (simp add: Prod_TA_Defs.N_s_length)
+      apply (subst state_ta_unfold)
+      apply (frule steps_P_guard(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply assumption
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply assumption
+        apply (simp; fail)
+      apply (drule N_transD)
+       apply assumption
+        apply safe
+      apply (rule step_sn_i)
+        apply assumption
+      apply (solves auto; fail)+
+      apply (simp add: state_ta_def p_def state_inv_def state_pred_def; fail)
+      apply (solves auto; fail)+
+      apply (simp add: Prod_TA_Defs.N_s_length; fail)
+      apply (solves auto; fail)+
+      by (fastforce simp: p_def intro: steps_P intro!: states'_updI)+
+  next
+    case (step_u_s P pc_g1 vb vc vd ve pc_g2 vf vg vh vi pc_u2 vj vk s1 vl r2 pc_u1 r1 vm vn vo vp N
+            I l1 a l1' p' l2 l2' q
+         )
+    note [simp] = A_simp[OF this(1)]
+    from \<open>q < length L\<close> have "q < p" by (simp add: Prod_TA_Defs.N_s_length)
+    from step_u_s(2-) show ?thesis
+      apply -
+      apply (simp add: Prod_TA_Defs.N_s_length)
+      apply (subst state_ta_unfold)
+      apply (frule steps_P_guard(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply assumption
+       apply (simp; fail)
+      apply (frule steps_P_guard(1))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (drule steps_P_guard(2))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply (rule \<open>q < p\<close>)
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply (rule \<open>q < p\<close>)
+       apply (simp; fail)
+      apply (drule steps_P_reset[simplified])
+        apply assumption
+        apply (simp; fail)
+      apply (frule steps_P_upd(1))
+        apply assumption
+        apply (simp; fail)
+      apply (drule steps_P_upd(2))
+        apply assumption
+        apply (simp; fail)
+      apply (drule N_transD)
+       apply assumption
+      apply (drule N_transD)
+       apply assumption
+        apply safe
+      apply (rule step_sn_s)
+                         apply assumption
+                        apply assumption
+                apply (all \<open>(auto; fail)?\<close>)
+        apply (simp add: state_ta_def p_def state_inv_def state_pred_def; fail)
+         apply (simp add: Prod_TA_Defs.N_s_length; fail)
+         apply (simp add: Prod_TA_Defs.N_s_length; fail)
+
+      subgoals \<open>(defs.P ! _) (L[p' := l1', q := l2'] ! _) s'\<close>
+        by (metis Equiv_TA_Defs.p_def states states'_updI'' steps_P)
+      subgoal
+        by simp (metis Equiv_TA_Defs.p_def states states'_updI'' steps_P)
+      apply (fastforce simp: p_def intro: steps_P intro!: states'_updI'')
+      done
+   qed
+
+  lemma equiv_complete'':
+    assumes step: "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>a \<langle>L', s', u'\<rangle>" "p > 0"
+      shows "(\<forall>q<p. \<exists>pc st s'' rs pcs.
+               exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+               Some ((pc, st, s'', True, rs), pcs))" (is ?A)
+            "bounded B s'" (is ?B)
+  proof -
+    from assms equiv_complete' have *: "\<forall>q<p. (defs.P ! q) (L' ! q) s'" by simp
+    then show ?A unfolding state_ta_def state_pred_def by (fastforce split: option.splits)
+    from \<open>p > 0\<close> * have "(defs.P ! 0) (L' ! 0) s'" by auto
+    with \<open>p > 0\<close> show ?B unfolding state_ta_def state_pred_def by (auto split: option.splits)
+  qed
+
+  lemma equiv_steps_sound':
+    assumes step: "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>"
+    shows "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<and> L' \<in> defs.states' s' \<and>
+        (\<forall>q<p. \<exists>pc st s'' rs pcs.
+             exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+             Some ((pc, st, s'', True, rs), pcs)) \<and> bounded B s'"
+    using step states inv
+  proof (induction A \<equiv> state_ta L \<equiv> L s \<equiv> s u \<equiv> u L' s' u' arbitrary: rule: steps_sn.induct)
+    case (refl)
+    with bounded show ?case by blast
+  next
+    case prems: (step L' s' u' a L'' s'' u'')
+    from prems have *:
+      "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" "L' \<in> defs.states' s'"
+      "(\<forall>q<p. \<exists>pc st s'' rs pcs.
+             exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+             Some ((pc, st, s'', True, rs), pcs))"
+      "bounded B s'"
+      by auto
+    interpret interp: Equiv_TA A n L' s'
+      using pred_time_indep upd_time_indep clock_conj * by unfold_locales (auto simp: Len intro!: *)
+    from prems(3) have
+      "A \<turnstile>n \<langle>L', s', u'\<rangle> \<rightarrow>a \<langle>L'', s'', u''\<rangle>" "L'' \<in> defs.states' s''"
+      "\<forall>q<p. \<exists>pc st s''' rs pcs.
+              exec PF n ((I ! q) (L'' ! q), [], s'', True, []) [] =
+              Some ((pc, st, s''', True, rs), pcs)"
+      "bounded B s''"
+      by (force dest!: interp.equiv_sound')+
+    with * interp.states show ?case
+      by - (assumption | rule conjI steps_un.intros)+
+  qed
+
+lemma equiv_steps_complete':
+    "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<and> L' \<in> defs.states' s' \<and>
+        (\<forall>q<p. \<exists>pc st s'' rs pcs.
+             exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+             Some ((pc, st, s'', True, rs), pcs)) \<and> bounded B s'"
+    if "A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" "p > 0"
+    using that states inv proof (induction A \<equiv> A n \<equiv> n L \<equiv> L s \<equiv> s u \<equiv> u _ _ _ rule: steps_un.induct)
+    case refl
+    with bounded show ?case by blast
+  next
+    case prems: (step L' s' u' a L'' s'' u'')
+    from prems have *:
+      "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" "L' \<in> defs.states' s'"
+      "(\<forall>q<p. \<exists>pc st s'' rs pcs.
+             exec PF n ((I ! q) (L' ! q), [], s', True, []) [] =
+             Some ((pc, st, s'', True, rs), pcs))"
+      "bounded B s'"
+      by auto
+    interpret interp: Equiv_TA A n L' s'
+      using pred_time_indep upd_time_indep clock_conj by unfold_locales (auto simp: Len intro!: *)
+    from interp.equiv_complete'[OF prems(3)] interp.equiv_complete''[OF prems(3) \<open>p > 0\<close>] have
+      "state_ta \<turnstile> \<langle>L', s', u'\<rangle> \<rightarrow>a \<langle>L'', s'', u''\<rangle>" "L'' \<in> defs.states' s''"
+      "\<forall>q<p. \<exists>pc st s''' rs pcs.
+              exec PF n ((I ! q) (L'' ! q), [], s'', True, []) [] =
+              Some ((pc, st, s''', True, rs), pcs)"
+      "bounded B s''"
+      by auto
+    with * interp.states show ?case
+      by auto
+  qed
+
+  lemmas equiv_steps_sound = equiv_steps_sound'[THEN conjunct1]
+  lemmas equiv_steps_complete = equiv_steps_complete'[THEN conjunct1]
+
+  lemma equiv_correct:
+    "state_ta \<turnstile> \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle> \<longleftrightarrow> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" if "p > 0"
+    using that equiv_steps_sound equiv_steps_complete by metis
+
+  lemma prod_correct:
+    "defs.prod_ta \<turnstile> \<langle>(L, s), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle> \<longleftrightarrow> A \<turnstile>n \<langle>L, s, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>" if "p > 0"
+    by (metis prod.prod_correct equiv_correct that)
+
+  end (* End context: UPPAAL network + valid start state *)
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks_Impl.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks_Impl.thy
@@ -0,0 +1,471 @@
+theory UPPAAL_State_Networks_Impl
+  imports Munta_Base.Normalized_Zone_Semantics_Impl UPPAAL_State_Networks
+begin
+
+section \<open>Implementation of UPPAAL Style Networks\<close>
+
+no_notation OR (infix "or" 60)
+
+lemma step_resets:
+  "\<forall> c \<in> set r'. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc"
+  if "stepc cmd u (pc, st, s, f, r) = Some (pc', st', s', f', r')"
+    "\<forall> c \<in> set r. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc" "P pc = Some cmd"
+  using that
+  apply -
+  apply (erule stepc.elims)
+  by (auto split: option.splits if_splits elim!: step.elims) metis+
+
+lemma step_resets':
+  "\<forall> c \<in> set r'. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc"
+  if "step instr (pc, st, s, f, r) = Some (pc', st', s', f', r')"
+    "\<forall> c \<in> set r. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc" "P pc = Some (INSTR instr)"
+  using that
+  by (auto split: option.splits if_splits elim!: step.elims) metis+
+
+lemma step_resets'':
+  "\<forall> c \<in> set r'. \<exists> x pc. Some (STOREC c x) = P pc"
+  if "step instr (pc, st, s, f, r) = Some (pc', st', s', f', r')"
+    "\<forall> c \<in> set r. \<exists> x pc. Some (STOREC c x) = P pc" "P pc = Some instr"
+  using that
+  by (auto split: option.splits if_splits elim!: step.elims) metis+
+
+lemma steps_reset:
+  "\<forall> c \<in> set r'. \<exists> x pc. Some (STOREC c x) = P pc"
+  if "steps P n (pc, st, s, f, r) (pc', st', s', f', r')" "\<forall> c \<in> set r. \<exists> x pc. Some (STOREC c x) = P pc"
+  using that
+  by (induction P \<equiv> P n "(pc, st, s, f, r :: nat list)" "(pc', st', s', f', r')" arbitrary: pc st s f r rule: steps.induct)
+     (auto dest!: step_resets''[where P = P])
+
+lemma exec_reset:
+  "\<forall> c \<in> set r'. \<exists> x pc. Some (STOREC c x) = P pc"
+  if "Some ((pc', st', s', f', r'), pcs') = exec P n (pc, st, s, f, []) pcs"
+  using exec_steps[OF that[symmetric]] steps_reset by force
+
+lemma exec_pointers:
+  "\<forall> pc \<in> set pcs'. \<exists> pc instr. Some instr = P pc"
+  if "Some ((pc', st', s', f', r'), pcs') = exec P n (pc, st, s, f, r) pcs"
+     "\<forall> pc \<in> set pcs. \<exists> pc instr. Some instr = P pc"
+  using that
+  apply (induction rule: exec.induct)
+  by (auto split: option.splits if_splits) metis+
+
+lemma exec_pointers':
+  "\<forall> pc \<in> set pcs'. \<exists> pc instr. Some instr = P pc"
+  if "Some ((pc', st', s', f', r'), pcs') = exec P n (pc, st, s, f, r) []"
+  using that exec_pointers by fastforce
+
+context Prod_TA_Defs
+begin
+
+lemma finite_range_I':
+  assumes "\<forall>A\<in>{0..<p}. finite (range (snd (N ! A)))"
+    shows "finite (range (I' s))"
+    using assms unfolding inv_of_def Product_TA_Defs.product_ta_def N_s_def
+    by (auto simp: inv_of_def p_def intro!: Product_TA_Defs.finite_invariant_of_product)
+
+lemma range_prod_invariant:
+  "range prod_invariant = range (I' s)"
+  unfolding prod_invariant_def using I'_simp by auto
+
+lemma finite_rangeI:
+  assumes "\<forall>A\<in>{0..<p}. finite (range (snd (N ! A)))"
+  shows "finite (range prod_invariant)"
+  using assms by (metis finite_range_I' range_prod_invariant)
+
+end
+
+
+context Equiv_TA_Defs
+begin
+
+lemma states'_len_simp[simp]:
+  "length L = p" if "L \<in> defs.states' s"
+  using that
+  using Product_TA_Defs.states_length defs.N_s_def state_ta_def by fastforce
+
+lemma defs_N_p[simp]:
+  "length defs.N = p"
+  unfolding state_ta_def by simp
+
+lemma defs_p[simp]:
+  "defs.p = p"
+  unfolding defs.p_def by simp
+
+lemma P_Storec_iff:
+  "(Some (INSTR (STOREC x xa)) = P pc) \<longleftrightarrow> (Some (STOREC x xa) = PF pc)"
+  unfolding stripfp_def apply (cases "P pc")
+   apply force
+  subgoal for a
+    by (cases a) auto
+  done
+
+(* Unused but is explaining what is going on below *)
+lemma product_trans_i_resets:
+  "collect_clkvt (Product_TA_Defs.product_trans_i (defs.N_s s))
+  \<subseteq> {c. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc}"
+  unfolding collect_clkvt_def
+  unfolding Product_TA_Defs.product_trans_i_def
+  apply clarsimp
+  unfolding defs.N_s_def
+  unfolding trans_of_def
+  unfolding defs.T_s_def
+  unfolding state_ta_def
+  unfolding state_trans_t_def
+  unfolding make_f_def
+  apply (clarsimp split: option.split_asm)
+  by (auto dest: exec_reset simp: P_Storec_iff)
+
+(* Unused but is explaining what is going on below *)
+lemma product_trans_s_resets:
+  "collect_clkvt (Product_TA_Defs.product_trans_s (defs.N_s s))
+  \<subseteq> {c. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc}"
+  unfolding collect_clkvt_def
+  unfolding Product_TA_Defs.product_trans_s_def
+  apply clarsimp
+  unfolding defs.N_s_def
+  unfolding trans_of_def
+  unfolding defs.T_s_def
+  unfolding state_ta_def
+  unfolding state_trans_t_def
+  unfolding make_f_def
+  apply (clarsimp split: option.split_asm)
+  by (auto dest: exec_reset simp: P_Storec_iff)
+
+lemma product_trans_resets:
+  "collect_clkvt (\<Union>s. defs.T' s) \<subseteq> {c. \<exists> x pc. Some (INSTR (STOREC c x)) = P pc}"
+  unfolding trans_of_def
+  unfolding Product_TA_Defs.product_ta_def
+  apply simp
+  unfolding Product_TA_Defs.product_trans_def
+  unfolding collect_clkvt_def
+  apply safe
+  unfolding Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+   apply clarsimp_all
+  unfolding defs.N_s_def
+  unfolding trans_of_def
+  unfolding defs.T_s_def
+  unfolding state_ta_def
+  unfolding state_trans_t_def
+  unfolding make_f_def
+  apply (clarsimp_all split: option.split_asm)
+  by (auto dest: exec_reset simp: P_Storec_iff)
+
+lemma product_trans_guards:
+  "Timed_Automata.collect_clkt (\<Union>s. defs.T' s)
+  \<subseteq> {constraint_pair ac | ac. \<exists> pc. Some (CEXP ac) = P pc}"
+  unfolding trans_of_def
+  unfolding Product_TA_Defs.product_ta_def
+  apply simp
+  unfolding Product_TA_Defs.product_trans_def
+  unfolding Timed_Automata.collect_clkt_def collect_clock_pairs_def
+  apply safe
+  unfolding Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+   apply clarsimp_all
+  unfolding defs.N_s_def
+  unfolding trans_of_def
+  unfolding defs.T_s_def
+  unfolding state_ta_def
+  unfolding state_trans_t_def
+  unfolding make_g_def
+   apply (clarsimp_all split: option.split_asm)
+  subgoal premises prems
+    using prems(1,3-)
+    unfolding set_map_filter
+    by (smt (verit, best) instrc.simps(5,6) is_instr.cases
+        mem_Collect_eq option.exhaust option.inject
+        option.simps(3,4,5))+
+  subgoal premises prems
+    using prems(1,2,4-)
+    apply safe
+    unfolding set_map_filter
+    apply (drule exec_pointers')
+    by (smt (verit, best) instrc.simps(5,6) is_instr.cases
+        mem_Collect_eq option.exhaust option.inject
+        option.simps(3,4,5))+
+  done
+
+end
+
+datatype bexp =
+  not bexp | "and" bexp bexp | or bexp bexp | imply bexp bexp | \<comment> \<open>Boolean connectives\<close>
+  loc nat nat | \<comment> \<open>Is process p in location l?\<close>
+  eq nat int \<comment> \<open>Does var i equal x?\<close> |
+  le nat int |
+  lt nat int |
+  ge nat int |
+  gt nat int
+
+fun check_bexp :: "bexp \<Rightarrow> nat list \<Rightarrow> int list \<Rightarrow> bool" where
+  "check_bexp (not a) L s \<longleftrightarrow> \<not> check_bexp a L s" |
+  "check_bexp (and a b ) L s \<longleftrightarrow> check_bexp a L s \<and> check_bexp b L s" |
+  "check_bexp (or a b ) L s \<longleftrightarrow> check_bexp a L s \<or> check_bexp b L s" |
+  "check_bexp (imply a b ) L s \<longleftrightarrow> (check_bexp a L s \<longrightarrow> check_bexp b L s)" |
+  "check_bexp (loc p l) L _ \<longleftrightarrow> L ! p = l" |
+  "check_bexp (eq i x) _ s \<longleftrightarrow> s ! i = x" |
+  "check_bexp (le i x) _ s \<longleftrightarrow> s ! i \<le> x" |
+  "check_bexp (lt i x) _ s \<longleftrightarrow> s ! i < x" |
+  "check_bexp (ge i x) _ s \<longleftrightarrow> s ! i \<ge> x" |
+  "check_bexp (gt i x) _ s \<longleftrightarrow> s ! i > x"
+
+datatype formula =
+  EX bexp | EG bexp | AX bexp | AG bexp | Leadsto bexp bexp
+
+abbreviation "repeat x n \<equiv> map (\<lambda> _. x) [0..<n]"
+
+abbreviation "conv_prog P pc \<equiv> map_option (map_instrc real_of_int) (P pc)"
+abbreviation "conv_A' \<equiv> \<lambda> (T, I). (T, conv_cc o I)"
+
+fun hd_of_formula :: "formula \<Rightarrow> nat list \<Rightarrow> int list \<Rightarrow> bool" where
+  "hd_of_formula (formula.EX \<phi>) = check_bexp \<phi>" |
+  "hd_of_formula (EG \<phi>) = check_bexp \<phi>" |
+  "hd_of_formula (AX \<phi>) = Not oo check_bexp \<phi>" |
+  "hd_of_formula (AG \<phi>) = Not oo check_bexp \<phi>" |
+  "hd_of_formula (Leadsto \<phi> _) = check_bexp \<phi>"
+
+subsection \<open>Pre-compiled Networks With States and Clocks as Natural Numbers\<close>
+locale UPPAAL_Reachability_Problem_precompiled_defs =
+  fixes p :: nat \<comment> \<open>Number of processes\<close>
+    and m :: nat \<comment> \<open>Number of clocks\<close>
+    (*
+    and k :: "nat list list"
+      -- "Clock ceiling. Maximal constant appearing in automaton for each state"
+    *)
+    and max_steps :: nat \<comment> \<open>Maximal number of execution for steps of programs in the automaton\<close>
+    and inv :: "(nat, int) cconstraint list list" \<comment> \<open>Clock invariants on locations per process\<close>
+    and pred :: "addr list list" \<comment> \<open>State invariants on locations per process\<close>
+    and trans :: "(addr * nat act * addr * nat) list list list"
+      \<comment> \<open>Transitions between states per process\<close>
+    and prog :: "int instrc option list"
+    and formula :: formula \<comment> \<open>Model checking formula\<close>
+    and bounds :: "(int * int) list"
+begin
+  definition "clkp_set' \<equiv>
+    \<Union> (collect_clock_pairs ` set (concat inv))
+    \<union> {constraint_pair ac | ac. Some (CEXP ac) \<in> set prog}"
+  definition clk_set'_def: "clk_set' =
+    (fst ` clkp_set' \<union> {c. \<exists> x. Some (INSTR (STOREC c x)) \<in> set prog})"
+
+  text \<open>Definition of the corresponding network\<close>
+  definition "I i l \<equiv> if l < length (inv ! i) then inv ! i ! l else []"
+  definition "T i \<equiv>
+    {(l, trans ! i ! l ! j) | l j. l < length (trans ! i) \<and> j < length (trans ! i ! l)}"
+  definition "P \<equiv> map (\<lambda> P l. P ! l) pred"
+  definition "PROG pc \<equiv> (if pc < length prog then prog ! pc else None)"
+  definition N :: "(nat, int, nat) unta" where
+    "N \<equiv> (PROG, map (\<lambda> i. (T i, I i)) [0..<p], P, bounds)"
+  definition "init \<equiv> repeat (0::nat) p"
+  definition "F \<equiv> hd_of_formula formula"
+
+  sublocale equiv: Equiv_TA_Defs N max_steps .
+
+  abbreviation "EA \<equiv> equiv.state_ta"
+
+  abbreviation "A \<equiv> equiv.defs.prod_ta"
+
+  lemma equiv_p_eq[simp]:
+    "equiv.p = p"
+    unfolding equiv.p_def N_def Equiv_TA_Defs.p_def by simp
+
+  lemma length_N_s[simp]:
+    "length (equiv.defs.N_s s) = p"
+    unfolding equiv.defs.N_s_def by simp
+
+  lemma length_N[simp]:
+    "length equiv.defs.N = p"
+    by simp
+
+  lemma
+    "equiv.defs.I' s L = concat (map (\<lambda> q. if q < p then I q (L ! q) else []) [0..<length L])"
+    unfolding inv_of_def
+    unfolding Product_TA_Defs.product_ta_def
+    apply simp
+    unfolding Product_TA_Defs.product_invariant_def
+    unfolding equiv.defs.N_s_def inv_of_def
+    apply (rule arg_cong[where f = concat])
+    unfolding Equiv_TA_Defs.state_ta_def
+      apply simp
+    unfolding N_def Equiv_TA_Defs.state_inv_def
+      by simp
+
+end (* End of definitions locale *)
+
+  lemma snd_comp[simp]:
+    "snd o (\<lambda> i. (f i, g i)) = g"
+  by auto
+
+locale UPPAAL_Reachability_Problem_precompiled =
+  UPPAAL_Reachability_Problem_precompiled_defs +
+  assumes process_length: "length inv = p" "length trans = p" "length pred = p"
+    and lengths:
+    "\<forall> i < p. length (pred ! i) = length (trans ! i) \<and> length (inv ! i) = length (trans ! i)"
+    and state_set: "\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (_, _, _, l) \<in> set xs. l < length T"
+  assumes consts_nats: "snd ` clkp_set' \<subseteq> \<nat>"
+  (* This could also be subset for now but is left like this as an input sanity check right now *)
+  assumes clock_set: "clk_set' = {1..m}"
+    and p_gt_0: "p > 0"
+    and m_gt_0: "m > 0"
+    and processes_have_trans: "\<forall> i < p. trans ! i \<noteq> []" \<comment> \<open>Necessary for refinement\<close>
+    and start_has_trans: "\<forall> q < p. trans ! q ! 0 \<noteq> []" \<comment> \<open>Necessary for refinement\<close>
+  (* Do not need this but a useful cautious check for the user? *)
+  assumes resets_zero: "\<forall> x c. Some (INSTR (STOREC c x)) \<in> set prog \<longrightarrow> x = 0"
+
+(*
+locale UPPAAL_Reachability_Problem_precompiled =
+  UPPAAL_Reachability_Problem_precompiled_raw +
+  assumes discrete_state_finite: "\<forall> i < p. \<forall> l < length (trans ! i). finite {s. (pred ! i ! l) s}"
+*)
+begin
+
+  lemma consts_nats':
+    "\<forall> I \<in> set inv. \<forall> cc \<in> set I. \<forall> (c, d) \<in> collect_clock_pairs cc. d \<in> \<nat>"
+    (* "\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (g, _) \<in> set xs. \<forall> (c, d) \<in> collect_clock_pairs g. d \<in> \<nat>" *)
+    "\<forall> ac. Some (CEXP ac) \<in> set prog \<longrightarrow> (snd (constraint_pair ac) \<in> \<nat>)"
+    using consts_nats unfolding clkp_set'_def by force+
+
+  lemma clk_pairs_N_inv:
+    "\<Union> (collect_clock_pairs ` range (snd x)) \<subseteq> \<Union> (collect_clock_pairs ` set (concat inv))"
+    if "x \<in> set equiv.defs.N" for x
+    using that process_length(1)
+    unfolding equiv.state_ta_def equiv.state_inv_def equiv.p_def
+    unfolding N_def I_def
+    by clarsimp (auto split: if_split_asm dest: nth_mem)+
+
+  lemma clkp_set_simp_1:
+    "\<Union> (collect_clock_pairs ` set (concat inv)) \<supseteq> Timed_Automata.collect_clki (snd A)"
+    unfolding equiv.defs.prod_ta_def inv_of_def
+    apply (rule subset_trans)
+     apply simp
+     apply (rule equiv.defs.collect_clki_prod_invariant')
+    unfolding Timed_Automata.collect_clki_def using clk_pairs_N_inv nth_mem by blast
+
+  lemma clk_set_simp_2:
+    "{c. \<exists> x. Some (INSTR (STOREC c x)) \<in> set prog} \<supseteq> collect_clkvt (trans_of A)"
+    unfolding equiv.defs.prod_ta_def trans_of_def
+    apply (rule subset_trans)
+     apply simp
+     apply (rule equiv.defs.collect_clkvt_prod_trans_subs)
+    apply (rule subset_trans)
+     apply (rule equiv.product_trans_resets)
+    unfolding N_def PROG_def by (auto dest!: nth_mem) metis
+
+  lemma clkp_set_simp_3:
+    "{constraint_pair ac | ac. Some (CEXP ac) \<in> set prog} \<supseteq> Timed_Automata.collect_clkt (trans_of A)"
+    unfolding equiv.defs.prod_ta_def trans_of_def
+    apply (rule subset_trans)
+     apply simp
+     apply (rule equiv.defs.collect_clkt_prod_trans_subs)
+    apply (rule subset_trans)
+     apply (rule equiv.product_trans_guards)
+    unfolding N_def PROG_def by (auto dest!: nth_mem)
+
+  lemma clkp_set'_subs:
+    "Timed_Automata.clkp_set A \<subseteq> clkp_set'"
+    using clkp_set_simp_1 clkp_set_simp_3
+    by (auto simp add: clkp_set'_def Timed_Automata.clkp_set_def inv_of_def)
+
+  lemma clk_set'_subs:
+    "clk_set A \<subseteq> clk_set'"
+    using clkp_set'_subs clk_set_simp_2 by (auto simp: clk_set'_def)
+
+  lemma clk_set:
+    "clk_set A \<subseteq> {1..m}"
+    using clock_set m_gt_0 clk_set'_subs by auto
+
+  lemma
+    "\<forall>(_, d)\<in>Timed_Automata.clkp_set A. d \<in> \<int>"
+    unfolding Ints_def by auto
+
+  lemma clkp_set'_consts_nat:
+    "\<forall>(_, d)\<in>clkp_set'. d \<in> \<nat>"
+    using consts_nats' unfolding clkp_set'_def
+    apply safe
+     apply force
+    by (metis snd_conv)
+
+  lemma clkp_set_consts_nat:
+    "\<forall>(_, d)\<in>Timed_Automata.clkp_set A. d \<in> \<nat>"
+    using clkp_set'_subs clkp_set'_consts_nat by auto
+
+  lemma finite_clkp_set':
+    "finite clkp_set'"
+    unfolding clkp_set'_def
+    using [[simproc add: finite_Collect]]
+    by (auto simp: inj_on_def intro!: finite_vimageI)
+
+  lemma finite_clkp_set_A[intro, simp]:
+    "finite (Timed_Automata.clkp_set A)"
+    using clkp_set'_subs finite_clkp_set' by (rule finite_subset)
+
+  lemma clkp_set'_bounds:
+    "a \<in> {Suc 0..m}" if "(a, b) \<in> clkp_set'"
+    using that clock_set unfolding clk_set'_def by auto
+
+  lemma finite_range_inv_of_A[intro, simp]:
+    "finite (range (inv_of A))"
+    unfolding inv_of_def equiv.defs.prod_ta_def
+    apply simp
+    apply (rule equiv.defs.finite_rangeI)
+      apply simp
+    unfolding Equiv_TA_Defs.state_ta_def
+    apply simp
+    unfolding equiv.state_inv_def
+    unfolding N_def
+    unfolding I_def
+    by (auto intro: finite_subset[where B = "{[]}"])
+
+  lemma Collect_fold_pair:
+    "{f a b | a b. P a b} = (\<lambda> (a, b). f a b) ` {(a, b). P a b}" for P
+    by auto
+
+  lemma finite_T[intro, simp]:
+    "finite (trans_of A)"
+    unfolding trans_of_def equiv.defs.prod_ta_def fst_conv
+  proof (rule equiv.defs.finite_prod_trans, goal_cases)
+    case (1 s)
+    show "\<forall>l q. q < equiv.defs.p \<longrightarrow> (equiv.defs.P ! q) l s \<longrightarrow> (bounded equiv.B) s"
+      apply simp
+      unfolding equiv.state_ta_def equiv.state_pred_def
+      by (simp split: option.split)
+  next
+    case 2
+    show "finite {s. bounded equiv.B s}" by (rule equiv.bounded_finite)
+  next
+    case 3
+    show ?case
+    proof
+      fix A assume A: "A \<in> set equiv.defs.N"
+      have
+        "{(l, j). l < length (trans ! i) \<and> j < length (trans ! i ! l)}
+        = \<Union> ((\<lambda> l. {(l, j) | j. j < length (trans ! i ! l)}) ` {l. l < length (trans ! i)})" for i
+        by auto
+      then have "finite (T q)" if "q < p" for q
+        using that unfolding T_def by (fastforce simp: Collect_fold_pair)
+      then have "finite (fst (equiv.N ! q))" if "q < p" for q
+        using that unfolding N_def by simp
+      then have "finite (equiv.state_trans q)" if "q < p" for q
+        using that
+        unfolding Equiv_TA_Defs.state_trans_t_def
+        using [[simproc add: finite_Collect]]
+          by auto
+      then show "finite (fst A)" using A unfolding Equiv_TA_Defs.state_ta_def by auto
+    qed
+  qed (auto simp: p_gt_0)
+
+  sublocale TA_Start_No_Ceiling A "(init, s0)" m
+    using clkp_set_consts_nat clk_set m_gt_0 by - (standard; blast)
+
+  lemma P_p[simp]:
+    "length P = p"
+    unfolding P_def using process_length(3) by simp
+
+  lemma length_I[simp]:
+    "length equiv.I = p"
+    unfolding N_def by simp
+
+  lemma length_N[simp]:
+    "length equiv.N = p"
+    unfolding N_def by simp
+
+end (* End of locale *)
+
+end (* End of theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks_Impl_Refine.thy
--- /dev/null
+++ thys/Munta_Model_Checker/Uppaal_Networks/UPPAAL_State_Networks_Impl_Refine.thy
@@ -0,0 +1,3009 @@
+theory UPPAAL_State_Networks_Impl_Refine
+  imports
+    Program_Analysis
+    Munta_Base.Normalized_Zone_Semantics_Impl_Refine
+    Munta_Base.TA_Impl_Misc
+    Munta_Base.TA_Syntax_Bundles
+begin
+
+unbundle no_library_syntax
+
+section \<open>Imperative Implementation of UPPAAL Style Networks\<close>
+
+subsection \<open>Executable Successor Computation\<close>
+
+lemma exec_state_length:
+  assumes "exec prog n (pc, st, s, f, rs) pcs = Some ((pc', st', s', f', rs'), pcs')"
+  shows "length s = length s'"
+  using assms
+proof (induction prog n "(pc, st, s, f, rs)" pcs arbitrary: pc st s f rs pcs rule: exec.induct)
+  case 1
+  then show ?case by simp
+next
+  case prems: (2 prog n pc st m f rs pcs)
+  from prems(2) show ?case
+    apply (clarsimp split: option.split_asm if_split_asm)
+    apply (drule prems(1), assumption+)
+    by (erule step.elims; simp split: if_split_asm)
+qed
+
+locale UPPAAL_Reachability_Problem_precompiled_defs' =
+  UPPAAL_Reachability_Problem_precompiled_defs +
+  fixes na :: nat
+begin
+
+text \<open>Definition of implementation auxiliaries (later connected to the automaton via proof)\<close>
+
+  definition
+    "trans_i_map =
+      map (map (List.map_filter
+        (\<lambda> (g, a, m, l'). case a of Sil a \<Rightarrow> Some (g, a, m, l') | _ \<Rightarrow> None)
+      )) trans"
+
+  definition
+    "trans_in_map \<equiv>
+      map (map (map
+        (\<lambda> (g, a, m, l'). case a of In a \<Rightarrow> (g, a, m, l')) o filter (\<lambda> (g, a, m, l').
+          case a of In a \<Rightarrow> True | _ \<Rightarrow> False))
+          ) trans"
+
+  definition
+    "trans_out_map \<equiv>
+      map (map (map
+        (\<lambda> (g, a, m, l'). case a of Out a \<Rightarrow> (g, a, m, l')) o filter (\<lambda> (g, a, m, l').
+          case a of Out a \<Rightarrow> True | _ \<Rightarrow> False))
+          ) trans"
+
+  abbreviation
+    "nested_list_to_iarray xs \<equiv> IArray (map IArray xs)"
+
+  (* Could be optimized by using a better data structure? *)
+  definition
+    "actions_by_state i \<equiv> fold (\<lambda> t acc. acc[fst (snd t) := (i, t) # (acc ! fst (snd t))])"
+
+  definition
+    "all_actions_by_state t L \<equiv>
+      fold (\<lambda> i. actions_by_state i (t !! i !! (L ! i))) [0..<p] (repeat [] na)"
+
+  definition "PROG' pc \<equiv> (if pc < length prog then (IArray prog) !! pc else None)"
+  abbreviation "PF \<equiv> stripfp PROG'"
+  abbreviation "PT \<equiv> striptp PROG'"
+  definition "runf pc s \<equiv> exec PF max_steps (pc, [], s, True, []) []"
+  definition "runt pc s \<equiv> exec PT max_steps (pc, [], s, True, []) []"
+
+  lemma PROG'_PROG [simp]:
+    "PROG' = PROG"
+    unfolding PROG'_def PROG_def by (rule ext) simp
+
+  definition "bounded' s \<equiv>
+    (\<forall>i<length s. fst (IArray bounds !! i) < s ! i \<and> s ! i < snd (IArray bounds !! i))"
+
+  definition
+    "check_pred L s \<equiv>
+      list_all
+        (\<lambda> q.
+          case runf (pred ! q ! (L ! q)) s of
+            Some ((_, _, _, f, _), _) \<Rightarrow> f \<and> bounded' s
+          | None \<Rightarrow> False
+        )
+        [0..<p]
+      "
+
+  definition
+    "make_cconstr pcs = List.map_filter
+      (\<lambda> pc.
+        case PROG pc of
+          Some (CEXP ac) \<Rightarrow> Some ac
+        | _ \<Rightarrow> None
+      )
+      pcs"
+
+  definition
+    "check_g pc s \<equiv>
+      case runt pc s of
+        Some ((_, _, _, True, _), pcs) \<Rightarrow> Some (make_cconstr pcs)
+      | _ \<Rightarrow> None
+      "
+
+   definition
+    "trans_i_from \<equiv> \<lambda> (L, s) i.
+      List.map_filter (\<lambda> (g, a, m, l').
+        case check_g g s of
+          Some cc \<Rightarrow>
+          (case runf m s of
+            Some ((_, _, s', _, r), _) \<Rightarrow>
+              if check_pred (L[i := l']) s'
+              then Some (cc, a, r, (L[i := l'], s'))
+              else None
+         | _ \<Rightarrow> None)
+      | _ \<Rightarrow> None)
+        ((IArray (map IArray trans_i_map)) !! i !! (L ! i))"
+
+  definition
+    "trans_i_fun L \<equiv> concat (map (trans_i_from L) [0..<p])"
+
+  definition
+    "make_reset m1 s \<equiv>
+      case runf m1 s of
+        Some ((_, _, _, _, r1), _) \<Rightarrow> r1
+      | None \<Rightarrow> []
+    "
+
+  definition
+    "pairs_by_action \<equiv> \<lambda> (L, s) OUT. concat o
+      map (\<lambda> (i, g1, a, m1, l1). List.map_filter
+      (\<lambda> (j, g2, a, m2, l2).
+        if i = j then None else
+        case (check_g g1 s, check_g g2 s) of
+          (Some cc1, Some cc2) \<Rightarrow> (
+          case runf m2 s of
+            Some ((_, _, s1, _, r2), _) \<Rightarrow> (
+            case runf m1 s1 of
+              Some (( _, _, s', _, _), _) \<Rightarrow>
+                if check_pred (L[i := l1, j := l2]) s'
+                then Some (cc1 @ cc2, a, make_reset m1 s @ r2, (L[i := l1, j := l2], s'))
+                else None
+            | _ \<Rightarrow> None)
+          | _ \<Rightarrow> None)
+        | _ \<Rightarrow> None
+      )
+      OUT)
+        "
+
+  definition
+    "trans_s_fun \<equiv> \<lambda> (L, s).
+      let
+        In = all_actions_by_state (nested_list_to_iarray trans_in_map) L;
+        Out = all_actions_by_state (nested_list_to_iarray trans_out_map) L
+      in
+        concat (map (\<lambda> a. pairs_by_action (L, s) (Out ! a) (In ! a)) [0..<na])
+    "
+
+  definition
+    "trans_fun L \<equiv> trans_s_fun L @ trans_i_fun L"
+
+  lemma trans_i_fun_trans_fun:
+    assumes "(g, a, r, L') \<in> set (trans_i_fun L)"
+    shows "(g, a, r, L') \<in> set (trans_fun L)"
+    using assms unfolding trans_fun_def by auto
+
+  lemma trans_s_fun_trans_fun:
+    assumes "(g, a, r, L') \<in> set (trans_s_fun L)"
+    shows "(g, a, r, L') \<in> set (trans_fun L)"
+    using assms unfolding trans_fun_def by auto
+
+end (* End of locale for implementation definitions *)
+
+context Prod_TA_Defs
+begin
+
+  (* Can overwrite this theorem from other context in SAME locale *)
+  lemma prod_trans_i_alt_def:
+    "prod_trans_i =
+      {((L, s), g, a, r, (L', s')) | L s g c a r m L' s'.
+       (L, g, (a, Networks.label.Act (c, m)), r, L') \<in> Product_TA_Defs.product_trans_i (N_s s) \<and>
+       (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+       \<and> c s \<and> Some s' = m s}"
+    unfolding
+      prod_trans_i_def trans_of_def Product_TA_Defs.product_ta_def
+      Product_TA_Defs.product_trans_def
+      Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+    by (safe; simp; metis)
+
+  lemma prod_trans_s_alt_def:
+    "prod_trans_s =
+      {((L, s), g, a, r, (L', s')) | L s g ci co a r mi mo L' s1 s'.
+        ci s \<and> co s
+        \<and> (\<forall> q < p. (P ! q) (L ! q) s) \<and> (\<forall> q < p. (P ! q) (L' ! q) s')
+        \<and> (L, g, (a, Networks.label.Syn (ci, mi) (co, mo)), r, L')
+          \<in> Product_TA_Defs.product_trans_s (N_s s)
+        \<and> Some s' = mi s1 \<and> Some s1 = mo s
+      }"
+    unfolding
+      prod_trans_s_def trans_of_def Product_TA_Defs.product_ta_def
+      Product_TA_Defs.product_trans_def
+      Product_TA_Defs.product_trans_i_def
+      Product_TA_Defs.product_trans_s_def
+    by (safe; simp; metis)
+
+end
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  lemma T_s_unfold_1:
+    "fst ` equiv.defs.T_s q s = fst ` fst (equiv.N ! q)" if "q < p"
+    using \<open>q < p\<close>
+    unfolding equiv.defs.T_s_def
+    unfolding equiv.state_ta_def
+    unfolding equiv.state_trans_t_def
+    by force
+
+  lemma T_s_unfold_2:
+    "(snd o snd o snd o snd) ` equiv.defs.T_s q s = (snd o snd o snd o snd) ` fst (equiv.N ! q)"
+    if "q < p"
+    using \<open>q < p\<close>
+    unfolding equiv.defs.T_s_def
+    unfolding equiv.state_ta_def
+    unfolding equiv.state_trans_t_def
+    by force
+
+end
+
+lemma (in Equiv_TA_Defs) p_p:
+  "defs.p = p"
+  by simp
+
+context
+  Prod_TA_Defs
+begin
+
+  lemma states'_alt_def:
+    "states' s =
+    {L. length L = p \<and>
+        (\<forall> q < p. (L ! q) \<in> fst ` (fst (fst A ! q)) \<union> (snd o snd o snd o snd) ` (fst (fst A ! q)))}"
+    unfolding trans_of_def Product_TA_Defs.product_ta_def N_s_def
+    unfolding Product_TA_Defs.product_trans_def
+    unfolding Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+    apply simp
+    apply safe
+    unfolding T_s_def
+      apply (fastforce simp: Product_TA_Defs.states_def trans_of_def p_def)
+     apply (force simp: Product_TA_Defs.states_def trans_of_def p_def)
+    by (fastforce simp: Product_TA_Defs.states_def trans_of_def p_def image_iff)
+
+end
+
+
+context UPPAAL_Reachability_Problem_precompiled
+begin
+
+lemma PF_unfold:
+  "equiv.PF = stripfp (conv_prog PROG)"
+  using [[show_abbrevs=false]]
+  unfolding N_def stripfp_def
+  apply (rule ext)
+  subgoal for x
+    apply (cases "PROG x")
+     apply (solves simp)
+    subgoal for a
+      by (cases a) auto
+    done
+  done
+
+lemma PT_unfold:
+  "equiv.PT = striptp (conv_prog PROG)"
+  using [[show_abbrevs=false]]
+  unfolding N_def striptp_def
+  apply (rule ext)
+  subgoal for x
+    apply (cases "PROG x")
+     apply (solves simp)
+    subgoal for a
+      by (cases a) auto
+    done
+  done
+
+lemma states'I:
+  "l \<in> equiv.defs.states' s" if "A \<turnstile> (l, s) \<longrightarrow>g,a,r (l', s')"
+  using equiv.defs.prod_ta_cases[OF that]
+  unfolding equiv.defs.prod_trans_i_alt_def equiv.defs.prod_trans_s_alt_def
+  unfolding Product_TA_Defs.product_trans_def
+  unfolding Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+  by fastforce
+
+lemma A_lengthD:
+  "length l = p" if "A \<turnstile> (l, s) \<longrightarrow>g,a,r (l', s')"
+  using that by (auto dest: states'I)
+
+lemma N_s_state_trans:
+  assumes "equiv.defs.N_s s ! q \<turnstile> l ! q \<longrightarrow>g,(a, c, m'),r l'" "q < p"
+  obtains f' g' where
+    "(l ! q, g', (a, c, m'), f', l') \<in> equiv.state_trans q" "g = g' s" "r = f' s"
+  using assms
+  unfolding equiv.defs.N_s_def trans_of_def equiv.defs.T_s_def
+  unfolding equiv.state_ta_def by auto
+
+lemma make_f_collect_store:
+  assumes "(l, pc_g, a, pc_u, l') \<in> fst (equiv.N ! q)" "c \<in> set (equiv.make_f pc_u s)" "q < p"
+  shows "c \<in> fst ` collect_store' pc_u"
+proof -
+  from assms(1) \<open>q < p\<close> have "(pc_g, a, pc_u, l') \<in> set (trans ! q ! l)"
+    unfolding N_def T_def by (auto dest!: nth_mem)
+  from assms obtain pc x2 x3 x4 pcs r where exec:
+    "exec equiv.PF max_steps (pc_u, [], s, True, []) [] = Some ((pc, x2, x3, x4, r), pcs)"
+    unfolding equiv.make_f_def by (auto split: option.split_asm) metis
+  with assms have "c \<in> set r" unfolding equiv.make_f_def by auto
+  with exec_reset'[OF exec] obtain pc' d where "Some (STOREC c d) = equiv.PF pc'" "pc' \<in> set pcs"
+    by force
+  with exec obtain y2 y3 y4 y5 where steps:
+    "steps equiv.PF max_steps (pc_u, [], s, True, []) (pc', y2, y3, y4, y5)"
+    by (auto intro: exec_steps')
+  from \<open>_ = equiv.PF pc'\<close> have "pc' < length prog"
+    unfolding N_def PROG_def stripfp_def by (simp split: if_split_asm)
+  from steps have "pc' \<in> steps_approx max_steps prog pc_u"
+    unfolding PF_unfold
+    unfolding stripfp_def
+    by (auto simp: PROG_def intro: steps_steps_approx[of stripf, OF _ _ _ \<open>pc' < length prog\<close>])
+  with \<open>_ = equiv.PF pc'\<close> \<open>_ \<in> set (trans ! q ! l)\<close> show ?thesis
+    unfolding collect_store'_def stripfp_def N_def PROG_def
+    apply (clarsimp split: if_split_asm)
+    apply (cases "prog ! pc'")
+     apply (simp; fail)
+    subgoal for x
+      by (cases x; force)
+    done
+qed
+
+lemma resets_approx:
+  "set r \<subseteq>
+  \<Union> {fst ` collect_store' r | i g a r. (g, a, r, (l' ! i)) \<in> set (trans ! i ! (l ! i))}"
+  if "A \<turnstile> (l, s) \<longrightarrow>g,a,r (l', s')"
+proof -
+  from that have [simp]: "length l = p" by (auto dest: A_lengthD)
+  show ?thesis using that
+    apply clarsimp
+    apply (drule equiv.defs.prod_ta_cases)
+    apply safe
+    subgoal for x
+      unfolding equiv.defs.prod_trans_i_alt_def
+      apply simp
+      unfolding Product_TA_Defs.product_trans_def
+      apply safe
+      unfolding Product_TA_Defs.product_trans_i_def
+      apply clarsimp
+      apply (erule N_s_state_trans, assumption)
+      unfolding equiv.state_trans_t_def
+      apply clarsimp
+      subgoal for q l'' pc_g pc_u
+        apply (frule make_f_collect_store, assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for b j
+        proof -
+          from prems(6,8,11-) have
+            "(pc_g, Sil a, pc_u, l[q := l''] ! q) \<in> set (trans ! q ! (l ! q))"
+            by simp
+          with prems(6,8,11-) show ?thesis by blast
+        qed
+        done
+      done
+    subgoal for x
+      unfolding equiv.defs.prod_trans_s_alt_def
+      apply simp
+      unfolding Product_TA_Defs.product_trans_def
+      apply safe
+      unfolding Product_TA_Defs.product_trans_s_def
+      apply clarsimp
+      apply (erule N_s_state_trans, assumption)
+      apply (erule N_s_state_trans, assumption)
+      unfolding equiv.state_trans_t_def
+      apply clarsimp
+      apply (erule disjE)
+      subgoal for s1 p' q l'' l'aa pc_g pc_ga pc_u pc_ua
+        apply (frule make_f_collect_store, assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for b j j'
+        proof -
+          from prems(9-) have
+            "(pc_g, In a, pc_u, l[p' := l'', q := l'aa] ! p') \<in> set (trans ! p' ! (l ! p'))"
+            by simp
+          with prems(9-) show ?thesis by blast
+        qed
+        done
+      subgoal for s1 q p' l'aa l'' pc_ga pc_g pc_ua pc_u
+        apply (frule make_f_collect_store, assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for b j j'
+        proof -
+          from prems(9-) have
+            "(pc_g, Out a, pc_u, l[q := l'aa, p' := l''] ! p') \<in> set (trans ! p' ! (l ! p'))"
+            by simp
+          with prems(9-) show ?thesis by blast
+        qed
+        done
+      done
+    done
+qed
+
+lemma make_g_clkp_set'':
+  fixes x
+  assumes
+    "(l, pc_g, a, pc_u, l') \<in> fst (equiv.N ! q)" "x \<in> collect_clock_pairs (equiv.make_g pc_g s)"
+    "q < p"
+  shows "x \<in> clkp_set'' q l"
+proof -
+  from assms(1) \<open>q < p\<close> have "(pc_g, a, pc_u, l') \<in> set (trans ! q ! l)"
+    unfolding N_def T_def by (auto dest!: nth_mem)
+  from assms obtain pc x2 x3 x4 r pcs where exec:
+    "exec equiv.PT max_steps (pc_g, [], s, True, []) [] = Some ((pc, x2, x3, x4, r), pcs)"
+    unfolding equiv.make_g_def by (auto split: option.split_asm)
+  with assms have "x \<in> collect_clock_pairs (List.map_filter (\<lambda> pc.
+        case equiv.P pc of
+          Some (CEXP ac) \<Rightarrow> Some ac
+        | _ \<Rightarrow> None
+          )
+        pcs)"
+    unfolding equiv.make_g_def by auto
+  then obtain pc' ac where
+    "equiv.P pc' = Some (CEXP ac)" "x = constraint_pair ac" "pc' \<in> set pcs"
+    unfolding equiv.make_g_def collect_clock_pairs_def set_map_filter
+    by (clarsimp split: option.split_asm; auto split: instrc.split_asm)
+  with exec obtain y2 y3 y4 y5 where steps:
+    "steps equiv.PT max_steps (pc_g, [], s, True, []) (pc', y2, y3, y4, y5)"
+    by (auto intro: exec_steps')
+  from \<open>equiv.P pc' = _\<close> have "pc' < length prog"
+    unfolding N_def PROG_def by (simp split: if_split_asm)
+  from steps have "pc' \<in> steps_approx max_steps prog pc_g"
+    unfolding PT_unfold
+    unfolding striptp_def
+    by (auto simp: PROG_def intro: steps_steps_approx[of stript, OF _ _ _ \<open>pc' < length prog\<close>])
+  with \<open>equiv.P pc' = _\<close> \<open>_ \<in> set (trans ! q ! l)\<close> \<open>x = _\<close> show ?thesis
+    unfolding clkp_set''_def collect_cexp'_def N_def PROG_def by (force split: if_split_asm)
+qed
+
+lemma guard_approx:
+  "collect_clock_pairs g \<subseteq>
+  \<Union> {clkp_set'' i (l ! i) | i g a r.
+      (g, a, r, (l' ! i)) \<in> set (trans ! i ! (l ! i)) \<and> l \<in> equiv.defs.states' s \<and> i < p
+    }"
+  if "A \<turnstile> (l, s) \<longrightarrow>g,a,r (l', s')"
+proof -
+  from that have [simp]: "length l = p" by (auto dest: A_lengthD)
+  show ?thesis using that
+    apply clarsimp
+    apply (drule equiv.defs.prod_ta_cases)
+    apply safe
+    subgoal for x b
+      unfolding equiv.defs.prod_trans_i_alt_def
+      apply simp
+      unfolding Product_TA_Defs.product_trans_def
+      apply safe
+      unfolding Product_TA_Defs.product_trans_i_def
+      apply clarsimp
+      apply (erule N_s_state_trans, assumption)
+      unfolding equiv.state_trans_t_def
+      apply clarsimp
+      subgoal for q l'' pc_g pc_u
+        apply (frule make_g_clkp_set'', assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for j
+        proof -
+          from prems(6,8,11-) have
+            "(pc_g, Sil a, pc_u, l[q := l''] ! q) \<in> set (trans ! q ! (l ! q))"
+            by simp
+          with prems(6,8,11-) prems show ?thesis by blast
+        qed
+        done
+      done
+    subgoal for x b
+      unfolding equiv.defs.prod_trans_s_alt_def
+      apply simp
+      unfolding Product_TA_Defs.product_trans_def
+      apply safe
+      unfolding Product_TA_Defs.product_trans_s_def
+      apply clarsimp
+      apply (erule N_s_state_trans, assumption)
+      apply (erule N_s_state_trans, assumption)
+      unfolding equiv.state_trans_t_def
+      apply clarsimp
+      apply (drule collect_clock_pairs_append_cases)
+      apply (erule disjE)
+      subgoal for s1 p' q l'' l'aa pc_g pc_ga pc_u pc_ua
+        unfolding equiv.make_c_def
+          apply (clarsimp split: option.split_asm)
+        apply (frule make_g_clkp_set'', assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for j j'
+        proof -
+          from prems(9-) have
+            "(pc_g, In a, pc_u, l[p' := l'', q := l'aa] ! p') \<in> set (trans ! p' ! (l ! p'))"
+            by simp
+          with prems(9-) show ?thesis by blast
+        qed
+        done
+      subgoal for s1 q p' l'aa l'' pc_ga pc_g pc_ua pc_u
+        apply (frule make_g_clkp_set'', assumption+)
+        unfolding N_def T_def
+        apply (clarsimp dest!: nth_mem)
+        subgoal premises prems for j j'
+        proof -
+          from prems(9-) have
+            "(pc_g, Out a, pc_u, l[q := l'aa, p' := l''] ! p') \<in> set (trans ! p' ! (l ! p'))"
+            by simp
+          with prems(9-) show ?thesis by blast
+        qed
+        done
+      done
+    done
+qed
+
+end (* End of context for pre-compiled reachability problem *)
+
+
+abbreviation "conv B \<equiv> (conv_prog (fst B), (map conv_A' (fst (snd B))), snd (snd B))"
+
+context UPPAAL_Reachability_Problem_precompiled
+begin
+
+  sublocale defs':
+    Equiv_TA_Defs "conv N" max_steps .
+
+  lemma equiv_states'_alt_def:
+    "equiv.defs.states' s =
+      {L. length L = p \<and>
+        (\<forall> q < p. L ! q \<in> fst ` fst (equiv.N ! q)
+                \<or> L ! q \<in> (snd o snd o snd o snd) ` fst (equiv.N ! q))}"
+    unfolding Product_TA_Defs.states_def
+    unfolding equiv.defs.N_s_def trans_of_def
+    using T_s_unfold_1 T_s_unfold_2 by simp
+
+  lemma init_states:
+    "init \<in> equiv.defs.states' s0"
+    using processes_have_trans start_has_trans
+    unfolding equiv_states'_alt_def
+    unfolding init_def N_def T_def by force
+
+  lemma p_p[simp]:
+    "defs'.p = p"
+    unfolding defs'.p_def by simp
+
+  lemma T_s_unfold_1':
+    "fst ` defs'.defs.T_s q s = fst ` fst (defs'.N ! q)" if "q < p"
+    using \<open>q < p\<close>
+    unfolding defs'.defs.T_s_def
+    unfolding defs'.state_ta_def
+    unfolding defs'.state_trans_t_def p_p
+    by force
+
+  lemma T_s_unfold_2':
+    "(snd o snd o snd o snd) ` defs'.defs.T_s q s = (snd o snd o snd o snd) ` fst (defs'.N ! q)"
+    if "q < p"
+    using \<open>q < p\<close>
+    unfolding defs'.defs.T_s_def
+    unfolding defs'.state_ta_def
+    unfolding defs'.state_trans_t_def p_p
+    by force
+
+  lemma product_states'_alt_def:
+    "defs'.defs.states' s =
+      {L. length L = p \<and>
+        (\<forall> q < p. L ! q \<in> fst ` fst (defs'.N ! q)
+                \<or> L ! q \<in> (snd o snd o snd o snd) ` fst (defs'.N ! q))}"
+    unfolding Product_TA_Defs.states_def
+    unfolding defs'.defs.N_s_def trans_of_def
+    using T_s_unfold_1' T_s_unfold_2'
+    by force
+
+  lemma states'_conv[simp]:
+    "defs'.defs.states' s = equiv.defs.states' s"
+    unfolding product_states'_alt_def equiv_states'_alt_def
+    unfolding N_def T_def by simp
+
+  lemma init_states_defs'[intro]:
+    "init \<in> defs'.defs.states' s0"
+    using init_states by simp
+
+  lemma
+    "defs'.I = equiv.I"
+    by simp
+
+  lemma PF_PF[simp]:
+    "defs'.PF = equiv.PF"
+    apply simp
+    unfolding stripfp_def
+    apply (rule ext)
+    apply clarsimp
+    subgoal for x
+      apply (cases "equiv.P x")
+       apply simp
+      subgoal for a
+        by (cases a) auto
+      done
+    done
+
+  lemma PF_PROG[simp]:
+    "equiv.PF = stripfp PROG"
+    unfolding N_def by simp
+
+  lemma I_simp[simp]:
+    "(equiv.I ! q) l = pred ! q ! l" if "q < p"
+    unfolding N_def P_def using \<open>q < p\<close> process_length(3) by simp
+
+  lemma
+    "defs'.P = conv_prog PROG"
+    by (simp add: N_def)
+
+  lemma states_len[intro]:
+    assumes
+      "q < p" "L \<in> equiv.defs.states' s"
+    shows
+      "L ! q < length (trans ! q)"
+    using assms unfolding Product_TA_Defs.states_def
+    apply simp
+    unfolding trans_of_def equiv.defs.N_s_def
+    apply (simp add: T_s_unfold_1[simplified] T_s_unfold_2[simplified])
+    unfolding N_def
+    apply simp
+    unfolding T_def
+      using state_set
+    unfolding process_length(2)[symmetric]
+    apply safe
+    apply (erule allE)
+    apply (erule impE)
+     apply assumption
+    apply safe
+    by (clarsimp dest!: nth_mem; force)+
+
+end (* End of context for precompiled reachability problem *)
+
+
+locale UPPAAL_Reachability_Problem_precompiled_ceiling =
+  UPPAAL_Reachability_Problem_precompiled +
+  fixes k :: "nat list list list"
+  assumes k_ceiling:
+    "\<forall> i < p. \<forall> l < length (trans ! i). \<forall> (x, m) \<in> clkp_set'' i l. m \<le> k ! i ! l ! x"
+    "\<forall> i < p. \<forall> l < length (trans ! i). \<forall> (x, m) \<in> collect_clock_pairs (inv ! i ! l).
+      m \<le> k ! i ! l ! x"
+  and k_resets:
+    "\<forall> i < p. \<forall> l < length (trans ! i). \<forall> (g, a, r, l') \<in> set (trans ! i ! l).
+     \<forall> c \<in> {0..<m+1} - fst ` collect_store'' r. k ! i ! l' ! c \<le> k ! i ! l ! c"
+  and k_length:
+    "length k = p" "\<forall> i < p. length (k ! i) = length (trans ! i)"
+    "\<forall> xs \<in> set k. \<forall> xxs \<in> set xs. length xxs = m + 1"
+  and k_0:
+    "\<forall> i < p. \<forall> l < length (trans ! i). k ! i ! l ! 0 = 0"
+  and guaranteed_resets:
+    "\<forall> i < p. \<forall> l < length (trans ! i). \<forall> (g, a, r, l') \<in> set (trans ! i ! l).
+      guaranteed_execution_cond prog r max_steps
+     "
+begin
+
+definition "k_fun l c \<equiv> if c > 0 \<and> c \<le> m then Max {k ! i ! (fst l ! i) ! c | i . i < p} else 0"
+
+
+lemma p_p':
+  "equiv.p = p"
+  by simp
+
+lemma clkp_set_clk_set_subs:
+  "fst ` clkp_set A (l, s) \<subseteq> clk_set A"
+  unfolding TA_clkp_set_unfold by auto
+
+lemma k_ceiling_1:
+  "\<forall> l. \<forall>(x,m) \<in> clkp_set A l. m \<le> k_fun l x"
+
+  apply safe
+  subgoal premises prems for l s x d
+  proof -
+    from \<open>(x, d) \<in> _\<close> have "0 < x" "x \<le> m"
+      using clkp_set_clk_set_subs[of l s] clk_set by force+
+    from prems show ?thesis
+      unfolding clkp_set_def
+      apply safe
+      subgoal
+        unfolding collect_clki_def
+        unfolding inv_of_def
+        unfolding equiv.defs.prod_ta_def
+        unfolding equiv.defs.prod_invariant_def
+        unfolding inv_of_def
+          Product_TA_Defs.product_ta_def
+          Product_TA_Defs.product_invariant_def
+          equiv.defs.N_s_def
+        unfolding length_N
+        unfolding equiv.state_ta_def
+        unfolding p_p'
+        unfolding equiv.state_inv_def
+        unfolding N_def
+        unfolding collect_clock_pairs_def
+        apply (clarsimp cong: if_cong simp: I_def)
+        subgoal premises prems for l' i
+        proof -
+          have "nat d \<le> k ! i ! (l ! i) ! x"
+            using prems lengths k_ceiling(2)
+            unfolding collect_clock_pairs_def
+            by (auto 4 4)
+          also from \<open>_ < p\<close> have "\<dots> \<le> Max {k ! i ! (l ! i) ! x |i. i < p}"
+            by (auto intro: Max_ge)
+          finally show ?thesis
+            unfolding k_fun_def using \<open>0 < x\<close> \<open>x \<le> m\<close> by auto
+        qed
+        done
+      subgoal
+        unfolding collect_clkt_def
+        apply clarsimp
+        subgoal premises prems for g a r l' s'
+        proof -
+          from guard_approx[OF prems(2)] prems(1) obtain i g a r where *:
+            "(x, d) \<in> clkp_set'' i (l ! i)" "(g, a, r, l' ! i) \<in> set (trans ! i ! (l ! i))"
+            "l \<in> equiv.defs.states' s" "i < p"
+            by auto
+          from \<open>i < p\<close> \<open>l \<in> _\<close> have "l ! i < length (trans ! i)"
+            by auto
+          with k_ceiling(1) * have "nat d \<le> k ! i ! (l ! i) ! x"
+            by force
+          also from \<open>_ < p\<close> have "\<dots> \<le> Max {k ! i ! (l ! i) ! x |i. i < p}"
+            by (auto intro: Max_ge)
+          finally show ?thesis
+            unfolding k_fun_def using \<open>0 < x\<close> \<open>x \<le> m\<close> by auto
+        qed
+        done
+      done
+  qed
+done
+
+lemma k_ceiling_2:
+    "\<forall> l g a r l' c. A \<turnstile> l \<longrightarrow>g,a,r l' \<and> c \<notin> set r \<longrightarrow> k_fun l' c \<le> k_fun l c"
+  unfolding trans_of_def equiv.defs.prod_ta_def equiv.defs.prod_trans_def
+  apply clarsimp
+  apply safe
+  subgoal premises prems for l s g a r l' s' c
+  proof -
+    from prems obtain p' l'' pc_g pc_u where *:
+      "p' < p" "l' = l[p' := l'']"
+      "r = equiv.make_f pc_u s" "Some s' = equiv.make_mf pc_u s"
+      "(l ! p', pc_g, Sil a, pc_u, l'') \<in> fst (equiv.N ! p')"
+      "l \<in> equiv.defs.states' s"
+      apply atomize_elim
+      unfolding equiv.defs.prod_trans_i_def
+      unfolding Product_TA_Defs.product_ta_def Product_TA_Defs.product_trans_def trans_of_def
+      apply clarsimp
+      apply safe
+      unfolding Product_TA_Defs.product_trans_i_def
+      unfolding trans_of_def
+       apply clarsimp
+      unfolding equiv.defs.N_s_def
+      unfolding equiv.defs.T_s_def
+      unfolding Equiv_TA_Defs.state_ta_def
+      unfolding equiv.state_trans_t_def
+      unfolding Product_TA_Defs.product_trans_s_def
+      by auto
+    from \<open>l \<in> _\<close> have [simp]: "length l = p"
+      by simp
+    from \<open>r = _\<close> have "fst ` collect_store'' pc_u \<subseteq> set r"
+      supply find_resets_start.simps[simp del]
+      unfolding collect_store''_def equiv.make_f_def
+      apply (clarsimp split: option.split_asm)
+      subgoal
+        using \<open>Some s' = _\<close> unfolding equiv.make_mf_def
+        by (auto split: option.split_asm)
+      subgoal premises prems for pc' g st f pcs c d pc_t pc''
+      proof -
+        from prems have
+          "steps (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None)) max_steps
+            (pc_u, [], s, True, []) (pc', g, st, f, r)"
+          "prog ! pc' = Some (INSTR HALT)"
+          unfolding PF_unfold stripfp_def N_def PROG_def
+          by (auto dest!: exec_steps split: if_split_asm elim!: stripf.elims)
+        with prems show ?thesis
+          by (force intro: sym dest!: resets_start')
+      qed
+      done
+    with \<open>c \<notin> _\<close> have "c \<notin> fst ` collect_store'' pc_u" by blast
+    show ?thesis
+    proof (cases "c > m")
+      case True
+      then show ?thesis
+        unfolding k_fun_def by auto
+    next
+      case False
+      with \<open>c \<notin> fst ` _ _\<close> have "c \<in> {0..<m+1} - fst ` collect_store'' pc_u"
+        by auto
+      from * have "(l ! p') < length (trans ! p')"
+        unfolding N_def T_def by auto
+      from * have "(pc_g, Sil a, pc_u, l'') \<in> set (trans ! p' ! (l ! p'))"
+        "(l ! p') < length (trans ! p')"
+        unfolding N_def T_def by auto
+      with k_resets \<open>c \<in> _\<close> \<open>p' < _\<close> have "k ! p' ! l'' ! c \<le> k ! p' ! (l ! p') ! c"
+        unfolding k_fun_def by force
+      with \<open>l' = _\<close> show ?thesis
+        unfolding k_fun_def
+        apply clarsimp
+        apply (rule Max.boundedI)
+          apply force
+        using p_gt_0 apply force
+        apply clarsimp
+        subgoal for i
+          apply (cases "i = p'")
+           apply simp
+           apply (rule le_trans)
+          by (auto intro: Max_ge)
+        done
+    qed
+  qed
+  subgoal premises prems for l s g a r l' s' c
+  proof -
+    from prems obtain p1 l1 pc_g1 pc_u1 p2 l2 pc_g2 pc_u2 s'' where *:
+      "p1 < p" "p2 < p" "l' = l[p1 := l1, p2 := l2]"
+      "r = equiv.make_f pc_u1 s @ equiv.make_f pc_u2 s"
+      "Some s' = equiv.make_mf pc_u1 s''" "Some s'' = equiv.make_mf pc_u2 s"
+      "(l ! p1, pc_g1, In a, pc_u1, l1) \<in> fst (equiv.N ! p1)"
+      "(l ! p2, pc_g2, Out a, pc_u2, l2) \<in> fst (equiv.N ! p2)"
+      "l \<in> equiv.defs.states' s"
+      apply atomize_elim
+      unfolding equiv.defs.prod_trans_s_def
+      unfolding Product_TA_Defs.product_ta_def Product_TA_Defs.product_trans_def trans_of_def
+      apply clarsimp
+      apply safe
+      subgoal
+        unfolding Product_TA_Defs.product_trans_i_def
+        by auto
+      unfolding Product_TA_Defs.product_trans_s_def
+      unfolding trans_of_def
+      apply clarsimp
+      unfolding equiv.defs.N_s_def
+      unfolding equiv.defs.T_s_def
+      unfolding Equiv_TA_Defs.state_ta_def
+      unfolding equiv.state_trans_t_def
+      apply clarsimp
+      by blast
+    from \<open>l \<in> _\<close> have [simp]: "length l = p"
+      by simp
+    from * have **:
+      "(pc_g1, In a, pc_u1, l1) \<in> set (trans ! p1 ! (l ! p1))" "(l ! p1) < length (trans ! p1)"
+      "(pc_g2, Out a, pc_u2, l2) \<in> set (trans ! p2 ! (l ! p2))" "(l ! p2) < length (trans ! p2)"
+      unfolding N_def T_def by auto
+    with \<open>p1 < p\<close> guaranteed_resets have guaranteed_execution:
+      "guaranteed_execution_cond prog pc_u1 max_steps"
+      by blast
+    from \<open>r = _\<close> have "fst ` collect_store'' pc_u1 \<subseteq> set r"
+      supply find_resets_start.simps[simp del]
+      unfolding collect_store''_def
+        equiv.make_f_def
+      apply (clarsimp split: option.split_asm)
+      subgoal
+        using \<open>Some s'' = _\<close> unfolding equiv.make_mf_def
+        by (auto split: option.split_asm)
+      subgoal
+        using \<open>Some s'' = _\<close> unfolding equiv.make_mf_def
+        by (auto split: option.split_asm)
+      subgoal
+        using \<open>Some s' = _\<close> unfolding equiv.make_mf_def
+        using guaranteed_execution'[OF guaranteed_execution, of "[]" s True "[]" "[]"]
+        unfolding stripfp_def PROG_def by auto
+      subgoal premises prems for _ _ _ _ r2 _ pc' g st f r1 pcs c d pc_t pc''
+      proof -
+        from prems have
+          "steps (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None)) max_steps
+            (pc_u1, [], s, True, []) (pc', g, st, f, r1)"
+          "prog ! pc' = Some (INSTR HALT)" "r = r1 @ r2"
+          unfolding PF_unfold stripfp_def N_def PROG_def
+          by (auto dest!: exec_steps split: if_split_asm elim!: stripf.elims)
+        with prems show ?thesis
+          by (force intro: sym dest!: resets_start')
+      qed
+      done
+    moreover from \<open>r = _\<close> have "fst ` collect_store'' pc_u2 \<subseteq> set r"
+      supply find_resets_start.simps[simp del]
+      unfolding collect_store''_def
+        equiv.make_f_def
+      apply (clarsimp split: option.split_asm)
+      subgoal
+        using \<open>Some s'' = _\<close> unfolding equiv.make_mf_def
+        by (auto split: option.split_asm)
+      subgoal
+        using \<open>Some s'' = _\<close> unfolding equiv.make_mf_def
+        by (auto split: option.split_asm)
+      subgoal
+        using \<open>Some s' = _\<close> unfolding equiv.make_mf_def
+        using guaranteed_execution'[OF guaranteed_execution, of "[]" s True "[]" "[]"]
+        unfolding stripfp_def PROG_def by auto
+      subgoal premises prems for pc' g st f r1 pcs _ _ _ _ r2 _ c d pc_t pc''
+      proof -
+        from prems have
+          "steps (map_option stripf o (\<lambda>pc. if pc < size prog then prog ! pc else None)) max_steps
+            (pc_u2, [], s, True, []) (pc', g, st, f, r1)"
+          "prog ! pc' = Some (INSTR HALT)" "r = r2 @ r1"
+          unfolding PF_unfold stripfp_def N_def PROG_def
+          by (auto dest!: exec_steps split: if_split_asm elim!: stripf.elims)
+        with prems show ?thesis
+          by (force intro: sym dest!: resets_start')
+      qed
+      done
+    ultimately have c_not_elem: "c \<notin> fst ` collect_store'' pc_u1" "c \<notin> fst ` collect_store'' pc_u2"
+      using \<open>c \<notin> _\<close> by auto
+    show ?thesis
+    proof (cases "c > m")
+      case True
+      then show ?thesis
+        unfolding k_fun_def by auto
+    next
+      case False
+      with c_not_elem have
+        "c \<in> {0..<m+1} - fst ` collect_store'' pc_u1"
+        "c \<in> {0..<m+1} - fst ` collect_store'' pc_u2"
+        by auto
+      with ** k_resets \<open>p1 < _\<close> \<open>p2 < _\<close> have
+        "k ! p1 ! l1 ! c \<le> k ! p1 ! (l ! p1) ! c" "k ! p2 ! l2 ! c \<le> k ! p2 ! (l ! p2) ! c"
+        by (auto split: prod.split_asm)
+      with \<open>l' = _\<close> show ?thesis
+        unfolding k_fun_def
+        apply clarsimp
+        apply (rule Max.boundedI)
+          apply force
+        using p_gt_0 apply force
+        apply clarsimp
+        subgoal for i
+          apply (cases "i = p2")
+          subgoal
+            apply simp
+            apply (rule le_trans)
+            by (auto intro: Max_ge)
+          apply (cases "i = p1")
+           apply simp
+           apply (rule le_trans)
+          by (auto intro: Max_ge)
+        done
+    qed
+  qed
+  done
+
+lemma
+  shows k_ceiling':
+    "\<forall> l. \<forall>(x,m) \<in> clkp_set A l. m \<le> k_fun l x"
+    "\<forall> l g a r l' c. A \<turnstile> l \<longrightarrow>g,a,r l' \<and> c \<notin> set r \<longrightarrow> k_fun l' c \<le> k_fun l c"
+  and k_bound':
+    "\<forall> l. \<forall> i > m. k_fun l i = 0"
+  and k_0':
+    "\<forall> l. k_fun l 0 = 0"
+  using k_ceiling_1 k_ceiling_2 unfolding k_fun_def by auto
+
+sublocale Reachability_Problem A "(init, s0)" m k_fun "PR_CONST (\<lambda> (l, s). F l s)"
+  by (standard; rule k_ceiling' k_bound' k_0')
+
+end (* End of context for precompiled reachability problem with ceiling *)
+
+
+locale UPPAAL_Reachability_Problem_precompiled_start_state =
+  UPPAAL_Reachability_Problem_precompiled _ _ _ _ pred
+  for pred :: "nat list list" +
+  fixes s0 :: "int list" (* Why does nat not work? *)
+  assumes start_pred:
+    "\<forall> q < p. \<exists> pc st s' rs pcs.
+       exec (stripfp PROG) max_steps ((pred ! q ! (init ! q)), [], s0, True, []) []
+     = Some ((pc, st, s', True, rs), pcs)"
+      and bounded: "bounded bounds s0"
+      and pred_time_indep: "\<forall> x \<in> set pred. \<forall> pc \<in> set x. time_indep_check prog pc max_steps"
+      and upd_time_indep:
+        "\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (_, _, pc_u, _) \<in> set xs.
+           time_indep_check prog pc_u max_steps"
+     and clock_conj:
+       "\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (pc_g, _, _, _) \<in> set xs.
+           conjunction_check prog pc_g max_steps"
+begin
+
+  lemma B0[intro]:
+    "bounded defs'.B s0"
+    using bounded unfolding bounded_def N_def by simp
+
+  lemma equiv_P_simp:
+    "equiv.P = PROG"
+    unfolding N_def by simp
+
+  lemma time_indep_P[intro]:
+    "time_indep (conv_prog equiv.P) max_steps (pred ! q ! (L ! q), [], s, True, [])"
+    if "q < p" "L \<in> equiv.defs.states' s"
+  proof -
+    from that lengths process_length have "q < length pred" "L ! q < length (pred ! q)" by auto
+    then have "pred ! q \<in> set pred" "pred ! q ! (L ! q) \<in> set (pred ! q)" by auto
+    with pred_time_indep time_indep_overapprox show ?thesis
+      by (auto simp: PROG_def equiv_P_simp)
+  qed
+
+  lemma time_indep_PROG[intro]:
+    "time_indep (conv_prog PROG) max_steps (pc_u, [], s, True, [])"
+    if "q < p" "(l, pc_g, a, pc_u, l') \<in> T q"
+  proof -
+    from that lengths process_length have
+      "q < length trans" "l < length (trans ! q)" "(pc_g, a, pc_u, l') \<in> set (trans ! q ! l)"
+      unfolding T_def by auto
+    moreover then have "trans ! q \<in> set trans" "trans ! q ! l \<in> set (trans ! q)" by auto
+    ultimately show ?thesis using upd_time_indep time_indep_overapprox
+      unfolding PROG_def by blast
+  qed
+
+  lemma facts[intro]:
+    "u \<turnstile>a ac" if
+    "q < defs'.p"
+    "(l, pc_g, a, pc_u, l') \<in> fst (defs'.N ! q)"
+    "stepst defs'.P max_steps u (pc_g, [], s, True, []) (pc_t, st_t, s_t, True, rs_t)"
+    "stepsc defs'.P max_steps u (pc_g, [], s, True, []) (pc', st, s', f', rs)"
+    "defs'.P pc' = Some (CEXP ac)"
+  proof -
+    let ?P = "conv_P prog"
+    from that(5) obtain ac' where
+      "ac = conv_ac ac'" "prog ! pc' = Some (CEXP ac')" "pc' < length prog"
+      apply (clarsimp split: option.split_asm if_split_asm simp add: PROG_def N_def)
+      subgoal for z
+        by (cases z) auto
+      done
+    with that have "u \<turnstile>a conv_ac ac'"
+      apply -
+      apply (rule conjunction_check)
+      using clock_conj apply simp_all
+      unfolding N_def apply simp_all
+      using lengths process_length(2) by (force dest!: nth_mem simp: PROG_def N_def T_def)+
+    with \<open>ac = _\<close> show ?thesis by simp
+  qed
+
+  sublocale product':
+    Equiv_TA "conv N" max_steps init s0
+    apply standard
+          apply standard
+         apply (simp; blast)
+         subgoal
+           apply clarsimp
+           apply (force simp: N_def)
+           done
+       apply blast
+      apply (simp; fail)
+    unfolding PF_PF using start_pred apply simp
+    by (rule B0)
+
+  lemma fst_S[simp]:
+    "fst ` (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` S = fst ` S"
+    by force
+
+  lemma snd_S[simp]:
+    "(snd \<circ> snd \<circ> snd \<circ> snd) ` (\<lambda>(l, g, a, r, l'). (l, map conv_ac g, a, r, l')) ` S
+    = (snd \<circ> snd \<circ> snd \<circ> snd) ` S"
+    by force
+
+end (* End of locale *)
+
+locale UPPAAL_Reachability_Problem_precompiled' =
+  UPPAAL_Reachability_Problem_precompiled_start_state +
+  UPPAAL_Reachability_Problem_precompiled_defs' +
+  UPPAAL_Reachability_Problem_precompiled_ceiling +
+  assumes action_set:
+    "\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (_, a, _) \<in> set xs. pred_act (\<lambda> a. a < na) a"
+begin
+
+  (* Why are we re-doing this here? *)
+  sublocale Reachability_Problem_Impl_Defs _ _ A "(init, s0)" m k_fun "PR_CONST (\<lambda> (l, s). F l s)" .
+
+  definition
+    "states' = {(L, s). L \<in> equiv.defs.states' s \<and> check_pred L s \<and> length s = length bounds}"
+
+  lemma in_trans_in_mapI:
+    assumes
+      "q < p" "l < length (trans ! q)" "i < length (trans ! q ! l)"
+      "(g1, In a, r1) = trans ! q ! l ! i"
+    shows "(g1, a, r1) \<in> set (IArray (map IArray trans_in_map) !! q !! l)"
+    using assms process_length(2) unfolding trans_in_map_def
+    by (force dest: nth_mem intro!: image_eqI[where x = "(g1, In a, r1)"])
+
+  lemma in_trans_out_mapI:
+    assumes
+      "q < p" "l < length (trans ! q)" "i < length (trans ! q ! l)"
+      "(g1, Out a, r1) = trans ! q ! l ! i"
+    shows "(g1, a, r1) \<in> set (IArray (map IArray trans_out_map) !! q !! l)"
+    using assms process_length(2) unfolding trans_out_map_def
+    by (force dest: nth_mem intro!: image_eqI[where x = "(g1, Out a, r1)"])
+
+  lemma in_trans_in_mapD:
+    assumes
+      "(g1, a, r1) \<in> set (IArray (map IArray trans_in_map) !! q !! l)"
+      "q < p" "l < length (trans ! q)"
+    obtains i where
+      "i < length (trans ! q ! l) \<and> trans ! q ! l ! i = (g1, In a, r1)"
+    using assms process_length(2) unfolding trans_in_map_def
+    by (fastforce dest: mem_nth split: act.split_asm)
+
+  lemma in_trans_out_mapD:
+    assumes
+      "(g1, a, r1) \<in> set (IArray (map IArray trans_out_map) !! q !! l)"
+      "q < p" "l < length (trans ! q)"
+    obtains i where
+      "i < length (trans ! q ! l) \<and> trans ! q ! l ! i = (g1, Out a, r1)"
+    using assms process_length(2) unfolding trans_out_map_def
+    by (fastforce dest: mem_nth split: act.split_asm)
+
+  lemma in_actions_by_stateI:
+    assumes
+      "(g1, a, r1) \<in> set xs" "a < length acc"
+    shows
+      "(q, g1, a, r1) \<in> set (actions_by_state q xs acc ! a)
+      \<and> a < length (actions_by_state q xs acc)"
+    using assms unfolding actions_by_state_def
+    apply (induction xs arbitrary: acc)
+     apply (simp; fail)
+    apply simp
+    apply (erule disjE)
+     apply (rule fold_acc_preserv
+        [where P = "\<lambda> acc. (q, g1, a, r1) \<in> set (acc ! a) \<and> a < length acc"]
+        )
+      apply (subst list_update_nth_split; auto)
+    by auto
+
+  lemma in_actions_by_state_preserv:
+    assumes
+      "(q, g1, a, r1) \<in> set (acc ! a)" "a < length acc"
+    shows
+      "(q, g1, a, r1) \<in> set (actions_by_state y xs acc ! a)
+      \<and> a < length (actions_by_state y xs acc)"
+    using assms unfolding actions_by_state_def
+    apply -
+    apply (rule fold_acc_preserv
+        [where P = "\<lambda> acc. (q, g1, a, r1) \<in> set (acc ! a) \<and> a < length acc"]
+        )
+    apply (subst list_update_nth_split; auto)
+    by auto
+
+  lemma length_actions_by_state_preserv[simp]:
+    shows "length (actions_by_state y xs acc) = length acc"
+    unfolding actions_by_state_def by (auto intro: fold_acc_preserv simp: list_update_nth_split)
+
+  lemma in_all_actions_by_stateI:
+    assumes
+      "a < na" "q < p" "(g1, a, r1) \<in> set (M !! q !! (L ! q))"
+    shows
+      "(q, g1, a, r1) \<in> set (all_actions_by_state M L ! a)"
+    unfolding all_actions_by_state_def
+    apply (rule fold_acc_ev_preserv
+        [where P = "\<lambda> acc. (q, g1, a, r1) \<in> set (acc ! a)" and Q = "\<lambda> acc. a < length acc",
+          THEN conjunct1]
+        )
+        apply (rule in_actions_by_state_preserv[THEN conjunct1])
+    using assms by (auto intro: in_actions_by_stateI[THEN conjunct1])
+
+  lemma actions_by_state_inj:
+    assumes "j < length acc"
+    shows "\<forall> (q, a) \<in> set (actions_by_state i xs acc ! j). (q, a) \<notin> set (acc ! j) \<longrightarrow> i = q"
+    unfolding actions_by_state_def
+    apply (rule fold_acc_preserv
+        [where P =
+          "\<lambda> acc'. (\<forall> (q, a) \<in> set (acc' ! j). (q, a) \<notin> set (acc ! j) \<longrightarrow> i = q) \<and> j < length acc'",
+          THEN conjunct1])
+    subgoal for x acc
+      by (cases "fst (snd x) = j"; simp)
+    using assms by auto
+
+  lemma actions_by_state_inj':
+    assumes "j < length acc" "(q, a) \<notin> set (acc ! j)" "(q, a) \<in> set (actions_by_state i xs acc ! j)"
+    shows "i = q"
+    using actions_by_state_inj[OF assms(1)] assms(2-) by fast
+
+  lemma in_actions_by_stateD:
+    assumes
+      "(q, g, a, t) \<in> set (actions_by_state i xs acc ! j)" "(q, g, a, t) \<notin> set (acc ! j)"
+      "j < length acc"
+    shows
+      "(g, a, t) \<in> set xs \<and> j = a"
+    using assms unfolding actions_by_state_def
+    apply -
+    apply (drule fold_evD
+        [where y = "(g, a, t)" and Q = "\<lambda> acc'. length acc' = length acc"
+          and R = "\<lambda> (_, a', t). a' = j"]
+        )
+         apply assumption
+        apply (subst (asm) list_update_nth_split[of j]; force)
+       apply simp+
+     apply (subst (asm) list_update_nth_split[of j]; force)
+    by auto
+
+  lemma in_all_actions_by_stateD:
+    assumes
+      "(q, g1, a, r1) \<in> set (all_actions_by_state M L ! a')" "a' < na"
+    shows
+      "(g1, a, r1) \<in> set (M !! q !! (L ! q)) \<and> q < p \<and> a' = a"
+    using assms
+    unfolding all_actions_by_state_def
+    apply -
+    apply (drule fold_evD''[where y = q and Q = "\<lambda> acc. length acc = na"])
+        apply (simp; fail)
+       apply (drule actions_by_state_inj'[rotated])
+         apply (simp; fail)+
+    apply safe
+      apply (drule in_actions_by_stateD)
+        apply assumption
+       apply (rule fold_acc_preserv)
+        apply (simp; fail)+
+    subgoal premises prems
+    proof -
+      from prems(2) have "q \<in> set [0..<p]" by auto
+      then show ?thesis by simp
+    qed
+    by (auto intro: fold_acc_preserv dest!: in_actions_by_stateD)
+
+  lemma length_all_actions_by_state_preserv:
+      "length (all_actions_by_state M L) = na"
+    unfolding all_actions_by_state_def by (auto intro: fold_acc_preserv)
+
+  lemma less_naI:
+    assumes
+      "q < p"
+      "(g1, a, r1) = trans ! q ! l ! j"
+      "l < length (trans ! q)"
+      "j < length (trans ! q ! l)"
+    shows "pred_act (\<lambda>a. a < na) a"
+    using action_set assms process_length(2) by (force dest!: nth_mem)
+
+  lemma in_actions_trans_in_mapI:
+    assumes
+      "pa < p"
+      "(g1, In a, r1) = trans ! pa ! (L ! pa) ! j"
+      "L ! pa < length (trans ! pa)"
+      "j < length (trans ! pa ! (L ! pa))"
+    shows "(pa, g1, a, r1) \<in> set (all_actions_by_state (IArray (map IArray trans_in_map)) L ! a)"
+    apply (rule in_all_actions_by_stateI)
+    using assms action_set process_length(2) apply (fastforce dest!: nth_mem)
+    using assms by (fastforce intro: in_trans_in_mapI)+
+
+  lemma in_actions_trans_out_mapI:
+    assumes
+      "pa < p"
+      "(g1, Out a, r1) = trans ! pa ! (L ! pa) ! j"
+      "L ! pa < length (trans ! pa)"
+      "j < length (trans ! pa ! (L ! pa))"
+    shows "(pa, g1, a, r1) \<in> set (all_actions_by_state (IArray (map IArray trans_out_map)) L ! a)"
+    apply (rule in_all_actions_by_stateI)
+    using assms action_set process_length(2) apply (fastforce dest!: nth_mem)
+    using assms by (fastforce intro: in_trans_out_mapI)+
+
+  lemma in_pairs_by_actionD2:
+    assumes
+      "(g, a, r, L', s') \<in> set (pairs_by_action (L, s) xs ys)"
+      "\<forall> (q, g, a'', m, l) \<in> set xs. a'' = a'"
+      "\<forall> (q, g, a'', m, l) \<in> set ys. a'' = a'"
+    shows "check_pred L' s'"
+    using assms(1) unfolding pairs_by_action_def using assms(2,3)
+    by (clarsimp split: option.split_asm simp: set_map_filter) (clarsimp split: if_split_asm)
+
+  lemma in_pairs_by_actionD1:
+    assumes
+      "(g, a, r, L', s') \<in> set (pairs_by_action (L, s) xs ys)"
+      "\<forall> (q, g, a'', m, l) \<in> set xs. a'' = a'"
+      "\<forall> (q, g, a'', m, l) \<in> set ys. a'' = a'"
+    obtains
+      pa q pc_g1 pc_g2 g1 g2 r1 r2 pc_u1 pc_u2 l1' l2' s1
+      x1 x2 x3 x4 x5 y1 y2 y3 y4
+    where
+      "pa \<noteq> q"
+      "(pa, pc_g1, a, pc_u1, l1') \<in> set ys"
+      "(q, pc_g2, a, pc_u2, l2') \<in> set xs"
+      "L' = L[pa := l1', q := l2']"
+      "runf pc_u1 s1 = Some ((x1, x2, s', x3, x4), x5)"
+      "runf pc_u2 s = Some ((y1, y2, s1, y3, r2), y4)"
+      "check_g pc_g1 s = Some g1" "check_g pc_g2 s = Some g2"
+      "r1 = make_reset pc_u1 s"
+      "g = g1 @ g2" "r = r1 @ r2"
+  proof -
+    obtain
+      pa q pc_g1 pc_g2 g1 g2 r1 r2 pc_u1 pc_u2 l1' l2' s1
+      x1 x2 x3 x4 x5 y1 y2 y3 y4
+      where
+      "(q, pc_g1, a, pc_u1, l1') \<in> set ys" "(pa, pc_g2, a, pc_u2, l2') \<in> set xs" "q \<noteq> pa"
+      "check_g pc_g1 s = Some g1" "check_g pc_g2 s = Some g2"
+      "runf pc_u1 s1 = Some ((x1, x2, s', x3, x4), x5)"
+      "runf pc_u2 s = Some ((y1, y2, s1, y3, r2), y4)"
+      "r1 = make_reset pc_u1 s"
+      "Some (g1 @ g2, a, r1 @ r2, L[q := l1', pa := l2'], s') = Some (g, a, r, L', s')"
+    proof -
+      from assms(1) show ?thesis
+      unfolding pairs_by_action_def using assms(2,3)
+      apply clarsimp
+      unfolding set_map_filter
+      apply clarsimp
+      apply (clarsimp split: option.split_asm if_split_asm)
+      by (force intro!: that)
+    qed
+    then show ?thesis by (fast intro: that)
+  qed
+
+  lemma in_pairs_by_actionD:
+    assumes
+      "(g, a, r, L', s') \<in> set (pairs_by_action (L, s) xs ys)"
+      "\<forall> (q, g, a'', m, l) \<in> set xs. a'' = a'"
+      "\<forall> (q, g, a'', m, l) \<in> set ys. a'' = a'"
+    obtains
+        pa q pc_g1 pc_g2 p1 p2 g1 g2 r1 r2 pc_u1 pc_u2 l1' l2' s1
+      x1 x2 x3 x4 x5 y1 y2 y3 y4
+    where
+      "pa \<noteq> q"
+      "(pa, pc_g1, a, pc_u1, l1') \<in> set ys"
+      "(q, pc_g2, a, pc_u2, l2') \<in> set xs"
+      "L' = L[pa := l1', q := l2']"
+      (*"s' = m1 (m2 s)"*)
+      "runf pc_u1 s1 = Some ((x1, x2, s', x3, x4), x5)"
+      "runf pc_u2 s = Some ((y1, y2, s1, y3, r2), y4)"
+      "check_g pc_g1 s = Some g1" "check_g pc_g2 s = Some g2"
+      "r1 = make_reset pc_u1 s"
+      "g = g1 @ g2" "r = r1 @ r2"
+      (* "\<forall> q' < p. (P ! q') (L' ! q') s'" *)
+      "check_pred L' s'"
+      (* Some ((_, _, s1, _, r2), _) *)
+    using in_pairs_by_actionD1[OF assms] in_pairs_by_actionD2[OF assms] by metis
+
+  lemma in_trans_funD:
+    assumes "y \<in> set (trans_fun L)"
+    shows "y \<in> set (trans_s_fun L) \<or> y \<in> set (trans_i_fun L)"
+      using assms unfolding trans_fun_def by auto
+
+  lemma states'_states'[intro]:
+    "L \<in> equiv.defs.states' s" if "(L, s) \<in> states'"
+    using that unfolding states'_def by auto
+
+  lemma bounded'_bounded:
+    "bounded' s \<longleftrightarrow> bounded bounds s" if "length s = length bounds"
+    using that unfolding bounded'_def bounded_def by simp
+
+  lemma bounded_bounded':
+    "bounded bounds s \<Longrightarrow> bounded' s"
+    unfolding bounded'_def bounded_def by simp
+
+  lemma P_unfold:
+    "(\<forall>q<p. (equiv.defs.P ! q) (L ! q) s) \<longleftrightarrow> (check_pred L s)" if "length s = length bounds"
+    unfolding equiv.state_ta_def equiv.state_pred_def check_pred_def using process_length(3) that
+    apply simp
+    unfolding list_all_iff
+    unfolding N_def
+    unfolding runf_def P_def
+    by (fastforce split: option.splits simp: bounded'_bounded)
+
+  lemma P_unfold_1:
+    "(\<forall>q<p. (equiv.defs.P ! q) (L ! q) s) \<Longrightarrow> (check_pred L s)"
+    unfolding equiv.state_ta_def equiv.state_pred_def check_pred_def using process_length(3)
+    apply simp
+    unfolding list_all_iff
+    unfolding N_def
+    unfolding runf_def P_def
+    by (fastforce split: option.split_asm simp: bounded_bounded')
+
+  lemma equiv_PT[simp]:
+    "equiv.PT = PT"
+    unfolding striptp_def N_def by simp
+
+  lemmas [simp] = equiv_P_simp
+
+  lemma transD:
+    assumes
+      "(pc_g, a, pc_u, l') = trans ! q ! (L ! q) ! j"
+      "L ! q < length (trans ! q)" "j < length (trans ! q ! (L ! q))"
+      "q < p"
+    shows "(L ! q, pc_g, a, pc_u, l') \<in> fst (equiv.N ! q)"
+    using assms unfolding N_def T_def by simp solve_ex_triv
+
+  lemma trans_ND:
+    assumes
+      "(L ! q, pc_g, a, pc_u, l') \<in> fst (equiv.N ! q)"
+      "q < p"
+    shows
+      "equiv.defs.N_s s ! q \<turnstile> L ! q
+        \<longrightarrow>equiv.make_g pc_g s,(a, equiv.make_c pc_g, equiv.make_mf pc_u),equiv.make_f pc_u s l'"
+    using assms
+    unfolding equiv.defs.N_s_def trans_of_def equiv.defs.T_s_def
+    unfolding equiv.state_ta_def equiv.state_trans_t_def
+    by clarsimp solve_ex_triv+
+
+  lemma make_f_unfold:
+    "equiv.make_f pc s = make_reset pc s"
+    unfolding make_reset_def equiv.make_f_def runf_def by simp
+
+  lemma make_g_simp:
+    assumes "check_g pc_g s = Some g1"
+    shows "g1 = equiv.make_g pc_g s"
+    using assms unfolding check_g_def equiv.make_g_def runt_def
+    by (clarsimp split: option.splits bool.splits simp: make_cconstr_def)
+
+  lemma make_c_simp:
+    assumes "check_g pc_g s = Some g1"
+    shows "equiv.make_c pc_g s"
+    using assms unfolding check_g_def equiv.make_c_def runt_def
+    by (clarsimp split: option.splits bool.splits simp: make_cconstr_def)
+
+  lemma make_reset_simp:
+    assumes "runf pc_u s = Some ((y1, y2, s1, y3, r2), y4)"
+    shows "make_reset pc_u s = r2"
+    using assms unfolding runf_def make_reset_def by (auto split: option.splits)
+
+  lemma make_mf_simp:
+    assumes "runf pc_u s = Some ((y1, y2, s1, y3, r2), y4)"
+    shows "equiv.make_mf pc_u s = Some s1"
+    using assms unfolding runf_def equiv.make_mf_def by (auto split: option.splits)
+
+  lemma trans_fun_trans_of':
+    "(trans_fun, trans_of A) \<in> transition_rel states'"
+    unfolding transition_rel_def T_def
+    apply simp
+    unfolding trans_of_def
+    apply safe
+    subgoal for L s g a r L' s'
+      unfolding equiv.defs.prod_ta_def equiv.defs.prod_trans_def
+      apply simp
+      apply safe
+      subgoal
+        apply (rule trans_i_fun_trans_fun)
+        unfolding equiv.defs.prod_trans_i_alt_def
+        apply safe
+          unfolding trans_fun_def trans_i_from_def trans_i_fun_def
+          unfolding Product_TA_Defs.product_trans_i_def
+          apply clarsimp
+          subgoal premises prems for c m p' l'
+          proof -
+            from prems have "L ! p' < length (trans ! p')" by auto
+            from prems obtain pc_g pc_u where
+              "(L ! p', pc_g, Sil a, pc_u, l') \<in> T p'"
+              "g = equiv.make_g pc_g s" "r = equiv.make_f pc_u s"
+              "c = equiv.make_c pc_g" "m = equiv.make_mf pc_u"
+              unfolding equiv.defs.N_s_def trans_of_def equiv.defs.T_s_def
+              unfolding equiv.state_ta_def equiv.state_trans_t_def
+              apply clarsimp
+              unfolding N_def T_def by clarsimp
+            from this(1) have "(pc_g, Sil a, pc_u, l') \<in> set (trans ! p' ! (L ! p'))"
+              unfolding T_def by auto
+            moreover have "check_g pc_g s = Some g"
+              using \<open>g = _\<close> \<open>c = _\<close> \<open>c s\<close>
+                unfolding check_g_def equiv.make_g_def equiv.make_c_def
+                by (auto split: option.splits simp: runt_def make_cconstr_def)
+            moreover obtain x1 x2 x3 pcs where "runf pc_u s = Some ((x1, x2, s', x3, r), pcs)"
+              using \<open>r = _\<close> \<open>m = _\<close> prems(5)
+              unfolding equiv.make_f_def equiv.make_mf_def runf_def trans_of_def
+              by (auto split: option.splits)
+            moreover have "check_pred (L[p' := l']) s'"
+              using prems(3) by (auto intro: P_unfold_1)
+            ultimately show ?thesis using process_length(2) \<open>p' < _\<close> \<open>L ! p' < _\<close>
+              by (force simp: set_map_filter trans_i_map_def)
+          qed
+          done
+        subgoal
+          apply (rule trans_s_fun_trans_fun)
+            unfolding equiv.defs.prod_trans_s_alt_def
+            apply safe
+          unfolding trans_fun_def trans_s_fun_def
+          unfolding Product_TA_Defs.product_trans_s_def
+          apply clarsimp
+          subgoal premises prems for ci co mi mo s1 q1 q2 g1 g2 r1 r2 l1' l2'
+          proof -
+            from prems have "L ! q1 < length (trans ! q1)" "L ! q2 < length (trans ! q2)" by auto
+            from prems obtain pc_g1 pc_u1 where
+              "(L ! q1, pc_g1, In a, pc_u1, l1') \<in> T q1"
+              "g1 = equiv.make_g pc_g1 s" "r1 = equiv.make_f pc_u1 s"
+              "ci = equiv.make_c pc_g1" "mi = equiv.make_mf pc_u1"
+              unfolding equiv.defs.N_s_def trans_of_def equiv.defs.T_s_def
+              unfolding equiv.state_ta_def equiv.state_trans_t_def
+              apply clarsimp
+              unfolding N_def T_def by clarsimp
+            from prems obtain pc_g2 pc_u2 where
+              "(L ! q2, pc_g2, Out a, pc_u2, l2') \<in> T q2"
+              "g2 = equiv.make_g pc_g2 s" "r2 = equiv.make_f pc_u2 s"
+              "co = equiv.make_c pc_g2" "mo = equiv.make_mf pc_u2"
+              unfolding equiv.defs.N_s_def trans_of_def equiv.defs.T_s_def
+              unfolding equiv.state_ta_def equiv.state_trans_t_def
+              apply clarsimp
+              unfolding N_def T_def by clarsimp
+            from \<open>_ \<in> T q1\<close> have "(pc_g1, In a, pc_u1, l1') \<in> set (trans ! q1 ! (L ! q1))"
+              unfolding T_def by auto
+            from \<open>_ \<in> T q2\<close> have "(pc_g2, Out a, pc_u2, l2') \<in> set (trans ! q2 ! (L ! q2))"
+              unfolding T_def by auto
+            moreover have "check_g pc_g1 s = Some g1" "check_g pc_g2 s = Some g2"
+              using \<open>g1 = _\<close> \<open>ci = _\<close> \<open>ci s\<close> \<open>g2 = _\<close> \<open>co = _\<close> \<open>co s\<close>
+                unfolding check_g_def equiv.make_g_def equiv.make_c_def
+                by (auto split: option.splits simp: runt_def make_cconstr_def)
+            moreover obtain x1 x2 x3 pcs where "runf pc_u2 s = Some ((x1, x2, s1, x3, r2), pcs)"
+              using \<open>r2 = _\<close> \<open>mo = _\<close> \<open>Some s1 = _\<close>
+              unfolding equiv.make_f_def equiv.make_mf_def runf_def trans_of_def
+              by (auto split: option.splits)
+            moreover obtain x1 x2 x3 x4 pcs where "runf pc_u1 s1 = Some ((x1, x2, s', x3, x4), pcs)"
+              using \<open>r1 = _\<close> \<open>mi = _\<close> \<open>Some s' = _\<close>
+              unfolding equiv.make_f_def equiv.make_mf_def runf_def trans_of_def
+              by (auto split: option.splits)
+             moreover have "r1 = make_reset pc_u1 s"
+              using \<open>r1 = _\<close> unfolding make_reset_def equiv.make_f_def runf_def by auto
+            moreover have "check_pred (L[q1 := l1', q2 := l2']) s'"
+              using prems(5) by (auto intro: P_unfold_1)
+            moreover have "a < na"
+                using action_set \<open>_ \<in> set (trans ! q1 ! (L ! q1))\<close> \<open>q1 < _\<close> \<open>L ! q1 < _\<close>
+                      process_length(2)
+                by (fastforce dest!: nth_mem)
+              moreover have
+                "(q1, pc_g1, a, pc_u1, l1')
+                \<in> set (all_actions_by_state (nested_list_to_iarray trans_in_map) L ! a)"
+              using \<open>L ! q1 < _\<close> \<open>_ \<in> set (trans ! q1 ! (L ! q1))\<close> \<open>q1 < p\<close>
+              by (force intro: in_actions_trans_in_mapI dest: mem_nth)
+            moreover have
+              "(q2, pc_g2, a, pc_u2, l2')
+              \<in> set (all_actions_by_state (nested_list_to_iarray trans_out_map) L ! a)"
+              using \<open>L ! q2 < _\<close> \<open>_ \<in> set (trans ! q2 ! (L ! q2))\<close> \<open>q2 < p\<close>
+              by (force intro: in_actions_trans_out_mapI dest: mem_nth)
+            ultimately show ?thesis
+              using process_length(2) \<open>q1 < _\<close> \<open>L ! q1 < _\<close> \<open>q2 < _\<close> \<open>L ! q2 < _\<close> \<open>_ \<noteq> _\<close>
+              unfolding pairs_by_action_def
+                apply -
+                apply (rule bexI[where x = a])
+                by auto (force simp: set_map_filter)
+          qed
+          done
+        done
+      subgoal for L s g a r L' s'
+      apply (drule in_trans_funD)
+      apply (erule disjE)
+      unfolding equiv.defs.prod_ta_def equiv.defs.prod_trans_def
+       apply simp
+       apply (rule disjI2)
+      subgoal
+        unfolding equiv.defs.prod_trans_s_alt_def
+        apply safe
+        unfolding trans_s_fun_def
+        apply clarsimp
+          subgoal for x
+        apply (erule in_pairs_by_actionD[where a' = x])
+        apply (solves \<open>auto dest: in_all_actions_by_stateD\<close>)
+             apply (solves \<open>auto dest: in_all_actions_by_stateD\<close>)
+              apply (drule in_all_actions_by_stateD, assumption)
+          apply (drule in_all_actions_by_stateD, assumption)
+            apply safe
+          apply (erule in_trans_in_mapD)
+              apply (simp; fail)
+             apply blast
+              apply (erule in_trans_out_mapD)
+            apply blast
+             apply blast
+            apply (simp only: ex_simps[symmetric])
+            unfolding states'_def
+            apply (clarsimp)
+            apply (subst P_unfold, assumption)
+            apply (subst P_unfold)
+            subgoal
+              unfolding runf_def by (auto dest!: exec_state_length)
+            apply simp
+            apply (drule transD[rotated 2], solve_triv+, blast)
+            apply (drule transD[rotated 2], solve_triv+, blast)
+            apply (drule_tac s = s in trans_ND, assumption)
+            apply (drule_tac s = s in trans_ND, assumption)
+            unfolding Product_TA_Defs.product_trans_s_def
+            apply clarsimp
+            unfolding trans_of_def
+            subgoal
+              apply (simp only: ex_simps[symmetric])
+              apply defer_ex
+              apply defer_ex
+              apply solve_ex_triv
+              apply solve_ex_triv
+              apply solve_ex_triv
+              unfolding make_f_unfold
+              by (auto simp add: make_c_simp make_mf_simp make_reset_simp make_g_simp[symmetric])
+            done
+          done
+      subgoal
+        apply simp
+        apply (rule disjI1)
+        using process_length(2)
+        unfolding equiv.defs.prod_trans_i_alt_def
+        apply simp
+        unfolding P_unfold
+        unfolding trans_i_fun_def trans_i_from_def states'_def
+        apply simp
+        apply (erule bexE)
+
+        unfolding set_map_filter
+        apply simp
+        subgoal premises prems for q
+        proof -
+          from prems have len: "q < length trans" "L ! q < length (trans ! q)" by auto
+          from prems(4) obtain pc_g pc_u l' x1 x2 x3 x4 where
+            "(pc_g, a, pc_u, l') \<in> set (IArray.list_of (map IArray trans_i_map ! q) ! (L ! q))"
+            "check_g pc_g s = Some g"
+            "r = make_reset pc_u s"
+            "runf pc_u s = Some ((x1, x2, s', x3, r), x4)"
+            "check_pred (L[q := l']) s'"
+            "L' = L[q := l']"
+            apply atomize_elim
+              unfolding make_reset_def
+              by (force split: option.splits if_split_asm)
+          moreover then have
+              "(L, g, (a, Networks.label.Act (equiv.make_c pc_g, equiv.make_mf pc_u)), r, L')
+              \<in> Product_TA_Defs.product_trans_i (equiv.defs.N_s s)"
+          unfolding Product_TA_Defs.product_trans_i_def
+          apply clarsimp
+          apply solve_ex_triv
+          apply safe
+          using prems apply simp
+          unfolding trans_i_map_def
+          using len \<open>q < p\<close> apply (clarsimp simp: set_map_filter)
+            apply (clarsimp split: act.split_asm)
+           apply (frule make_c_simp)
+           apply (drule mem_nth)
+           apply safe
+           apply (drule transD[rotated], solve_triv+)
+           apply (drule trans_ND)
+            apply solve_triv
+             apply (subst make_g_simp)
+          using \<open>q < p\<close> prems(1) by (auto simp add: make_f_unfold)
+        ultimately show ?thesis
+          apply (subst P_unfold)
+          subgoal
+            using prems(1) by fast
+          apply (subst P_unfold)
+          subgoal
+            using \<open>runf _ _ = _\<close> prems(1)
+            unfolding runf_def by (auto dest!: exec_state_length)
+          using prems(1) by (force simp: make_mf_simp dest: make_c_simp)
+      qed
+      done
+    done
+  done
+
+  (* Unused *)
+  lemma transition_rel_mono:
+    "(a, b) \<in> transition_rel B" if "(a, b) \<in> transition_rel C" "B \<subseteq> C"
+    using that unfolding transition_rel_def b_rel_def fun_rel_def by auto
+
+end
+
+context
+  Equiv_TA_Defs
+begin
+
+  lemma state_set_subs:
+    "state_set (trans_of (defs.product s''))
+  \<subseteq> {L. length L = defs.p \<and> (\<forall>q<defs.p. L ! q \<in> State_Networks.state_set (fst (defs.N ! q)))}"
+    unfolding defs.states'_alt_def[symmetric]
+    unfolding defs.N_s_def
+    unfolding state_set_def Product_TA_Defs.states_def
+    unfolding trans_of_def
+    unfolding Product_TA_Defs.product_ta_def Product_TA_Defs.product_trans_def
+    unfolding Product_TA_Defs.product_trans_i_def Product_TA_Defs.product_trans_s_def
+    unfolding defs.T_s_def
+      unfolding Product_TA_Defs.states_def trans_of_def
+      apply simp
+      apply safe
+              apply (simp; fail)
+      using [[goals_limit = 1]]
+             apply force
+            apply force
+           apply (force simp: image_iff)
+         apply force
+        subgoal for x _ _ _ _ _ _ L p g _ _ _ r l' q
+          apply (cases "p = q")
+          apply (force simp: image_iff)
+          apply (force simp: image_iff)
+          done
+       apply force
+      subgoal for _ _ _ _ _ _ _ _ p q g1 g2 _ _ _ _ _ r1 r2 l1' l2' qa
+        apply (cases "qa = q")
+        apply force
+        apply (cases "qa = p")
+        by (force simp: image_iff)+
+      done
+
+end
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  lemma N_s_state_indep:
+    assumes "(L ! q, g, a, r, l') \<in> map trans_of (equiv.defs.N_s s) ! q" "q < p"
+    obtains g r where "(L ! q, g, a, r, l') \<in> map trans_of (equiv.defs.N_s s') ! q"
+    using assms unfolding trans_of_def equiv.defs.N_s_def equiv.defs.T_s_def by force
+
+  lemma fst_product_state_indep:
+    "fst ` fst (equiv.defs.product s) = fst ` fst (equiv.defs.product s')"
+    unfolding Product_TA_Defs.product_ta_def Product_TA_Defs.product_trans_def
+    unfolding Product_TA_Defs.product_trans_s_def Product_TA_Defs.product_trans_i_def
+    apply simp
+    unfolding equiv.defs.states'_alt_def
+    apply clarsimp
+    apply (rule equalityI)
+    subgoal
+      apply (rule subsetI)
+      apply clarsimp
+      apply (erule disjE)
+       apply (erule conjE exE)+
+       apply (erule N_s_state_indep)
+        apply (simp; fail)
+       apply (rule img_fst)
+       apply (rule Set.UnI1)
+       apply (subst mem_Collect_eq)
+       apply solve_ex_triv+
+      apply (erule conjE exE)+
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (rule img_fst)
+      apply (rule Set.UnI2)
+      apply (subst mem_Collect_eq)
+      apply (rule exI)+
+      apply (rule conjI)
+       defer
+      by solve_ex_triv+
+    subgoal
+      apply (rule subsetI)
+      apply clarsimp
+      apply (erule disjE)
+       apply (erule conjE exE)+
+       apply (erule N_s_state_indep)
+        apply (simp; fail)
+       apply (rule img_fst)
+       apply (rule Set.UnI1)
+       apply (subst mem_Collect_eq)
+       apply solve_ex_triv+
+      apply (erule conjE exE)+
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (rule img_fst)
+      apply (rule Set.UnI2)
+      apply (subst mem_Collect_eq)
+      apply (rule exI)+
+      apply (rule conjI)
+       defer
+      by solve_ex_triv+
+    done
+
+  lemma last_product_state_indep:
+    "(snd o snd o snd o snd) ` fst (equiv.defs.product s)
+   = (snd o snd o snd o snd) ` fst (equiv.defs.product s')"
+    unfolding Product_TA_Defs.product_ta_def Product_TA_Defs.product_trans_def
+    unfolding Product_TA_Defs.product_trans_s_def Product_TA_Defs.product_trans_i_def
+    apply simp
+    unfolding equiv.defs.states'_alt_def
+    apply clarsimp
+    apply (rule equalityI)
+    subgoal
+      apply (rule subsetI)
+      apply clarsimp
+      apply (erule disjE)
+       apply (erule conjE exE)+
+       apply (erule N_s_state_indep)
+        apply (simp; fail)
+       apply (rule image_eqI)
+        prefer 2
+        apply (rule Set.UnI1)
+        apply (subst mem_Collect_eq)
+        apply solve_ex_triv+
+      apply (erule conjE exE)+
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (rule image_eqI)
+       defer
+       apply (rule Set.UnI2)
+       apply (subst mem_Collect_eq)
+       apply (rule exI)+
+       apply (rule conjI)
+        defer
+        apply solve_ex_triv+
+       defer
+       apply (rule HOL.refl)
+      by simp
+    subgoal
+      apply (rule subsetI)
+      apply clarsimp
+      apply (erule disjE)
+      apply (erule conjE exE)+
+      apply (erule N_s_state_indep)
+      apply (simp; fail)
+      apply (rule image_eqI)
+      prefer 2
+      apply (rule Set.UnI1)
+      apply (subst mem_Collect_eq)
+      apply solve_ex_triv+
+      apply (erule conjE exE)+
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (erule N_s_state_indep, (simp; fail))
+      apply (rule image_eqI)
+      defer
+      apply (rule Set.UnI2)
+      apply (subst mem_Collect_eq)
+      apply (rule exI)+
+      apply (rule conjI)
+      defer
+      apply solve_ex_triv+
+      defer
+      apply (rule HOL.refl)
+      by simp
+    done
+
+  lemma state_set_T':
+    "state_set (equiv.defs.T' s'') \<supseteq> fst ` state_set (trans_of A)"
+    unfolding trans_of_def
+    unfolding state_set_def
+    unfolding Prod_TA_Defs.prod_ta_def equiv.defs.prod_trans_def
+    apply simp
+    unfolding equiv.defs.prod_trans_i_def equiv.defs.prod_trans_s_def
+    unfolding trans_of_def
+    apply safe
+       apply (subst fst_product_state_indep; force)
+      apply (subst fst_product_state_indep; force)
+     apply (subst (asm) last_product_state_indep[simplified]; force)
+    by (subst (asm) last_product_state_indep[simplified]; force)
+
+  lemma state_set_T'2:
+    "length L = equiv.defs.p"
+    "\<forall>q<equiv.defs.p. L ! q \<in> State_Networks.state_set (fst (equiv.defs.N ! q))"
+    if "(L, s) \<in> state_set (trans_of A)"
+    using subset_trans [OF state_set_T' equiv.state_set_subs] that by blast+
+
+  lemma state_set_states':
+    "L \<in> equiv.defs.states' s" if "(L, s) \<in> state_set (trans_of A)"
+    using state_set_T'2[OF that] unfolding equiv.defs.states'_alt_def by simp
+
+  lemma state_set_pred:
+    "\<forall>q<p. (equiv.defs.P ! q) (L ! q) s" if "(L, s) \<in> state_set (trans_of A)"
+    using that
+    unfolding Normalized_Zone_Semantics_Impl_Refine.state_set_def
+    unfolding trans_of_def Prod_TA_Defs.prod_ta_def Prod_TA_Defs.prod_trans_def
+    unfolding Prod_TA_Defs.prod_trans_i_def Prod_TA_Defs.prod_trans_s_def
+    by force
+
+end
+
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+  lemma bounded_bounded'':
+    "bounded bounds s \<Longrightarrow> length s = length bounds"
+    unfolding bounded'_def bounded_def by simp
+
+  lemma P_bounded:
+    "(\<forall>q<p. (equiv.defs.P ! q) (L ! q) s) \<Longrightarrow> bounded bounds s"
+    unfolding equiv.state_ta_def equiv.state_pred_def check_pred_def using process_length(3) p_gt_0
+    apply simp
+    unfolding list_all_iff
+    unfolding N_def
+    unfolding runf_def P_def
+    apply (drule spec[of _ 0])
+    by (simp split: option.split; auto split: option.split_asm)
+
+  lemma P_state_length:
+    "(\<forall>q<p. (equiv.defs.P ! q) (L ! q) s) \<Longrightarrow> length s = length bounds"
+    by (intro P_bounded bounded_bounded'')
+
+  lemma state_set_state_length:
+    "length s = length bounds" if "(L, s) \<in> state_set (trans_of A)"
+    using that unfolding state_set_def
+    apply (safe dest!: equiv.defs.prod_ta_cases)
+    unfolding equiv.defs.prod_trans_i_alt_def equiv.defs.prod_trans_s_alt_def
+    by safe (auto dest: P_state_length)
+
+  lemma state_set_states:
+    "state_set (trans_of A) \<subseteq> states'"
+    using state_set_states' state_set_pred unfolding states'_def
+    by (auto intro: P_unfold_1 state_set_state_length)
+
+  lemma p_p_2[simp]:
+  "defs'.defs.p = p"
+  unfolding defs'.p_p p_p ..
+
+  lemma len_product'_N[simp]:
+    "length defs'.defs.N = p"
+    unfolding defs'.defs.p_def[symmetric] by (rule p_p_2)
+
+  lemma len_equiv_N:
+    "length equiv.defs.N = p"
+    unfolding equiv.defs.p_def[symmetric] by simp
+
+  lemma
+    "defs'.p = p"
+    unfolding defs'.p_def by simp
+
+  lemma equiv_p_p: "equiv.p = p"
+    by simp
+
+      (* R *)
+  lemma init_states:
+    "init \<in> equiv.defs.states' s0"
+    using processes_have_trans start_has_trans
+    unfolding equiv_states'_alt_def
+    unfolding init_def N_def T_def by force
+
+  lemma start_pred':
+    "check_pred init s0"
+    using start_pred bounded unfolding check_pred_def runf_def list_all_iff
+    by (fastforce split: option.split intro: bounded_bounded')
+
+  lemma start_states':
+    "(init, s0) \<in> states'"
+    using start_pred' init_states bounded unfolding states'_def bounded_def by auto
+
+  lemma trans_fun_trans_of[intro, simp]:
+    "(trans_fun, trans_of A) \<in> transition_rel states"
+    using trans_fun_trans_of' state_set_states start_states' unfolding transition_rel_def by blast
+
+  definition
+    "inv_fun \<equiv> \<lambda> (L, _). concat (map (\<lambda> i. IArray (map IArray inv) !! i !! (L ! i)) [0..<p])"
+
+  lemma states_states':
+    "states \<subseteq> states'"
+    using state_set_states start_states' by auto
+
+  lemma states'_length[simp]:
+    "length L = p" if "(L, s) \<in> states'"
+    using that  unfolding states'_def by auto
+
+  lemma inv_simp:
+    "I q (L ! q) = inv ! q ! (L ! q)" if "q < p" "(L, s) \<in> states'"
+    unfolding I_def using that states'_states'[OF that(2)] lengths by (auto dest!: states_len)
+
+  lemma inv_fun_inv_of':
+    "(inv_fun, inv_of A) \<in> inv_rel Id states'"
+    unfolding inv_rel_def
+    apply (clarsimp simp: equiv.defs.inv_of_simp Product_TA_Defs.inv_of_product)
+    using process_length(1)
+    unfolding inv_fun_def Product_TA_Defs.product_invariant_def init_def
+    unfolding equiv.defs.N_s_def
+    apply simp
+    apply (rule arg_cong[where f = concat])
+    unfolding inv_of_def Equiv_TA_Defs.state_ta_def apply simp
+    unfolding equiv.state_inv_def N_def Equiv_TA_Defs.state_inv_def
+    by (auto simp: inv_simp)
+
+  lemma inv_rel_mono:
+    "(a, b) \<in> inv_rel Id B" if "(a, b) \<in> inv_rel Id C" "B \<subseteq> C"
+    using that unfolding inv_rel_def b_rel_def fun_rel_def by auto
+
+  lemma inv_fun_inv_of[intro, simp]:
+    "(inv_fun, inv_of A) \<in> inv_rel Id states"
+    using inv_fun_inv_of' states_states' by (rule inv_rel_mono)
+
+  definition "final_fun \<equiv> \<lambda> (L, s). hd_of_formula formula L s"
+
+  lemma final_fun_final':
+    "(final_fun, (\<lambda> (l, s). F l s)) \<in> inv_rel Id states'"
+    unfolding F_def final_fun_def inv_rel_def in_set_member[symmetric] list_ex_iff
+     by (force dest!: states'_states')
+
+  lemma final_fun_final[intro, simp]:
+    "(final_fun, (\<lambda> (l, s). F l s)) \<in> inv_rel Id states"
+    using final_fun_final' states_states' by (rule inv_rel_mono)
+
+  lemma fst_clkp_setD:
+    assumes "(c, d) \<in> clkp_set A l"
+    shows "c > 0" "c \<le> m" "d \<in> range int"
+    using assms clock_set consts_nats clkp_set'_subs
+    unfolding Nats_def clk_set'_def TA_clkp_set_unfold by force+
+
+  lemma init_has_trans:
+    "(init, s0) \<in> fst ` (trans_of A) \<longleftrightarrow> trans_fun (init, s0) \<noteq> []"
+    apply standard
+    using trans_fun_trans_of unfolding transition_rel_def apply force
+    using trans_fun_trans_of' start_states' unfolding transition_rel_def by fast
+
+end (* End of context *)
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+  abbreviation "k_i \<equiv> IArray (map (IArray o (map (IArray o map int))) k)"
+
+  definition
+    "k_impl \<equiv> \<lambda> (l, _). IArray (map (\<lambda> c. Max {k_i !! i !! (l ! i) !! c | i. i < p}) [0..<m+1])"
+
+  lemma k_impl_alt_def:
+    "k_impl =
+    (\<lambda> (l, _). IArray (map (\<lambda> c. Max ((\<lambda> i. k_i !! i !! (l ! i) !! c) ` {0..<p})) [0..<m+1]))"
+  proof -
+    have "{i. i < p} = {0..<p}"
+      by auto
+    then show ?thesis unfolding k_impl_def setcompr_eq_image by auto
+  qed
+
+  lemma k_length_alt:
+    "\<forall> i < p. \<forall> j < length (k ! i). length (k ! i ! j) = m + 1"
+    using k_length(1,3) by (auto dest: nth_mem)
+
+  lemma Max_int_commute:
+    "int (Max S) = Max (int ` S)" if "finite S" "S \<noteq> {}"
+    apply (rule mono_Max_commute)
+      apply standard
+    using that by auto
+
+  lemma k_impl_k'[intro]:
+    "k_impl (l, s) = IArray (k' (l, s))" if "(l, s) \<in> states'"
+  proof -
+    have l_len[simp]: "l ! i < length (trans ! i)" if "i < p" for i
+      using \<open>i < p\<close> \<open>(l, s) \<in> _\<close> by auto
+    have *: "k_i !! i !! (l ! i) !! c = k ! i ! (l ! i) ! c"
+      if "c \<le> m" "i < p" for c i
+    proof -
+      from k_length_alt that k_length(1,2) have "length (k ! i ! (l ! i)) = m + 1"
+        by auto
+      with that k_length process_length(2) processes_have_trans start_has_trans show ?thesis
+        unfolding init_def by auto
+    qed
+    show ?thesis
+      unfolding k_impl_def k'_def k_fun_def
+
+      apply clarsimp
+      apply safe
+      subgoal
+        apply (subst Max_int_commute)
+        subgoal
+          by auto
+        subgoal
+          using p_gt_0 by auto
+        apply (rule arg_cong[where f = Max])
+        apply safe
+        using * apply (solves auto)
+        by (solves \<open>auto simp add: *[symmetric]\<close>)
+
+      subgoal
+        apply (rule Max_eqI)
+          apply (solves auto)
+        using k_length_alt k_length processes_have_trans k_0 p_gt_0 unfolding init_def
+         apply (solves auto)
+
+        using k_length_alt k_length processes_have_trans k_0 p_gt_0 unfolding init_def
+        apply clarsimp
+        apply (rule exI[where x = 0])
+        by simp
+
+      subgoal
+        apply (subst Max_int_commute)
+        subgoal
+          by auto
+        subgoal
+          using p_gt_0 by auto
+        apply (rule arg_cong[where f = Max])
+        apply safe
+        using * apply (solves auto)
+        by (solves \<open>auto simp add: *[symmetric]\<close>)
+      done
+  qed
+
+  lemma k_impl_k'2[intro]:
+    "k_impl (l, s) = IArray (k' (l, s))" if
+    "(l, s) \<in> Normalized_Zone_Semantics_Impl_Refine.state_set (trans_of A)"
+    using that states_states' by auto
+
+  lemma k_impl_k'_0[intro]:
+    "k_impl (init, s0) = IArray (k' (init, s0))"
+    using states_states' by auto
+
+  sublocale impl:
+    Reachability_Problem_Impl
+    where trans_fun = trans_fun
+    and trans_impl = trans_fun
+    and inv_fun = inv_fun
+    and F_fun = final_fun
+    and ceiling = k_impl
+    and A = A
+    and l0 = "(init, s0)"
+    and l0i = "(init, s0)"
+    and F = "PR_CONST ((\<lambda> (l, s). F l s))"
+    and n = m
+    and k = k_fun
+    and loc_rel = Id
+    and show_clock = "show"
+    and show_state = "show"
+    and states' = states
+    unfolding PR_CONST_def
+    apply standard
+           apply (fastforce simp: inv_rel_def b_rel_def)
+    subgoal
+      by auto (metis IdI list_rel_id_simp relAPP_def)
+         by (fastforce simp: inv_rel_def b_rel_def)+
+
+  lemma F_reachable_correct':
+    "impl.op.F_reachable
+    \<longleftrightarrow> (\<exists> L' s' u u'.
+        conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>
+        \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+      )" if "formula = formula.EX \<phi>"
+    using that E_op''.E_from_op_reachability_check[of F_rel "PR_CONST (\<lambda>(x, y). F x y)",
+        unfolded F_rel_def, OF HOL.refl]
+      reachability_check
+    unfolding impl.E_op_F_reachable E_op''.F_reachable_def E_op''.reachable_def
+    unfolding F_rel_def unfolding F_def by force
+
+  lemma PT_PT:
+    "defs'.PT = equiv.PT"
+    apply simp
+    unfolding striptp_def
+    apply (rule ext)
+    apply clarsimp
+    subgoal for x
+      apply (cases "PROG x")
+       apply (simp; fail)
+      subgoal for a
+        by (cases a) auto
+      done
+    done
+
+  lemma P_P[simp]:
+    "defs'.defs.P = equiv.defs.P"
+    unfolding Equiv_TA_Defs.state_ta_def
+    unfolding Equiv_TA_Defs.p_def
+    unfolding Equiv_TA_Defs.state_pred_def
+    using PF_PF by (auto split: option.split)
+
+  lemma map_map_filter:
+    "map f (List.map_filter g xs) = List.map_filter (map_option f o g) xs"
+    by (induction xs; simp add: List.map_filter_simps split: option.split)
+
+  lemma make_g_conv:
+    "defs'.make_g = conv_cc oo equiv.make_g"
+    unfolding Equiv_TA_Defs.make_g_def PT_PT PF_PF apply simp
+    apply (intro ext)
+    apply (clarsimp split: option.splits simp: map_map_filter)
+    apply (rule arg_cong2[where f = List.map_filter])
+    by (auto split: option.split instrc.split)
+
+  lemma make_c_conv:
+    "defs'.make_c = equiv.make_c"
+    unfolding Equiv_TA_Defs.make_c_def PT_PT by simp
+
+  lemma make_f_conv:
+    "defs'.make_f = equiv.make_f"
+    unfolding Equiv_TA_Defs.make_f_def PF_PF by simp
+
+  lemma make_mf_conv:
+    "defs'.make_mf = equiv.make_mf"
+    unfolding Equiv_TA_Defs.make_mf_def PF_PF by simp
+
+  lemmas make_convs = make_g_conv make_c_conv make_f_conv make_mf_conv
+
+  lemma state_trans_conv:
+    "Equiv_TA_Defs.state_trans_t (conv N) max_steps q
+    = (\<lambda> (a, b, c). (a, \<lambda> x. conv_cc (b x), c)) ` equiv.state_trans q" if \<open>q < p\<close>
+    unfolding Equiv_TA_Defs.state_trans_t_def image_Collect
+    using \<open>q < _\<close> make_convs by (force split: prod.splits)+
+
+  lemma map_conv_t:
+    "map trans_of (defs'.defs.N_s s) ! q = conv_t ` (map trans_of (equiv.defs.N_s s) ! q)"
+    if \<open>q < p\<close>
+    using \<open>q < p\<close>
+    apply (subst nth_map)
+    unfolding defs'.defs.N_s_length p_p_2
+     apply assumption
+    apply (subst nth_map)
+    unfolding equiv.defs.N_s_length
+     apply simp
+    unfolding trans_of_def Prod_TA_Defs.N_s_def
+    unfolding len_equiv_N len_product'_N
+    apply simp
+    unfolding Prod_TA_Defs.T_s_def
+    unfolding image_Collect
+    unfolding Equiv_TA_Defs.state_ta_def Equiv_TA_Defs.p_def
+    apply simp
+    using state_trans_conv[of q]
+    apply simp
+    apply safe
+     apply force
+    apply solve_ex_triv+
+    unfolding image_iff by force (* Slow *)
+
+  lemma product_trans_t_conv:
+    "Product_TA_Defs.product_trans_s (defs'.defs.N_s s)
+     = conv_t ` Product_TA_Defs.product_trans_s (equiv.defs.N_s s)"
+    unfolding Product_TA_Defs.product_trans_s_def
+    apply (simp only: states'_conv)
+      apply safe
+     apply (simp only: equiv.states'_len_simp equiv_p_p map_conv_t)
+      unfolding image_Collect
+      apply (simp split: prod.split)
+       apply safe
+      subgoal
+        by defer_ex solve_ex_triv+
+      subgoal
+        by solve_ex_triv+ (force simp only: equiv.states'_len_simp equiv_p_p map_conv_t)
+      done
+
+  lemma product_trans_t_conv':
+    "Product_TA_Defs.product_trans_i (defs'.defs.N_s s)
+     = conv_t ` Product_TA_Defs.product_trans_i (equiv.defs.N_s s)"
+    unfolding Product_TA_Defs.product_trans_i_def
+    apply (simp only: states'_conv)
+      apply safe
+     apply (simp only: equiv.states'_len_simp equiv_p_p map_conv_t)
+      unfolding image_Collect
+      apply (simp split: prod.split)
+       apply safe
+      subgoal
+        by defer_ex solve_ex_triv+
+      subgoal
+        by solve_ex_triv+ (force simp only: equiv.states'_len_simp equiv_p_p map_conv_t)
+      done
+
+  lemma prod_trans_s_conv:
+    "defs'.defs.prod_trans_s = conv_t ` equiv.defs.prod_trans_s"
+    unfolding defs'.defs.prod_trans_s_alt_def
+    unfolding equiv.defs.prod_trans_s_alt_def
+    unfolding product_trans_t_conv
+    unfolding p_p_2 P_P
+    apply simp
+    apply safe
+    unfolding p_p P_P
+     apply (simp add: image_Collect)
+     apply solve_ex_triv
+    subgoal
+      apply defer_ex
+      apply defer_ex
+      by solve_ex_triv+
+    subgoal
+      apply defer_ex
+      apply defer_ex
+      by solve_ex_triv+
+    done
+
+  lemma prod_trans_i_conv:
+    "defs'.defs.prod_trans_i = conv_t ` equiv.defs.prod_trans_i"
+    unfolding defs'.defs.prod_trans_i_alt_def
+    unfolding equiv.defs.prod_trans_i_alt_def
+    unfolding product_trans_t_conv'
+    unfolding p_p_2 P_P
+    apply simp
+    apply safe
+    unfolding p_p P_P
+     apply (simp add: image_Collect)
+     apply solve_ex_triv
+    subgoal
+      apply defer_ex
+      apply defer_ex
+      by solve_ex_triv+
+    subgoal
+      apply defer_ex
+      apply defer_ex
+      by solve_ex_triv+
+    done
+
+  lemma prod_trans_conv:
+    "defs'.defs.prod_trans = conv_t ` equiv.defs.prod_trans"
+    unfolding defs'.defs.prod_trans_def
+    unfolding equiv.defs.prod_trans_def
+    unfolding prod_trans_s_conv
+    unfolding prod_trans_i_conv image_Un ..
+
+  lemma prod_invariant_conv:
+    "defs'.defs.prod_invariant = (map conv_ac \<circ>\<circ> Prod_TA_Defs.prod_invariant) EA"
+    apply (rule ext)
+    apply safe
+    unfolding defs'.defs.prod_invariant_def equiv.defs.prod_invariant_def
+    unfolding Product_TA_Defs.product_ta_def inv_of_def
+    apply simp
+    unfolding Product_TA_Defs.product_invariant_def List.map_concat
+    apply (simp add: Prod_TA_Defs.N_s_length)
+    unfolding Equiv_TA_Defs.p_p Equiv_TA_Defs.p_def apply simp
+    apply (rule cong[where f = concat])
+     apply (rule HOL.refl)
+    unfolding Prod_TA_Defs.N_s_def inv_of_def Equiv_TA_Defs.state_ta_def
+    unfolding Equiv_TA_Defs.p_def unfolding Equiv_TA_Defs.state_inv_def
+    by (simp split: prod.split)
+
+  lemma prod_conv: "defs'.defs.prod_ta = conv_A A"
+    unfolding defs'.defs.prod_ta_def
+    unfolding equiv.defs.prod_ta_def
+    unfolding conv_A_def
+    by (simp add: prod_invariant_conv[symmetric] prod_trans_conv[symmetric])
+
+  lemma F_reachable_correct:
+    "impl.op.F_reachable
+    \<longleftrightarrow> (\<exists> L' s' u u'.
+        conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+        \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+      )" if "formula = formula.EX \<phi>" "start_inv_check"
+      unfolding F_reachable_correct'[OF that(1)]
+      apply (subst product'.prod_correct[symmetric])
+      using prod_conv p_p p_gt_0 apply simp
+      using prod_conv p_p p_gt_0 apply simp
+      using F_reachable_equiv[OF that(2)]
+      by (simp add: F_def, simp add: that(1))
+
+  definition
+    "reachability_checker_old \<equiv>
+      worklist_algo2_impl
+        impl.subsumes_impl impl.a0_impl impl.F_impl impl.succs_impl impl.emptiness_check_impl"
+
+  definition
+    "reachability_checker' \<equiv>
+       pw_impl
+        (return o fst) impl.state_copy_impl impl.tracei impl.subsumes_impl impl.a0_impl impl.F_impl
+        impl.succs_impl impl.emptiness_check_impl"
+
+  theorem reachability_check':
+    "(uncurry0 reachability_checker',
+      uncurry0 (
+        Refine_Basic.RETURN (\<exists> L' s' u u'.
+        conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>
+        \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+       )
+      )
+     )
+    \<in> unit_assnk \<rightarrow>a bool_assn" if "formula = formula.EX \<phi>"
+    using impl.pw_impl_hnr_F_reachable
+    unfolding reachability_checker'_def F_reachable_correct'[OF that] .
+
+  corollary reachability_checker'_hoare:
+    "<emp> reachability_checker'
+    <\<lambda> r. \<up>(r = (\<exists> L' s' u u'.
+        conv_A A \<turnstile>' \<langle>(init, s0), u\<rangle> \<rightarrow>* \<langle>(L', s'), u'\<rangle>
+        \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+       ))
+    >t" if "formula = formula.EX \<phi>"
+   apply (rule cons_post_rule)
+   using reachability_check'[OF that, to_hnr] apply (simp add: hn_refine_def)
+   by (sep_auto simp: pure_def)
+
+  definition reachability_checker where
+    "reachability_checker \<equiv> do
+      {
+        init_sat \<leftarrow> impl.start_inv_check_impl;
+        if init_sat then do
+          { x \<leftarrow> reachability_checker';
+            return (if x then REACHABLE else UNREACHABLE)
+          }
+        else
+          return INIT_INV_ERR
+      }"
+
+  theorem reachability_check:
+    "(uncurry0 reachability_checker,
+      uncurry0 (
+        Refine_Basic.RETURN (
+          if start_inv_check
+          then
+            if
+              (
+              \<exists> L' s' u u'.
+                conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+              \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+              )
+            then REACHABLE
+            else UNREACHABLE
+          else INIT_INV_ERR
+      )
+     ))
+    \<in> unit_assnk \<rightarrow>a id_assn" if "formula = formula.EX \<phi>"
+    apply (simp only: F_reachable_correct[OF that, symmetric] cong: if_cong)
+    supply
+      impl.pw_impl_hnr_F_reachable
+      [unfolded reachability_checker'_def[symmetric], to_hnr, unfolded hn_refine_def,
+       rule_format, sep_heap_rules]
+    supply
+      impl.start_inv_check_impl.refine[to_hnr, unfolded hn_refine_def, rule_format, sep_heap_rules]
+    unfolding reachability_checker_def
+    by sepref_to_hoare (sep_auto simp: pure_def)
+
+  corollary reachability_checker_hoare:
+    "<emp> reachability_checker
+    <\<lambda> r. \<up>(r =
+        (
+          if start_inv_check
+          then
+            if
+              (
+              \<exists> L' s' u u'.
+                conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+              \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+              )
+            then REACHABLE
+            else UNREACHABLE
+          else INIT_INV_ERR
+      )
+       )
+    >t" if "formula = formula.EX \<phi>"
+   apply (rule cons_post_rule)
+   using reachability_check[OF that, to_hnr] apply (simp add: hn_refine_def)
+   by (sep_auto simp: pure_def)
+
+  lemma list_all_concat:
+    "list_all Q (concat xxs) \<longleftrightarrow> (\<forall> xs \<in> set xxs. list_all Q xs)"
+    unfolding list_all_iff by auto
+
+  lemma inv_of_init_unfold:
+    "u \<turnstile> inv_of (conv_A A) (init, s0) \<longleftrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (inv ! i ! 0))"
+  proof -
+    have *: "inv_of (conv_A A) (init, s0) = conv_cc (equiv.defs.I' s0 init)"
+      using equiv.defs.inv_of_simp[of init s0]
+      unfolding inv_of_def conv_A_def by (auto split: prod.split)
+    have "u \<turnstile> inv_of (conv_A A) (init, s0) \<longleftrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (I i 0))"
+      unfolding * Product_TA_Defs.inv_of_product Product_TA_Defs.product_invariant_def
+      apply (simp only: product'.prod.length_L p_p_2 cong: list.map_cong_simp)
+      unfolding equiv.defs.N_s_def length_N
+      apply (simp cong: list.map_cong_simp)
+      unfolding inv_of_def
+      apply (simp cong: list.map_cong_simp)
+      unfolding init_def
+      apply (simp cong: list.map_cong_simp)
+      unfolding Equiv_TA_Defs.state_ta_def
+      apply (simp cong: list.map_cong_simp)
+      unfolding equiv.state_inv_def
+      unfolding N_def
+      by (force simp: map_concat list_all_concat clock_val_def cong: list.map_cong_simp)
+    also have "(\<forall> i < p. u \<turnstile> conv_cc (I i 0)) \<longleftrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (inv ! i ! 0))"
+      unfolding I_def using lengths processes_have_trans by fastforce
+    finally show ?thesis .
+  qed
+
+  corollary reachability_checker_hoare':
+    "<emp> reachability_checker
+    <\<lambda> r. \<up>(r =
+        (
+          if (\<forall>u. (\<forall>c\<in>{1..m}. u c = 0) \<longrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (inv ! i ! 0)))
+          then
+            if
+              (
+              \<exists> L' s' u u'.
+                conv N \<turnstile>max_steps \<langle>init, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+              \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp \<phi> L' s'
+              )
+            then REACHABLE
+            else UNREACHABLE
+          else INIT_INV_ERR
+      )
+       )
+    >t" if "formula = formula.EX \<phi>"
+    using reachability_checker_hoare[OF that] unfolding start_inv_check_correct inv_of_init_unfold .
+
+  subsubsection \<open>Post-processing\<close>
+
+  schematic_goal succs_impl_alt_def:
+    "impl.succs_impl \<equiv> ?impl"
+    unfolding impl.succs_impl_def
+    unfolding k_impl_alt_def
+    apply (abstract_let
+          "\<lambda> (l, _ :: int list). IArray (map (\<lambda> c. MAX i\<in>{0..<p}. k_i !! i !! (l ! i) !! c) [0..<m+1])"
+          k_i
+        )
+    apply (abstract_let "inv_fun :: nat list \<times> int list \<Rightarrow> (nat, int) acconstraint list" inv_fun)
+    apply (abstract_let "trans_fun" trans_fun)
+    unfolding inv_fun_def[abs_def] trans_fun_def[abs_def] trans_s_fun_def trans_i_fun_def trans_i_from_def
+    apply (abstract_let "IArray (map IArray inv)" inv)
+    apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+    apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+    apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+    by (rule Pure.reflexive)
+
+  lemma reachability_checker'_alt_def':
+    "reachability_checker' \<equiv>
+      let
+        key = return \<circ> fst;
+        sub = impl.subsumes_impl;
+        copy = impl.state_copy_impl;
+        start = impl.a0_impl;
+        final = impl.F_impl;
+        succs = impl.succs_impl;
+        empty = impl.emptiness_check_impl;
+        trace = impl.tracei
+      in pw_impl key copy trace sub start final succs empty"
+    unfolding reachability_checker'_def by simp
+
+  schematic_goal reachability_checker_alt_def:
+    "reachability_checker \<equiv> ?impl"
+    unfolding reachability_checker_def
+    unfolding reachability_checker'_alt_def' impl.succs_impl_def
+    unfolding impl.E_op''_impl_def impl.abstr_repair_impl_def impl.abstra_repair_impl_def
+    unfolding impl.start_inv_check_impl_def unbounded_dbm_impl_def unbounded_dbm'_def
+    unfolding k_impl_alt_def
+   apply (abstract_let k_i k_i)
+   apply (abstract_let "inv_fun :: nat list \<times> int list \<Rightarrow> (nat, int) acconstraint list" )
+    apply (abstract_let "trans_fun" trans_fun)
+   unfolding impl.init_dbm_impl_def impl.a0_impl_def
+   unfolding impl.F_impl_def
+   unfolding final_fun_def[abs_def]
+   unfolding impl.subsumes_impl_def
+   unfolding impl.emptiness_check_impl_def
+   unfolding impl.state_copy_impl_def
+  by (rule Pure.reflexive)
+
+end (* End of locale *)
+
+lemmas [code] = UPPAAL_Reachability_Problem_precompiled'.k_impl_def
+
+paragraph \<open>Some post refinements\<close>
+code_thms "fw_upd'"
+code_thms fw_impl'
+code_thms fw_impl
+
+abbreviation plus_int :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "plus_int a b \<equiv> a + b"
+
+fun dbm_add_int :: "int DBMEntry \<Rightarrow> int DBMEntry \<Rightarrow> int DBMEntry"
+where
+  "dbm_add_int \<infinity>     _      = \<infinity>" |
+  "dbm_add_int _      \<infinity>     = \<infinity>" |
+  "dbm_add_int (Le a) (Le b) = (Le (plus_int a b))" |
+  "dbm_add_int (Le a) (Lt b) = (Lt (plus_int a b))" |
+  "dbm_add_int (Lt a) (Le b) = (Lt (plus_int a b))" |
+  "dbm_add_int (Lt a) (Lt b) = (Lt (plus_int a b))"
+
+lemma dbm_add_int:
+  "dbm_add = dbm_add_int"
+  apply (rule ext)+
+  subgoal for x y
+    by (cases x; cases y) auto
+  done
+
+definition
+  "fw_upd'_int m k i j =
+    Refine_Basic.RETURN
+     (op_mtx_set m (i, j)
+       (min (op_mtx_get m (i, j)) (dbm_add_int (op_mtx_get m (i, k)) (op_mtx_get m (k, j)))))"
+
+definition
+  "fw_upd_impl_int n \<equiv> \<lambda>ai bib bia bi. do {
+                      xa \<leftarrow> mtx_get (Suc n) ai (bia, bib);
+                      xb \<leftarrow> mtx_get (Suc n) ai (bib, bi);
+                      x \<leftarrow> mtx_get (Suc n) ai (bia, bi);
+                      let e = (dbm_add_int xa xb);
+                      if e < x then mtx_set (Suc n) ai (bia, bi) e else Heap_Monad.return ai
+                    }"
+
+lemma fw_upd_impl_int_eq:
+  "fw_upd_impl_int = fw_upd_impl"
+  unfolding fw_upd_impl_int_def fw_upd_impl_def
+  unfolding dbm_add_int add
+  unfolding Let_def ..
+
+definition
+  "fw_impl_int n \<equiv>
+    imp_for' 0 (n + 1)
+     (\<lambda>xb. imp_for' 0 (n + 1)
+            (\<lambda>xd. imp_for' 0 (n + 1) (\<lambda>xf \<sigma>'''''. fw_upd_impl_int n \<sigma>''''' xb xd xf)))"
+
+lemma fw_impl'_int:
+  "fw_impl = fw_impl_int"
+  unfolding fw_impl_def fw_impl_int_def
+  unfolding fw_upd_impl_int_eq ..
+
+context UPPAAL_Reachability_Problem_precompiled_defs'
+begin
+
+  definition "run_impl program pc s \<equiv> exec program max_steps (pc, [], s, True, []) []"
+
+  lemma runf_impl:
+    "runf = run_impl PF"
+    unfolding runf_def run_impl_def ..
+
+  lemma runt_impl:
+    "runt = run_impl PT"
+    unfolding runt_def run_impl_def ..
+
+  definition
+    "make_cconstr_impl program pcs =
+    List.map_filter
+     (\<lambda>pc. case program pc of None \<Rightarrow> None | Some (INSTR x) \<Rightarrow> Map.empty x
+           | Some (CEXP ac) \<Rightarrow> Some ac)
+     pcs"
+
+  lemma make_cconstr_impl:
+    "make_cconstr = make_cconstr_impl PROG"
+    unfolding make_cconstr_def make_cconstr_impl_def ..
+
+  definition
+    "check_g_impl programf program pc s \<equiv>
+    case run_impl programf pc s of None \<Rightarrow> None
+    | Some ((x, xa, xb, True, xc), pcs) \<Rightarrow> Some (make_cconstr_impl program pcs)
+    | Some ((x, xa, xb, False, xc), pcs) \<Rightarrow> None"
+
+  lemma check_g_impl:
+    "check_g = check_g_impl PT PROG'"
+    unfolding check_g_impl_def check_g_def runt_impl PROG'_PROG make_cconstr_impl ..
+
+  (*
+  definition
+    "trans_i_from \<equiv> \<lambda> (L, s) i.
+      List.map_filter (\<lambda> (g, a, m, l').
+        case check_g g s of
+          Some cc \<Rightarrow>
+          case runf m s of
+            Some ((_, _, s', _, r), _) \<Rightarrow>
+              if check_pred (L[i := l']) s'
+              then Some (cc, a, r, (L[i := l'], s'))
+              else None
+         | _ \<Rightarrow> None
+      | _ \<Rightarrow> None)
+        ((IArray (map IArray trans_i_map)) !! i !! (L ! i))"
+
+  definition
+    "trans_i_fun L \<equiv> concat (map (trans_i_from L) [0..<p])"
+  *)
+
+  definition
+    "make_reset_impl program m1 s \<equiv>
+      case run_impl program m1 s of
+        Some ((_, _, _, _, r1), _) \<Rightarrow> r1
+      | None \<Rightarrow> []
+    "
+
+  lemma make_reset_impl:
+    "make_reset = make_reset_impl PF"
+    unfolding make_reset_def make_reset_impl_def runf_impl ..
+
+  definition
+    "check_pred_impl program bnds L s \<equiv>
+    list_all
+     (\<lambda>q. case run_impl program (pred ! q ! (L ! q)) s of None \<Rightarrow> False
+          | Some ((x, xa, xb, f, xc), xd) \<Rightarrow>
+              f \<and> (\<forall>i<length s. fst (bnds !! i) < s ! i \<and> s ! i < snd (bnds !! i)))
+     [0..<p]"
+
+  lemma check_pred_impl:
+    "check_pred = check_pred_impl PF (IArray bounds)"
+    unfolding check_pred_def check_pred_impl_def runf_impl bounded'_def ..
+
+  definition
+    "pairs_by_action_impl pf pt porig bnds \<equiv> \<lambda> (L, s) OUT. concat o
+      map (\<lambda> (i, g1, a, m1, l1). List.map_filter
+      (\<lambda> (j, g2, a, m2, l2).
+        if i = j then None else
+        case (check_g_impl pt porig g1 s, check_g_impl pt porig g2 s) of
+          (Some cc1, Some cc2) \<Rightarrow>
+          (case run_impl pf m2 s of
+            Some ((_, _, s1, _, r2), _) \<Rightarrow>
+            (case run_impl pf m1 s1 of
+              Some (( _, _, s', _, _), _) \<Rightarrow>
+                if check_pred_impl pf bnds (L[i := l1, j := l2]) s'
+                then Some (cc1 @ cc2, a, make_reset_impl pf m1 s @ r2, (L[i := l1, j := l2], s'))
+                else None
+            | _ \<Rightarrow> None)
+          | _ \<Rightarrow> None)
+        | _ \<Rightarrow> None
+      )
+      OUT)"
+
+  lemma pairs_by_action_impl:
+    "pairs_by_action = pairs_by_action_impl PF PT PROG' (IArray bounds)"
+    unfolding pairs_by_action_def pairs_by_action_impl_def
+    unfolding check_g_impl make_reset_impl check_pred_impl runf_impl ..
+
+  definition
+    "all_actions_by_state_impl upt_p empty_ran i L \<equiv>
+    fold (\<lambda>ia. actions_by_state ia (i !! ia !! (L ! ia))) upt_p empty_ran"
+
+  lemma all_actions_by_state_impl:
+    "all_actions_by_state = all_actions_by_state_impl [0..<p] (repeat [] na)"
+    unfolding all_actions_by_state_def all_actions_by_state_impl_def ..
+
+  definition
+    "trans_i_from_impl programf programt program bnds trans_i_array \<equiv>
+    \<lambda>(L, s) i.
+       List.map_filter
+        (\<lambda>(g, a, m, l').
+            case check_g_impl programt program g s of None \<Rightarrow> None
+            | Some cc \<Rightarrow>
+                (case run_impl programf m s of None \<Rightarrow> None
+                 | Some ((xx, xa, s', xb, r), xc) \<Rightarrow>
+                    if check_pred_impl programf (IArray bounds) (L[i := l']) s'
+                    then Some (cc, a, r, L[i := l'], s') else None))
+        (trans_i_array !! i !! (L ! i))"
+
+  lemma trans_i_from_impl:
+    "trans_i_from = trans_i_from_impl PF PT PROG' (IArray bounds) (IArray (map IArray trans_i_map))"
+    unfolding trans_i_from_def trans_i_from_impl_def
+    unfolding check_g_impl runf_impl check_pred_impl ..
+
+end
+
+context UPPAAL_Reachability_Problem_precompiled'
+begin
+
+  lemma PF_alt_def:
+    "PF = (\<lambda> pc. if pc < length prog then (IArray (map (map_option stripf) prog)) !! pc else None)"
+    unfolding stripfp_def PROG'_def by auto
+
+  lemma PT_alt_def:
+    "PT = (\<lambda> pc. if pc < length prog then (IArray (map (map_option stript) prog)) !! pc else None)"
+    unfolding striptp_def PROG'_def by auto
+
+  schematic_goal reachability_checker_alt_def_refined:
+    "reachability_checker \<equiv> ?impl"
+    unfolding reachability_checker_alt_def
+    unfolding fw_impl'_int
+    unfolding inv_fun_def trans_fun_def trans_s_fun_def trans_i_fun_def
+    unfolding trans_i_from_impl
+    unfolding runf_impl runt_impl check_g_impl pairs_by_action_impl check_pred_impl
+    apply (abstract_let "IArray (map IArray inv)" inv)
+    apply (abstract_let "IArray (map IArray trans_out_map)" trans_out_map)
+    apply (abstract_let "IArray (map IArray trans_in_map)" trans_in_map)
+    apply (abstract_let "IArray (map IArray trans_i_map)" trans_i_map)
+    apply (abstract_let "IArray bounds" bounds)
+    apply (abstract_let PF PF)
+    apply (abstract_let PT PT)
+    unfolding PF_alt_def PT_alt_def
+    apply (abstract_let PROG' PROG')
+    unfolding PROG'_def
+    apply (abstract_let "length prog" len_prof)
+    apply (abstract_let "IArray (map (map_option stripf) prog)" prog_f)
+    apply (abstract_let "IArray (map (map_option stript) prog)" prog_t)
+    apply (abstract_let "IArray prog" prog)
+    unfolding all_actions_by_state_impl
+    apply (abstract_let "[0..<p]")
+    apply (abstract_let "[0..<na]")
+    apply (abstract_let "{0..<p}")
+    apply (abstract_let "[0..<m+1]")
+    by (rule Pure.reflexive)
+
+end (* End of precompiled' locale context *)
+
+subsection \<open>Check Preconditions\<close>
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  abbreviation
+    "check_nat_subs \<equiv> \<forall> (_, d) \<in> clkp_set'. d \<ge> 0"
+
+  lemma check_nat_subs:
+    "check_nat_subs \<longleftrightarrow> snd ` clkp_set' \<subseteq> \<nat>"
+    unfolding Nats_def apply safe
+    subgoal for _ _ b using rangeI[of int "nat b"] by auto
+    by auto
+
+  definition
+    "check_resets \<equiv> \<forall> x c. Some (INSTR (STOREC c x)) \<in> set prog \<longrightarrow> x = 0"
+
+  lemma check_resets_alt_def:
+    "check_resets =
+      (\<forall> (c, x) \<in> collect_store. x = 0)"
+    unfolding check_resets_def collect_store_def by auto
+
+  definition
+    "check_pre \<equiv>
+      length inv = p \<and> length trans = p \<and> length pred = p
+      \<and> (\<forall> i < p. length (pred ! i) = length (trans ! i) \<and> length (inv ! i) = length (trans ! i))
+      \<and> (\<forall> T \<in> set trans. \<forall> xs \<in> set T. \<forall> (_, _, _, l) \<in> set xs. l < length T)
+      \<and> p > 0 \<and> m > 0
+      \<and> (\<forall> i < p. trans ! i \<noteq> []) \<and> (\<forall> q < p. trans ! q ! 0 \<noteq> [])
+      \<and> check_nat_subs \<and> clk_set' = {1..m}
+      \<and> check_resets
+      "
+
+  lemma finite_clkp_set'[intro, simp]:
+    "finite clkp_set'"
+    unfolding clkp_set'_def
+    using [[simproc add: finite_Collect]]
+    by (auto intro!: finite_vimageI simp: inj_on_def)
+
+  lemma check_pre:
+    "UPPAAL_Reachability_Problem_precompiled p m inv pred trans prog \<longleftrightarrow> check_pre"
+    unfolding
+      UPPAAL_Reachability_Problem_precompiled_def
+      check_pre_def check_nat_subs check_resets_def
+    by auto
+
+end (* End of definitions context for precompiled reachachability problem*)
+
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+context
+  fixes k :: "nat list list list"
+begin
+
+  definition
+    "check_ceiling \<equiv>
+    UPPAAL_Reachability_Problem_precompiled_ceiling_axioms p m max_steps inv trans prog k"
+
+  lemma check_axioms:
+    "UPPAAL_Reachability_Problem_precompiled_ceiling p m max_steps inv pred trans prog k
+    \<longleftrightarrow> check_pre \<and> check_ceiling"
+    unfolding UPPAAL_Reachability_Problem_precompiled_ceiling_def check_pre check_ceiling_def
+    by auto
+
+end
+
+end
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs.collect_cexp_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.collect_store_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.check_resets_alt_def
+
+export_code UPPAAL_Reachability_Problem_precompiled_defs.collect_cexp in SML module_name Test
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs.check_pre
+  UPPAAL_Reachability_Problem_precompiled_defs.check_axioms
+  UPPAAL_Reachability_Problem_precompiled_defs.clkp_set'_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.clk_set'_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.check_pre_def
+  UPPAAL_Reachability_Problem_precompiled_defs.check_ceiling_def
+  UPPAAL_Reachability_Problem_precompiled_defs.init_def
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs'.trans_out_map_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.trans_in_map_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.trans_i_map_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.all_actions_by_state_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.actions_by_state_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.pairs_by_action_def
+
+code_pred clock_val_a .
+
+concrete_definition reachability_checker_impl
+  uses UPPAAL_Reachability_Problem_precompiled'.reachability_checker_alt_def_refined
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs'.make_cconstr_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.make_reset_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.check_pred_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.check_g_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.runf_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.runt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.PROG_def
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs.P_def
+
+lemma exec_code[code]:
+  "exec prog n (pc, st, m, f, rs) pcs =
+  (case n of 0 \<Rightarrow> None
+   | Suc n \<Rightarrow>
+    (case prog pc of None \<Rightarrow> None
+     | Some instr \<Rightarrow>
+         if instr = HALT
+         then Some ((pc, st, m, f, rs), pc # pcs)
+         else
+           (case UPPAAL_Asm.step instr (pc, st, m, f, rs) of
+             None \<Rightarrow> None | Some s \<Rightarrow> exec prog n s (pc # pcs))))"
+  by (cases n) auto
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled'_axioms_def
+  UPPAAL_Reachability_Problem_precompiled'_def
+  pred_act_def
+
+definition
+  "init_pred_check \<equiv> \<lambda> p prog max_steps pred s0.
+    (\<forall> q < p.
+       case (exec
+        (stripfp (UPPAAL_Reachability_Problem_precompiled_defs.PROG prog))
+          max_steps
+          ((pred ! q ! (UPPAAL_Reachability_Problem_precompiled_defs.init p ! q)), [], s0, True, [])
+          [])
+      of Some ((pc, st, s', True, rs), pcs) \<Rightarrow> True | _ \<Rightarrow> False)
+  "
+
+definition
+  "time_indep_check1 \<equiv> \<lambda> pred prog max_steps.
+   (\<forall>x\<in>set pred. \<forall>pc\<in>set x. time_indep_check prog pc max_steps)
+  "
+
+definition
+  "time_indep_check2 \<equiv> \<lambda> trans prog max_steps.
+  (\<forall>T\<in>set trans. \<forall>xs\<in>set T. \<forall>(_, _, pc_u, _)\<in>set xs. time_indep_check prog pc_u max_steps)
+  "
+
+definition
+  "conjunction_check2 \<equiv> \<lambda> trans prog max_steps.
+  (\<forall>T\<in>set trans. \<forall>xs\<in>set T. \<forall>(pc_g, _, _, _)\<in>set xs. conjunction_check prog pc_g max_steps)
+  "
+
+lemma start_pred[code]:
+  "UPPAAL_Reachability_Problem_precompiled_start_state_axioms = (\<lambda> p max_steps trans prog bounds pred s0.
+    init_pred_check p prog max_steps pred s0
+  \<and> bounded bounds s0
+  \<and> time_indep_check1 pred prog max_steps
+  \<and> time_indep_check2 trans prog max_steps
+  \<and> conjunction_check2 trans prog max_steps
+  )"
+  unfolding UPPAAL_Reachability_Problem_precompiled_start_state_axioms_def
+  unfolding init_pred_check_def bounded_def time_indep_check1_def time_indep_check2_def
+    conjunction_check2_def
+  apply (rule ext)+
+  apply safe
+   apply (fastforce split: option.split_asm bool.split_asm)
+  subgoal premises prems
+    using prems(1,7) by (fastforce split: option.split_asm bool.split_asm)
+  done
+
+export_code UPPAAL_Reachability_Problem_precompiled_start_state_axioms
+
+context UPPAAL_Reachability_Problem_precompiled_defs
+begin
+
+  lemma collect_store''_alt_def:
+    "collect_store'' pc \<equiv>
+    case find_resets_start prog pc of
+      None \<Rightarrow> {} |
+      Some pc' \<Rightarrow>
+        \<Union> (
+          (\<lambda> cmd. case cmd of Some (INSTR (STOREC c x)) \<Rightarrow> {(c, x)} | _ \<Rightarrow> {}) `
+            ((!) prog) ` {pc .. pc'}
+        )"
+    unfolding collect_store''_def
+    apply (rule eq_reflection)
+    apply safe
+    subgoal
+      apply (simp del: find_resets_start.simps split: option.split_asm)
+      by (auto intro: sym intro!: bexI split: option.split)
+    subgoal
+      apply (clarsimp simp del: find_resets_start.simps split: option.split_asm)
+      by (auto intro: sym split: instrc.split_asm instr.split_asm)
+    done
+
+  lemma collect_cexp'_alt_def:
+    "collect_cexp' pc \<equiv>
+      \<Union> ((\<lambda> cmd. case cmd of Some (CEXP ac) \<Rightarrow> {ac} | _ \<Rightarrow> {}) `
+          ((!) prog) ` steps_approx max_steps prog pc
+      )"
+    unfolding collect_cexp'_def
+    by (auto 4 3 intro!: eq_reflection bexI intro: sym split: option.splits instrc.split_asm)
+
+end
+
+lemmas [code] =
+  UPPAAL_Reachability_Problem_precompiled_defs'.PROG'_def
+  UPPAAL_Reachability_Problem_precompiled_start_state_def
+  UPPAAL_Reachability_Problem_precompiled_ceiling_axioms_def
+  UPPAAL_Reachability_Problem_precompiled_defs.N_def
+  UPPAAL_Reachability_Problem_precompiled_defs.collect_store''_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs.clkp_set''_def
+  UPPAAL_Reachability_Problem_precompiled_defs.collect_cexp'_alt_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.pairs_by_action_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.make_reset_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.check_g_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.run_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.make_cconstr_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.check_pred_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.all_actions_by_state_impl_def
+  UPPAAL_Reachability_Problem_precompiled_defs'.trans_i_from_impl_def
+
+lemmas [code] =
+  Equiv_TA_Defs.state_ta_def Prod_TA_Defs.N_s_def Product_TA_Defs.states_def
+
+export_code UPPAAL_Reachability_Problem_precompiled'_axioms in SML module_name Test
+
+export_code UPPAAL_Reachability_Problem_precompiled' in SML module_name Test
+
+(* export_code reachability_checker_impl in SML_imp module_name TA *)
+
+hide_const check_and_verify
+
+definition [code]:
+  "check_and_verify p m k max_steps I T prog final bounds P s0 na \<equiv>
+    if UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k
+    then
+      reachability_checker_impl p m max_steps I T prog bounds P s0 na k final
+      \<bind> (\<lambda> x. return (Some x))
+    else return None"
+
+abbreviation "N \<equiv> UPPAAL_Reachability_Problem_precompiled_defs.N"
+
+theorem reachability_check:
+  "(uncurry0 (check_and_verify p m k max_steps I T prog (formula.EX formula) bounds P s0 na),
+    uncurry0 (
+       Refine_Basic.RETURN (
+        if UPPAAL_Reachability_Problem_precompiled' p m max_steps I T prog bounds P s0 na k
+        then Some (
+          if (\<forall>u. (\<forall>c\<in>{1..m}. u c = 0) \<longrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (I ! i ! 0)))
+            then
+              if
+                (
+                \<exists> L' s' u u'.
+                  conv (N p I P T prog bounds) \<turnstile>max_steps
+                  \<langle>repeat 0 p, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+                \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp formula L' s'
+                )
+              then REACHABLE
+              else UNREACHABLE
+            else INIT_INV_ERR
+            )
+        else None
+       )
+    )
+   )
+    \<in> unit_assnk \<rightarrow>a id_assn"
+proof -
+  define A where "A \<equiv> conv (N p I P T prog bounds)"
+  define start_inv where
+    "start_inv \<equiv> (\<forall>u. (\<forall>c\<in>{1..m}. u c = 0) \<longrightarrow> (\<forall> i < p. u \<turnstile> conv_cc (I ! i ! 0)))"
+  define reach where
+    "reach \<equiv>
+      \<exists> L' s' u u'.
+        A \<turnstile>max_steps
+        \<langle>repeat 0 p, s0, u\<rangle> \<rightarrow>* \<langle>L', s', u'\<rangle>
+      \<and> (\<forall> c \<in> {1..m}. u c = 0) \<and> check_bexp formula L' s'"
+  note [sep_heap_rules] =
+    UPPAAL_Reachability_Problem_precompiled'.reachability_checker_hoare'
+    [ OF _ HOL.refl[of "formula.EX formula"],
+      unfolded UPPAAL_Reachability_Problem_precompiled_defs.init_def,
+      of p m max_steps I T prog bounds P s0 na k,
+      unfolded A_def[symmetric] start_inv_def[symmetric] reach_def[symmetric]
+    ]
+  show ?thesis
+    unfolding A_def[symmetric] start_inv_def[symmetric] reach_def[symmetric]
+    unfolding check_and_verify_def
+    by sepref_to_hoare (sep_auto simp: reachability_checker_impl.refine[symmetric])
+qed
+
+export_code open
+  check_and_verify init_pred_check time_indep_check1 time_indep_check1 conjunction_check2
+  checking SML_imp
+
+end (* End of theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/HDDI_02.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/HDDI_02.muntax
@@ -0,0 +1,251 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 2, 
+                    "name": "ring_to_2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "ring_2"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 0, 
+                    "name": "ring_to_1"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "ring_1"
+                }
+            ], 
+            "initial": 0, 
+            "name": "RING", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "tt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 2, 
+                    "label": "rt1?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "tt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 0, 
+                    "label": "rt2?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST1_y<=120", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST1_z<=120", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST1", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST1_y := 0, ST1_x:= 0", 
+                    "target": 1, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z >= 120", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z < 120", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST1_z := 0, ST1_x:= 0", 
+                    "target": 4, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y >= 120", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y < 120", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "ST2_z<=120", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST2_y<=120", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST2_y := 0, ST2_x:= 0", 
+                    "target": 1, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z >= 120", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z < 120", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST2_z := 0, ST2_x:= 0", 
+                    "target": 4, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y >= 120", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y < 120", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }
+            ]
+        }
+    ], 
+    "clocks": "RING_x, ST1_x, ST1_y, ST1_z, ST2_x, ST2_y, ST2_z", 
+    "vars": "",
+    "formula": "E<> ((ST1.station_z_sync || ST1.station_z_async || ST1.station_y_sync || ST1.station_y_async) && (ST2.station_z_sync || ST2.station_z_async || ST2.station_y_sync || ST2.station_y_async) )"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/HDDI_02_broadcast.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/HDDI_02_broadcast.muntax
@@ -0,0 +1,252 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 2, 
+                    "name": "ring_to_2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "ring_2"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 0, 
+                    "name": "ring_to_1"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "ring_1"
+                }
+            ], 
+            "initial": 0, 
+            "name": "RING", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "tt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 2, 
+                    "label": "rt1?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "tt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 0, 
+                    "label": "rt2?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST1_y<=120", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST1_z<=120", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST1", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST1_y := 0, ST1_x:= 0", 
+                    "target": 1, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z >= 120", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z < 120", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST1_z := 0, ST1_x:= 0", 
+                    "target": 4, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y >= 120", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y < 120", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "ST2_z<=120", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST2_y<=120", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST2_y := 0, ST2_x:= 0", 
+                    "target": 1, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z >= 120", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z < 120", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST2_z := 0, ST2_x:= 0", 
+                    "target": 4, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y >= 120", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y < 120", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }
+            ]
+        }
+    ], 
+    "clocks": "RING_x, ST1_x, ST1_y, ST1_z, ST2_x, ST2_y, ST2_z", 
+    "vars": "",
+    "broadcast": ["rt1", "rt2", "tt1", "tt2"],
+    "formula": "E<> ((ST1.station_z_sync || ST1.station_z_async || ST1.station_y_sync || ST1.station_y_async) && 	(ST2.station_z_sync || ST2.station_z_async || ST2.station_y_sync || ST2.station_y_async) )"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/HDDI_08_broadcast.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/HDDI_08_broadcast.muntax
@@ -0,0 +1,960 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 14, 
+                    "name": "ring_to_5"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 15, 
+                    "name": "ring_5"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 12, 
+                    "name": "ring_to_4"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 13, 
+                    "name": "ring_4"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 10, 
+                    "name": "ring_to_3"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 11, 
+                    "name": "ring_3"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 8, 
+                    "name": "ring_to_2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 9, 
+                    "name": "ring_2"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 0, 
+                    "name": "ring_to_1"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "ring_1"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 2, 
+                    "name": "ring_to_6"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "ring_6"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 4, 
+                    "name": "ring_to_7"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 5, 
+                    "name": "ring_7"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 6, 
+                    "name": "ring_to_8"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 7, 
+                    "name": "ring_8"
+                }
+            ], 
+            "initial": 0, 
+            "name": "RING", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "tt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 8, 
+                    "label": "rt1?"
+                }, 
+                {
+                    "source": 8, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 9, 
+                    "label": "tt2!"
+                }, 
+                {
+                    "source": 9, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 10, 
+                    "label": "rt2?"
+                }, 
+                {
+                    "source": 10, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 11, 
+                    "label": "tt3!"
+                }, 
+                {
+                    "source": 11, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 12, 
+                    "label": "rt3?"
+                }, 
+                {
+                    "source": 12, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 13, 
+                    "label": "tt4!"
+                }, 
+                {
+                    "source": 13, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 14, 
+                    "label": "rt4?"
+                }, 
+                {
+                    "source": 14, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 15, 
+                    "label": "tt5!"
+                }, 
+                {
+                    "source": 15, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 2, 
+                    "label": "rt5?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "tt6!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 4, 
+                    "label": "rt6?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": "tt7!"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 6, 
+                    "label": "rt7?"
+                }, 
+                {
+                    "source": 6, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 7, 
+                    "label": "tt8!"
+                }, 
+                {
+                    "source": 7, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 0, 
+                    "label": "rt8?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "ST1_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST1_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST1", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST1_y := 0, ST1_x:= 0", 
+                    "target": 1, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST1_z := 0, ST1_x:= 0", 
+                    "target": 4, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST2_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST2_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST2_y := 0, ST2_x:= 0", 
+                    "target": 1, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST2_z := 0, ST2_x:= 0", 
+                    "target": 4, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST3_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST3_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST3_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST3_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST3", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST3_y := 0, ST3_x:= 0", 
+                    "target": 1, 
+                    "label": "tt3?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST3_x >= 20 && ST3_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST3_x >= 20 && ST3_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST3_z := 0, ST3_x:= 0", 
+                    "target": 4, 
+                    "label": "tt3?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST3_x >= 20 && ST3_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST3_x >= 20 && ST3_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt3!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST4_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST4_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST4_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST4_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST4", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST4_y := 0, ST4_x:= 0", 
+                    "target": 1, 
+                    "label": "tt4?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST4_x >= 20 && ST4_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST4_x >= 20 && ST4_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST4_z := 0, ST4_x:= 0", 
+                    "target": 4, 
+                    "label": "tt4?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST4_x >= 20 && ST4_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST4_x >= 20 && ST4_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt4!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST5_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST5_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST5_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST5_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST5", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST5_y := 0, ST5_x:= 0", 
+                    "target": 1, 
+                    "label": "tt5?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST5_x >= 20 && ST5_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST5_x >= 20 && ST5_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST5_z := 0, ST5_x:= 0", 
+                    "target": 4, 
+                    "label": "tt5?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST5_x >= 20 && ST5_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST5_x >= 20 && ST5_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt5!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST6_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST6_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST6_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST6_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST6", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST6_y := 0, ST6_x:= 0", 
+                    "target": 1, 
+                    "label": "tt6?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST6_x >= 20 && ST6_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST6_x >= 20 && ST6_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST6_z := 0, ST6_x:= 0", 
+                    "target": 4, 
+                    "label": "tt6?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST6_x >= 20 && ST6_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST6_x >= 20 && ST6_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt6!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST7_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST7_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST7_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST7_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST7", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST7_y := 0, ST7_x:= 0", 
+                    "target": 1, 
+                    "label": "tt7?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST7_x >= 20 && ST7_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST7_x >= 20 && ST7_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST7_z := 0, ST7_x:= 0", 
+                    "target": 4, 
+                    "label": "tt7?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST7_x >= 20 && ST7_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST7_x >= 20 && ST7_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt7!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST8_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST8_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST8_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST8_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST8", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST8_y := 0, ST8_x:= 0", 
+                    "target": 1, 
+                    "label": "tt8?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST8_x >= 20 && ST8_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST8_x >= 20 && ST8_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST8_z := 0, ST8_x:= 0", 
+                    "target": 4, 
+                    "label": "tt8?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST8_x >= 20 && ST8_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST8_x >= 20 && ST8_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt8!"
+                }
+            ]
+        }
+    ], 
+    "clocks": "RING_x, ST1_x, ST1_y, ST1_z, ST2_x, ST2_y, ST2_z, ST3_x, ST3_y, ST3_z, ST4_x, ST4_y, ST4_z, ST5_x, ST5_y, ST5_z, ST6_x, ST6_y, ST6_z, ST7_x, ST7_y, ST7_z, ST8_x, ST8_y, ST8_z", 
+    "vars": "",
+    "broadcast": ["rt1", "rt2", "rt3", "rt4", "rt5", "rt6", "rt7", "rt8", "tt1", "tt2", "tt3", "tt4", "tt5", "tt6", "tt7", "tt8"],
+    "formula": "E<> ((ST1.station_z_sync || ST1.station_z_async || ST1.station_y_sync || ST1.station_y_async) && (ST2.station_z_sync || ST2.station_z_async || ST2.station_y_sync || ST2.station_y_async))"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_1.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_1.muntax
@@ -0,0 +1,858 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk<400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PAVI", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "PAVI_t<=150", 
+                    "name": "AVI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "ASed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_t>=850", 
+                    "label": "AP!", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }
+            ], 
+            "initial": 1, 
+            "name": "PLRI", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "PLRI_t<=850", 
+                    "name": "LRI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "ASed"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "A_act!", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "V_act!", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pcond", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Retro"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Ante"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_t, PLRI_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pcond_t, PURI_test_t, Pvv_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI.Idle) && PAVI.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_2.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_2.muntax
@@ -0,0 +1,1134 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk<400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PAVI", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "PAVI_t<=150", 
+                    "name": "AVI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "ASed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_t>=850", 
+                    "label": "AP!", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }
+            ], 
+            "initial": 1, 
+            "name": "PLRI", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "PLRI_t<=850", 
+                    "name": "LRI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "ASed"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "A_act!", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "V_act!", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pcond", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Retro"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Ante"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABLOCK?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VB?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABLOCK?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VB?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABLOCK?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VB?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "VP1", 
+                    "target": 2, 
+                    "target_name": "AS1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "AS1", 
+                    "target": 1, 
+                    "target_name": "VP1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 1, 
+                    "target_name": "VP1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 3, 
+                    "source_name": "VSed", 
+                    "target": 2, 
+                    "target_name": "AS1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 3, 
+            "name": "PELT_det", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "AS1"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "VSed"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "VP1"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pvp_vp_t<=400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }, 
+                {
+                    "guard": "Pvp_vp_t>400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V1", 
+                    "target": 1, 
+                    "target_name": "V2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pvp_vp", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_t, PLRI_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pcond_t, Pvp_vp_t, Pvv_t, PURI_test_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI.Idle) && PAVI.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_3.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_3.muntax
@@ -0,0 +1,1171 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk<400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PAVI", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "PAVI_t<=150", 
+                    "name": "AVI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "ASed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_t>=850", 
+                    "label": "AP!", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }
+            ], 
+            "initial": 1, 
+            "name": "PLRI", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "PLRI_t<=850", 
+                    "name": "LRI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "ASed"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "A_act!", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "V_act!", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pcond", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Retro"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Ante"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "PVP_AS_t>200", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PVP_AS_t<150", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PVP_AS_t>=150 && PVP_AS_t<=200", 
+                    "label": "VP_AS!", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset!", 
+                    "source": 0, 
+                    "source_name": "cancel", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 1, 
+                    "target_name": "AS1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "vp1", 
+                    "update": "PVP_AS_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVP_AS", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "vp1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "AS1"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "cancel"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "E8", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 1, 
+                    "source_name": "E7", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 2, 
+                    "source_name": "E6", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 3, 
+                    "source_name": "E5", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 4, 
+                    "source_name": "E4", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 5, 
+                    "source_name": "E3", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 6, 
+                    "source_name": "E2", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 7, 
+                    "source_name": "E1", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 1, 
+                    "source_name": "E7", 
+                    "target": 0, 
+                    "target_name": "E8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 2, 
+                    "source_name": "E6", 
+                    "target": 1, 
+                    "target_name": "E7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 3, 
+                    "source_name": "E5", 
+                    "target": 2, 
+                    "target_name": "E6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 4, 
+                    "source_name": "E4", 
+                    "target": 3, 
+                    "target_name": "E5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 5, 
+                    "source_name": "E3", 
+                    "target": 4, 
+                    "target_name": "E4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 6, 
+                    "source_name": "E2", 
+                    "target": 5, 
+                    "target_name": "E3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 7, 
+                    "source_name": "E1", 
+                    "target": 6, 
+                    "target_name": "E2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 8, 
+                    "source_name": "Init", 
+                    "target": 7, 
+                    "target_name": "E1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 8, 
+            "name": "PELT_count", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "E7"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "E8"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "E6"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "E5"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "E4"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "E3"
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "", 
+                    "name": "E2"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "", 
+                    "name": "E1"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "", 
+                    "name": "Init"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_t, PLRI_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pcond_t, PVP_AS_t, Pvv_t, PURI_test_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI.Idle) && PAVI.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_4.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_4.muntax
@@ -0,0 +1,1673 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk<=400", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }
+            ], 
+            "initial": 4, 
+            "name": "PAVI_MS", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "VDI_idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "VDI_AVI"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "AVI"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=1000", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "aSensed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=850", 
+                    "label": "AP!", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PLRI_MS", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PLRI_MS_t<=1000", 
+                    "name": "VDI_LRI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "aSensed"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PLRI_MS_t<=850", 
+                    "name": "LRI"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pinterval_t<=400", 
+                    "label": "Fast!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow!", 
+                    "source": 0, 
+                    "source_name": "APed", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 0, 
+                    "target_name": "APed", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "Pinterval_t>400", 
+                    "label": "Slow!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pinterval", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait2nd"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Wait1st"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "APed"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "Resett"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                7, 
+                6, 
+                5, 
+                4, 
+                3, 
+                2, 
+                1, 
+                0, 
+                12, 
+                8
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg!", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 1, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 2, 
+                    "target_name": "switch2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 2, 
+                    "source_name": "switch2", 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 3, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 3, 
+                    "source_name": null, 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 4, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 4, 
+                    "source_name": null, 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 5, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 5, 
+                    "source_name": null, 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 6, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 6, 
+                    "source_name": null, 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 7, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 7, 
+                    "source_name": null, 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 8, 
+                    "source_name": null, 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 8, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD!", 
+                    "source": 12, 
+                    "source_name": "switch1", 
+                    "target": 18, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 12, 
+                    "target_name": "switch1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 18, 
+                    "source_name": "Init", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 18, 
+            "name": "Pcounter", 
+            "nodes": [
+                {
+                    "id": 15, 
+                    "invariant": "", 
+                    "name": "fast3"
+                }, 
+                {
+                    "id": 16, 
+                    "invariant": "", 
+                    "name": "fast2"
+                }, 
+                {
+                    "id": 17, 
+                    "invariant": "", 
+                    "name": "fast1"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "urge<=0", 
+                    "name": "switch2"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 10, 
+                    "invariant": "", 
+                    "name": "fast7"
+                }, 
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "fast8"
+                }, 
+                {
+                    "id": 12, 
+                    "invariant": "urge<=0", 
+                    "name": "switch1"
+                }, 
+                {
+                    "id": 18, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 13, 
+                    "invariant": "", 
+                    "name": "fast5"
+                }, 
+                {
+                    "id": 11, 
+                    "invariant": "", 
+                    "name": "fast6"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 14, 
+                    "invariant": "", 
+                    "name": "fast4"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "du_end!", 
+                    "source": 0, 
+                    "source_name": "V8", 
+                    "target": 9, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg?", 
+                    "source": 9, 
+                    "source_name": "Init", 
+                    "target": 8, 
+                    "target_name": "V0", 
+                    "update": ""
+                }
+            ], 
+            "initial": 9, 
+            "name": "Pdu_count", 
+            "nodes": [
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "", 
+                    "name": "V0"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V6"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "V7"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "V4"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "V5"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "V8"
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "V3"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pvp_vp_t<=400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }, 
+                {
+                    "guard": "Pvp_vp_t>400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V1", 
+                    "target": 1, 
+                    "target_name": "V2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pvp_vp", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_MS_t, PLRI_MS_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pinterval_t, PURI_test_t, Pvv_t, Pvp_vp_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI_MS.Idle) && PAVI_MS.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_5.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_5.muntax
@@ -0,0 +1,1768 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk<=400", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }
+            ], 
+            "initial": 4, 
+            "name": "PAVI_MS", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "VDI_idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "VDI_AVI"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "AVI"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=1000", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "aSensed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=850", 
+                    "label": "AP!", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PLRI_MS", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PLRI_MS_t<=1000", 
+                    "name": "VDI_LRI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "aSensed"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PLRI_MS_t<=850", 
+                    "name": "LRI"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pinterval_t<=400", 
+                    "label": "Fast!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow!", 
+                    "source": 0, 
+                    "source_name": "APed", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 0, 
+                    "target_name": "APed", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "Pinterval_t>400", 
+                    "label": "Slow!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pinterval", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait2nd"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Wait1st"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "APed"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "Resett"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                7, 
+                6, 
+                5, 
+                4, 
+                3, 
+                2, 
+                1, 
+                0, 
+                12, 
+                8
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg!", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 1, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 2, 
+                    "target_name": "switch2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 2, 
+                    "source_name": "switch2", 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 3, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 3, 
+                    "source_name": null, 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 4, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 4, 
+                    "source_name": null, 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 5, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 5, 
+                    "source_name": null, 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 6, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 6, 
+                    "source_name": null, 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 7, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 7, 
+                    "source_name": null, 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 8, 
+                    "source_name": null, 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 8, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD!", 
+                    "source": 12, 
+                    "source_name": "switch1", 
+                    "target": 18, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 12, 
+                    "target_name": "switch1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 18, 
+                    "source_name": "Init", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 18, 
+            "name": "Pcounter", 
+            "nodes": [
+                {
+                    "id": 15, 
+                    "invariant": "", 
+                    "name": "fast3"
+                }, 
+                {
+                    "id": 16, 
+                    "invariant": "", 
+                    "name": "fast2"
+                }, 
+                {
+                    "id": 17, 
+                    "invariant": "", 
+                    "name": "fast1"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "urge<=0", 
+                    "name": "switch2"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 10, 
+                    "invariant": "", 
+                    "name": "fast7"
+                }, 
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "fast8"
+                }, 
+                {
+                    "id": 12, 
+                    "invariant": "urge<=0", 
+                    "name": "switch1"
+                }, 
+                {
+                    "id": 18, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 13, 
+                    "invariant": "", 
+                    "name": "fast5"
+                }, 
+                {
+                    "id": 11, 
+                    "invariant": "", 
+                    "name": "fast6"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 14, 
+                    "invariant": "", 
+                    "name": "fast4"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }
+            ], 
+            "initial": 1, 
+            "name": "PMS", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": ""
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "du_end!", 
+                    "source": 0, 
+                    "source_name": "V8", 
+                    "target": 9, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg?", 
+                    "source": 9, 
+                    "source_name": "Init", 
+                    "target": 8, 
+                    "target_name": "V0", 
+                    "update": ""
+                }
+            ], 
+            "initial": 9, 
+            "name": "Pdu_count", 
+            "nodes": [
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "", 
+                    "name": "V0"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V6"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "V7"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "V4"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "V5"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "V8"
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "V3"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Paa_t<=400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Paa_t=0"
+                }, 
+                {
+                    "guard": "Paa_t>400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V1", 
+                    "target": 1, 
+                    "target_name": "V2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Paa_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Paa", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pvp_vp_t<=400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }, 
+                {
+                    "guard": "Pvp_vp_t>400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "V2", 
+                    "target": 0, 
+                    "target_name": "err", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V1", 
+                    "target": 1, 
+                    "target_name": "V2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "V1", 
+                    "update": "Pvp_vp_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pvp_vp", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "err"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_MS_t, PLRI_MS_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pinterval_t, PURI_test_t, Paa_t, Pvv_t, Pvp_vp_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI_MS.Idle) && PAVI_MS.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_6.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_6.muntax
@@ -0,0 +1,2069 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_AVI", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "VDI_AVI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 1, 
+                    "source_name": "VDI_idle", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "VDI_idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 2, 
+                    "source_name": "WaitURI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk<=400", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_MS_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "AVI", 
+                    "target": 4, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 4, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_MS_t=0"
+                }
+            ], 
+            "initial": 4, 
+            "name": "PAVI_MS", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "VDI_idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "VDI_AVI"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "PAVI_MS_t<=150", 
+                    "name": "AVI"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=1000", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD?", 
+                    "source": 0, 
+                    "source_name": "VDI_LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "VDI_LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "aSensed", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "aSensed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_MS_t>=850", 
+                    "label": "AP!", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "LRI", 
+                    "target": 2, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_MS_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PLRI_MS", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PLRI_MS_t<=1000", 
+                    "name": "VDI_LRI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "aSensed"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PLRI_MS_t<=850", 
+                    "name": "LRI"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "A_act!", 
+                    "source": 0, 
+                    "source_name": "Retro", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "Pcond_t>=150", 
+                    "label": "V_act!", 
+                    "source": 1, 
+                    "source_name": "Ante", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 0, 
+                    "target_name": "Retro", 
+                    "update": "Pcond_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "Ante", 
+                    "update": "Pcond_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pcond", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Retro"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "Pcond_t<=200", 
+                    "name": "Ante"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "Pinterval_t<=400", 
+                    "label": "Fast!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow!", 
+                    "source": 0, 
+                    "source_name": "APed", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 0, 
+                    "target_name": "APed", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "Pinterval_t>400", 
+                    "label": "Slow!", 
+                    "source": 1, 
+                    "source_name": "Resett", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "wait2nd", 
+                    "target": 1, 
+                    "target_name": "Resett", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 3, 
+                    "source_name": "Wait1st", 
+                    "target": 2, 
+                    "target_name": "wait2nd", 
+                    "update": "Pinterval_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Pinterval", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait2nd"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Wait1st"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "APed"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "Resett"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                7, 
+                6, 
+                5, 
+                4, 
+                3, 
+                2, 
+                1, 
+                0, 
+                12, 
+                8
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg!", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 1, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 9, 
+                    "source_name": "fast8", 
+                    "target": 2, 
+                    "target_name": "switch2", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 2, 
+                    "source_name": "switch2", 
+                    "target": 9, 
+                    "target_name": "fast8", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 10, 
+                    "source_name": "fast7", 
+                    "target": 3, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 3, 
+                    "source_name": null, 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 4, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 4, 
+                    "source_name": null, 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 5, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 5, 
+                    "source_name": null, 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 6, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 6, 
+                    "source_name": null, 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 7, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 7, 
+                    "source_name": null, 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VDI!", 
+                    "source": 8, 
+                    "source_name": null, 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_end?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 8, 
+                    "target_name": null, 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 11, 
+                    "source_name": "fast6", 
+                    "target": 10, 
+                    "target_name": "fast7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 11, 
+                    "target_name": "fast6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "DDD!", 
+                    "source": 12, 
+                    "source_name": "switch1", 
+                    "target": 18, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 12, 
+                    "target_name": "switch1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Slow?", 
+                    "source": 13, 
+                    "source_name": "fast5", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 14, 
+                    "source_name": "fast4", 
+                    "target": 13, 
+                    "target_name": "fast5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 15, 
+                    "source_name": "fast3", 
+                    "target": 14, 
+                    "target_name": "fast4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 16, 
+                    "source_name": "fast2", 
+                    "target": 15, 
+                    "target_name": "fast3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 17, 
+                    "source_name": "fast1", 
+                    "target": 16, 
+                    "target_name": "fast2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Fast?", 
+                    "source": 18, 
+                    "source_name": "Init", 
+                    "target": 17, 
+                    "target_name": "fast1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 18, 
+            "name": "Pcounter", 
+            "nodes": [
+                {
+                    "id": 15, 
+                    "invariant": "", 
+                    "name": "fast3"
+                }, 
+                {
+                    "id": 16, 
+                    "invariant": "", 
+                    "name": "fast2"
+                }, 
+                {
+                    "id": 17, 
+                    "invariant": "", 
+                    "name": "fast1"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "urge<=0", 
+                    "name": "switch2"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 10, 
+                    "invariant": "", 
+                    "name": "fast7"
+                }, 
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "fast8"
+                }, 
+                {
+                    "id": 12, 
+                    "invariant": "urge<=0", 
+                    "name": "switch1"
+                }, 
+                {
+                    "id": 18, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 13, 
+                    "invariant": "", 
+                    "name": "fast5"
+                }, 
+                {
+                    "id": 11, 
+                    "invariant": "", 
+                    "name": "fast6"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "urge<=0", 
+                    "name": ""
+                }, 
+                {
+                    "id": 14, 
+                    "invariant": "", 
+                    "name": "fast4"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "du_end!", 
+                    "source": 0, 
+                    "source_name": "V8", 
+                    "target": 9, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "V7", 
+                    "target": 0, 
+                    "target_name": "V8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "V6", 
+                    "target": 1, 
+                    "target_name": "V7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 3, 
+                    "source_name": "V5", 
+                    "target": 2, 
+                    "target_name": "V6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 4, 
+                    "source_name": "V4", 
+                    "target": 3, 
+                    "target_name": "V5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "V3", 
+                    "target": 4, 
+                    "target_name": "V4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 6, 
+                    "source_name": "V2", 
+                    "target": 5, 
+                    "target_name": "V3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 7, 
+                    "source_name": "V1", 
+                    "target": 6, 
+                    "target_name": "V2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 8, 
+                    "source_name": "V0", 
+                    "target": 7, 
+                    "target_name": "V1", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "du_beg?", 
+                    "source": 9, 
+                    "source_name": "Init", 
+                    "target": 8, 
+                    "target_name": "V0", 
+                    "update": ""
+                }
+            ], 
+            "initial": 9, 
+            "name": "Pdu_count", 
+            "nodes": [
+                {
+                    "id": 9, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "", 
+                    "name": "V1"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "", 
+                    "name": "V0"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "V6"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "V7"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "V4"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "V5"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "V8"
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "", 
+                    "name": "V2"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "V3"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "PVP_AS_t>200", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PVP_AS_t<150", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PVP_AS_t>=150 && PVP_AS_t<=200", 
+                    "label": "VP_AS!", 
+                    "source": 1, 
+                    "source_name": "AS1", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset!", 
+                    "source": 0, 
+                    "source_name": "cancel", 
+                    "target": 3, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 1, 
+                    "target_name": "AS1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "vp1", 
+                    "target": 0, 
+                    "target_name": "cancel", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Init", 
+                    "target": 2, 
+                    "target_name": "vp1", 
+                    "update": "PVP_AS_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVP_AS", 
+            "nodes": [
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Init"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "vp1"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "AS1"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "cancel"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "E8", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 1, 
+                    "source_name": "E7", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 2, 
+                    "source_name": "E6", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 3, 
+                    "source_name": "E5", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 4, 
+                    "source_name": "E4", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 5, 
+                    "source_name": "E3", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 6, 
+                    "source_name": "E2", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "reset?", 
+                    "source": 7, 
+                    "source_name": "E1", 
+                    "target": 8, 
+                    "target_name": "Init", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 1, 
+                    "source_name": "E7", 
+                    "target": 0, 
+                    "target_name": "E8", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 2, 
+                    "source_name": "E6", 
+                    "target": 1, 
+                    "target_name": "E7", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 3, 
+                    "source_name": "E5", 
+                    "target": 2, 
+                    "target_name": "E6", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 4, 
+                    "source_name": "E4", 
+                    "target": 3, 
+                    "target_name": "E5", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 5, 
+                    "source_name": "E3", 
+                    "target": 4, 
+                    "target_name": "E4", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 6, 
+                    "source_name": "E2", 
+                    "target": 5, 
+                    "target_name": "E3", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 7, 
+                    "source_name": "E1", 
+                    "target": 6, 
+                    "target_name": "E2", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP_AS?", 
+                    "source": 8, 
+                    "source_name": "Init", 
+                    "target": 7, 
+                    "target_name": "E1", 
+                    "update": ""
+                }
+            ], 
+            "initial": 8, 
+            "name": "PELT_count", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "E7"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "E8"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "E6"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "E5"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "", 
+                    "name": "E4"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "E3"
+                }, 
+                {
+                    "id": 6, 
+                    "invariant": "", 
+                    "name": "E2"
+                }, 
+                {
+                    "id": 7, 
+                    "invariant": "", 
+                    "name": "E1"
+                }, 
+                {
+                    "id": 8, 
+                    "invariant": "", 
+                    "name": "Init"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_MS_t, PLRI_MS_t, PPVARP_t, PVRP_t, PV_x, PA_x, Pcond_t, Pinterval_t, PVP_AS_t, PURI_test_t, Pvv_t, urge", 
+    "vars": "",
+    "formula": "E<> !(PAVI_MS.Idle) && PAVI_MS.Idle"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_all_urgent.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_all_urgent.muntax
@@ -0,0 +1,708 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "clk>=400", 
+                    "label": "VP!", 
+                    "source": 0, 
+                    "source_name": "WaitURI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk<400", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 0, 
+                    "target_name": "WaitURI", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PAVI_t>=150 && clk>=400", 
+                    "label": "VP!", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "AVI", 
+                    "target": 2, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 2, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "AVI", 
+                    "update": "PAVI_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PAVI", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "clk<=400", 
+                    "name": "WaitURI"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "PAVI_t<=150", 
+                    "name": "AVI"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "ASed", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 0, 
+                    "target_name": "ASed", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PLRI_t>=850", 
+                    "label": "AP!", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "LRI", 
+                    "target": 1, 
+                    "target_name": "LRI", 
+                    "update": "PLRI_t=0"
+                }
+            ], 
+            "initial": 1, 
+            "name": "PLRI", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "PLRI_t<=850", 
+                    "name": "LRI"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "ASed"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0, 
+                1, 
+                3
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "ABlock!", 
+                    "source": 0, 
+                    "source_name": "ABl", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 0, 
+                    "target_name": "ABl", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AR!", 
+                    "source": 1, 
+                    "source_name": "inter1", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 1, 
+                    "target_name": "inter1", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PPVARP_t>=100", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "PVARP", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "PPVARP_t>=50", 
+                    "label": "", 
+                    "source": 4, 
+                    "source_name": "PVAB", 
+                    "target": 2, 
+                    "target_name": "PVARP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AS!", 
+                    "source": 3, 
+                    "source_name": "inter", 
+                    "target": 5, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Aget?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 3, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 5, 
+                    "source_name": "Idle", 
+                    "target": 4, 
+                    "target_name": "PVAB", 
+                    "update": "PPVARP_t=0"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "ABl"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter1"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PPVARP_t<=100", 
+                    "name": "PVARP"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 4, 
+                    "invariant": "PPVARP_t<=50", 
+                    "name": "PVAB"
+                }, 
+                {
+                    "id": 5, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1, 
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VBlock!", 
+                    "source": 0, 
+                    "source_name": "VB", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 0, 
+                    "target_name": "VB", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS!", 
+                    "source": 1, 
+                    "source_name": "inter", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "Vget?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 1, 
+                    "target_name": "inter", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "PVRP_t>=150", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "VRP", 
+                    "target": 3, 
+                    "target_name": "Idle", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 3, 
+                    "source_name": "Idle", 
+                    "target": 2, 
+                    "target_name": "VRP", 
+                    "update": "PVRP_t=0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "PVRP", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "urge<=0", 
+                    "name": "inter"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "VB"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "Idle"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "PVRP_t<=150", 
+                    "name": "VRP"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "URI", 
+                    "target": 0, 
+                    "target_name": "URI", 
+                    "update": "clk=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PURI", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "URI"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "interval", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_vp", 
+                    "target": 0, 
+                    "target_name": "interval", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_v", 
+                    "target": 1, 
+                    "target_name": "wait_vp", 
+                    "update": "PURI_test_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "PURI_test", 
+            "nodes": [
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_vp"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_v"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "interval"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                0
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "two_a", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 1, 
+                    "source_name": "wait_2nd", 
+                    "target": 0, 
+                    "target_name": "two_a", 
+                    "update": "urge=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VS?", 
+                    "source": 2, 
+                    "source_name": "wait_1st", 
+                    "target": 1, 
+                    "target_name": "wait_2nd", 
+                    "update": "Pvv_t=0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "urge<=0", 
+                    "name": "two_a"
+                }, 
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "wait_2nd"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "V_act?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "VP?", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }, 
+                {
+                    "guard": "PV_x>500", 
+                    "label": "Vget!", 
+                    "source": 0, 
+                    "source_name": "VReady", 
+                    "target": 0, 
+                    "target_name": "VReady", 
+                    "update": "PV_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PV", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PV_x<800", 
+                    "name": "VReady"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "A_act?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "AP?", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }, 
+                {
+                    "guard": "PA_x>500", 
+                    "label": "Aget!", 
+                    "source": 0, 
+                    "source_name": "AReady", 
+                    "target": 0, 
+                    "target_name": "AReady", 
+                    "update": "PA_x=0"
+                }
+            ], 
+            "initial": 0, 
+            "name": "PA", 
+            "nodes": [
+                {
+                    "id": 0, 
+                    "invariant": "PA_x<800", 
+                    "name": "AReady"
+                }
+            ]
+        }
+    ], 
+    "broadcast": [
+        "AP", 
+        "AS", 
+        "VP", 
+        "VS", 
+        "Vget", 
+        "Aget", 
+        "AR", 
+        "A_act", 
+        "V_act", 
+        "ABlock", 
+        "VBlock", 
+        "reset", 
+        "VP_AS", 
+        "Fast", 
+        "Slow", 
+        "du_beg", 
+        "du_end", 
+        "DDD", 
+        "VDI"
+    ], 
+    "clocks": "clk, PAVI_t, PLRI_t, PPVARP_t, PVRP_t, PURI_test_t, Pvv_t, PV_x, PA_x, urge", 
+    "vars": "",
+    "formula": "E<> (Pvv.two_a)"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/PM_test.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/PM_test.muntax
@@ -0,0 +1,195 @@
+{
+    "broadcast": "AP, AS, VP, VS, Vget, Aget, AR, A_act, V_act, ABlock, VBlock, reset, VP_AS, Fast, Slow, du_beg, du_end, DDD, VDI", 
+    "automata": [ 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "ABl"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "inter1"
+                }, 
+                {
+                    "invariant": "PPVARP_t<=100", 
+                    "id": 2, 
+                    "name": "PVARP"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "inter"
+                }, 
+                {
+                    "invariant": "PPVARP_t<=50", 
+                    "id": 4, 
+                    "name": "PVAB"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 5, 
+                    "name": "Idle"
+                }
+            ], 
+            "initial": 5, 
+            "name": "PPVARP", 
+            "edges": [
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "A_act?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "A_act?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 4, 
+                    "label": "ABlock!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "Aget?"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "A_act?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": "AR!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "Aget?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "PPVARP_t>=100", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "PPVARP_t>=50", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": "AS!"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "Aget?"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "PPVARP_t=0", 
+                    "target": 4, 
+                    "label": "VP?"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "PPVARP_t=0", 
+                    "target": 4, 
+                    "label": "VS?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "two_a"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "wait_1st"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "wait_2nd"
+                }
+            ], 
+            "initial": 2, 
+            "name": "Pvv", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "Pvv_t=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "VP?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "VS?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "Pvv_t=0", 
+                    "target": 1, 
+                    "label": "VP?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "Pvv_t=0", 
+                    "target": 1, 
+                    "label": "VS?"
+                }
+            ]
+        }
+    ], 
+    "clocks": "clk, PAVI_t, PLRI_t, PPVARP_t, PVRP_t, PAVI_MS_t, PLRI_MS_t, Pinterval_t, PVP_AS_t, PV_x, PA_x, Pcond_t, PURI_test_t, Pvv_t, Pvp_vp_t, PELT_det_t", 
+    "vars": "",
+    "formula": "E<> (Pvv.two_a)"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/bridge.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/bridge.muntax
@@ -0,0 +1,343 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "Viking1_y >= 5", 
+                    "label": "release!", 
+                    "source": 2, 
+                    "source_name": null, 
+                    "target": 3, 
+                    "target_name": "unsafe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 1", 
+                    "label": "take !", 
+                    "source": 1, 
+                    "source_name": "safe", 
+                    "target": 2, 
+                    "target_name": null, 
+                    "update": "Viking1_y = 0"
+                }, 
+                {
+                    "guard": "Viking1_y >= 5", 
+                    "label": "release!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 1, 
+                    "target_name": "safe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 0", 
+                    "label": "take !", 
+                    "source": 3, 
+                    "source_name": "unsafe", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "Viking1_y = 0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Viking1", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "unsafe"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "safe"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "Viking2_y >= 10", 
+                    "label": "release!", 
+                    "source": 2, 
+                    "source_name": null, 
+                    "target": 3, 
+                    "target_name": "unsafe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 1", 
+                    "label": "take !", 
+                    "source": 1, 
+                    "source_name": "safe", 
+                    "target": 2, 
+                    "target_name": null, 
+                    "update": "Viking2_y = 0"
+                }, 
+                {
+                    "guard": "Viking2_y >= 10", 
+                    "label": "release!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 1, 
+                    "target_name": "safe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 0", 
+                    "label": "take !", 
+                    "source": 3, 
+                    "source_name": "unsafe", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "Viking2_y = 0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Viking2", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "unsafe"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "safe"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "Viking3_y >= 20", 
+                    "label": "release!", 
+                    "source": 2, 
+                    "source_name": null, 
+                    "target": 3, 
+                    "target_name": "unsafe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 1", 
+                    "label": "take !", 
+                    "source": 1, 
+                    "source_name": "safe", 
+                    "target": 2, 
+                    "target_name": null, 
+                    "update": "Viking3_y = 0"
+                }, 
+                {
+                    "guard": "Viking3_y >= 20", 
+                    "label": "release!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 1, 
+                    "target_name": "safe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 0", 
+                    "label": "take !", 
+                    "source": 3, 
+                    "source_name": "unsafe", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "Viking3_y = 0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Viking3", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "unsafe"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "safe"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "Viking4_y >= 25", 
+                    "label": "release!", 
+                    "source": 2, 
+                    "source_name": null, 
+                    "target": 3, 
+                    "target_name": "unsafe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 1", 
+                    "label": "take !", 
+                    "source": 1, 
+                    "source_name": "safe", 
+                    "target": 2, 
+                    "target_name": null, 
+                    "update": "Viking4_y = 0"
+                }, 
+                {
+                    "guard": "Viking4_y >= 25", 
+                    "label": "release!", 
+                    "source": 0, 
+                    "source_name": null, 
+                    "target": 1, 
+                    "target_name": "safe", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "L == 0", 
+                    "label": "take !", 
+                    "source": 3, 
+                    "source_name": "unsafe", 
+                    "target": 0, 
+                    "target_name": null, 
+                    "update": "Viking4_y = 0"
+                }
+            ], 
+            "initial": 3, 
+            "name": "Viking4", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "unsafe"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": ""
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": "safe"
+                }
+            ]
+        }, 
+        {
+            "committed": [
+                1
+            ], 
+            "edges": [
+                {
+                    "guard": "", 
+                    "label": "take?", 
+                    "source": 2, 
+                    "source_name": "free", 
+                    "target": 1, 
+                    "target_name": null, 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 0, 
+                    "target_name": "one", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "take?", 
+                    "source": 1, 
+                    "source_name": null, 
+                    "target": 3, 
+                    "target_name": "two", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "release?", 
+                    "source": 0, 
+                    "source_name": "one", 
+                    "target": 2, 
+                    "target_name": "free", 
+                    "update": "L = (1 ? L == 0 : 0)"
+                }, 
+                {
+                    "guard": "", 
+                    "label": "release?", 
+                    "source": 3, 
+                    "source_name": "two", 
+                    "target": 0, 
+                    "target_name": "one", 
+                    "update": ""
+                }
+            ], 
+            "initial": 2, 
+            "name": "Torch", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "free"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "two"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "one"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "", 
+                    "name": ""
+                }
+            ]
+        }
+    ], 
+    "clocks": "time, Viking1_y, Viking2_y, Viking3_y, Viking4_y", 
+    "vars": "L[-1:2]",
+    "formula": "E<> Viking1.safe && Viking2.safe && Viking3.safe && Viking4.safe"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/csma_05.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/csma_05.muntax
@@ -0,0 +1,527 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 6, 
+                    "name": "bus_collision5"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 4, 
+                    "name": "bus_collision3"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 5, 
+                    "name": "bus_collision4"
+                }, 
+                {
+                    "invariant": "P0_x < 26", 
+                    "id": 2, 
+                    "name": "bus_collision1"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 3, 
+                    "name": "bus_collision2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "bus_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "bus_active"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P0", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P0_x:= 0", 
+                    "target": 1, 
+                    "label": "begin?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P0_x:= 0", 
+                    "target": 0, 
+                    "label": "end?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P0_x >= 26", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "busy!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P0_x < 26", 
+                    "update": "P0_x:= 0", 
+                    "target": 2, 
+                    "label": "begin?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P0_x < 26", 
+                    "update": "P0_x:= 0", 
+                    "target": 3, 
+                    "label": "cd1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 4, 
+                    "label": "cd2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 5, 
+                    "label": "cd3!"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 6, 
+                    "label": "cd4!"
+                }, 
+                {
+                    "source": 6, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 0, 
+                    "label": "cd5!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "P1_x < 52", 
+                    "id": 0, 
+                    "name": "sender_retry"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P1_x<= 808", 
+                    "id": 2, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "sender_transm_check"
+                }
+            ], 
+            "initial": 1, 
+            "name": "P1", 
+            "edges": [
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 2, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 0, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 0, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P1_x <= 808 && P1_x >= 808", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 0, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P1_x >= 52", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 2, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 0, 
+                    "label": "cd1?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P2_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P2_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 0, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P2_x <= 808 && P2_x >= 808", 
+                    "update": "P2_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P3_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P3_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P3", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 0, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P3_x <= 808 && P3_x >= 808", 
+                    "update": "P3_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "P4_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P4_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P4", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 0, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P4_x <= 808 && P4_x >= 808", 
+                    "update": "P4_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "P5_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P5_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P5", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 0, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P5_x <= 808 && P5_x >= 808", 
+                    "update": "P5_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }
+            ]
+        }
+    ], 
+    "clocks": "P0_x, P1_x, P2_x, P3_x, P4_x, P5_x",
+    "formula": "E<> ( P1.sender_transm_check && P2.sender_transm )",
+    "vars": ""
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/csma_06.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/csma_06.muntax
@@ -0,0 +1,618 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 6, 
+                    "name": "bus_collision5"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 7, 
+                    "name": "bus_collision6"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 4, 
+                    "name": "bus_collision3"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 5, 
+                    "name": "bus_collision4"
+                }, 
+                {
+                    "invariant": "P0_x < 26", 
+                    "id": 2, 
+                    "name": "bus_collision1"
+                }, 
+                {
+                    "invariant": "P0_x <= 0", 
+                    "id": 3, 
+                    "name": "bus_collision2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "bus_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "bus_active"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P0", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P0_x:= 0", 
+                    "target": 1, 
+                    "label": "begin?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "P0_x:= 0", 
+                    "target": 0, 
+                    "label": "end?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P0_x >= 26", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "busy!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P0_x < 26", 
+                    "update": "P0_x:= 0", 
+                    "target": 2, 
+                    "label": "begin?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P0_x < 26", 
+                    "update": "P0_x:= 0", 
+                    "target": 3, 
+                    "label": "cd1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 4, 
+                    "label": "cd2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 5, 
+                    "label": "cd3!"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 6, 
+                    "label": "cd4!"
+                }, 
+                {
+                    "source": 6, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 7, 
+                    "label": "cd5!"
+                }, 
+                {
+                    "source": 7, 
+                    "guard": "P0_x <= 0", 
+                    "update": "P0_x:= 0", 
+                    "target": 0, 
+                    "label": "cd6!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_transm_check"
+                }, 
+                {
+                    "invariant": "P1_x < 52", 
+                    "id": 1, 
+                    "name": "sender_retry"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P1_x<= 808", 
+                    "id": 3, 
+                    "name": "sender_transm"
+                }
+            ], 
+            "initial": 2, 
+            "name": "P1", 
+            "edges": [
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 3, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 2, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "P1_x <= 808 && P1_x >= 808", 
+                    "update": "P1_x:= 0", 
+                    "target": 2, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "cd1?"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "P1_x >= 52", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": ""
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 3, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P1_x < 52", 
+                    "update": "P1_x:= 0", 
+                    "target": 1, 
+                    "label": "cd1?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P2_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P2_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 0, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P2_x <= 808 && P2_x >= 808", 
+                    "update": "P2_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P2_x < 52", 
+                    "update": "P2_x:= 0", 
+                    "target": 2, 
+                    "label": "cd2?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P3_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P3_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P3", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 0, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P3_x <= 808 && P3_x >= 808", 
+                    "update": "P3_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P3_x < 52", 
+                    "update": "P3_x:= 0", 
+                    "target": 2, 
+                    "label": "cd3?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P4_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "P4_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P4", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 0, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P4_x <= 808 && P4_x >= 808", 
+                    "update": "P4_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P4_x < 52", 
+                    "update": "P4_x:= 0", 
+                    "target": 2, 
+                    "label": "cd4?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P5_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }, 
+                {
+                    "invariant": "P5_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P5", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 0, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P5_x <= 808 && P5_x >= 808", 
+                    "update": "P5_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P5_x < 52", 
+                    "update": "P5_x:= 0", 
+                    "target": 2, 
+                    "label": "cd5?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "P6_x<= 808", 
+                    "id": 1, 
+                    "name": "sender_transm"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "sender_wait"
+                }, 
+                {
+                    "invariant": "P6_x < 52", 
+                    "id": 2, 
+                    "name": "sender_retry"
+                }
+            ], 
+            "initial": 0, 
+            "name": "P6", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P6_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P6_x:= 0", 
+                    "target": 0, 
+                    "label": "cd6?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P6_x:= 0", 
+                    "target": 2, 
+                    "label": "cd6?"
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "P6_x:= 0", 
+                    "target": 2, 
+                    "label": "busy?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P6_x <= 808 && P6_x >= 808", 
+                    "update": "P6_x:= 0", 
+                    "target": 0, 
+                    "label": "end!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "P6_x < 52", 
+                    "update": "P6_x:= 0", 
+                    "target": 2, 
+                    "label": "cd6?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P6_x < 52", 
+                    "update": "P6_x:= 0", 
+                    "target": 1, 
+                    "label": "begin!"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "P6_x < 52", 
+                    "update": "P6_x:= 0", 
+                    "target": 2, 
+                    "label": "cd6?"
+                }
+            ]
+        }
+    ], 
+    "clocks": "P0_x, P1_x, P2_x, P3_x, P4_x, P5_x, P6_x",
+    "formula": "E<> ( P1.sender_transm_check && P2.sender_transm )",
+    "vars": ""
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/fischer.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/fischer.muntax
@@ -0,0 +1,303 @@
+{
+    "automata": [
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "A", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P1_x = 0"
+                }, 
+                {
+                    "guard": "P1_x<=2", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "req", 
+                    "target": 0, 
+                    "target_name": "wait", 
+                    "update": "P1_x = 0, id = 1"
+                }, 
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P1_x = 0"
+                }, 
+                {
+                    "guard": "P1_x>2 && id==1", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 3, 
+                    "target_name": "cs", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "cs", 
+                    "target": 2, 
+                    "target_name": "A", 
+                    "update": "id = 0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "P1", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "A"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "cs"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "wait"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "P1_x<=2", 
+                    "name": "req"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "A", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P2_x = 0"
+                }, 
+                {
+                    "guard": "P2_x<=2", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "req", 
+                    "target": 0, 
+                    "target_name": "wait", 
+                    "update": "P2_x = 0, id = 2"
+                }, 
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P2_x = 0"
+                }, 
+                {
+                    "guard": "P2_x>2 && id==2", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 3, 
+                    "target_name": "cs", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "cs", 
+                    "target": 2, 
+                    "target_name": "A", 
+                    "update": "id = 0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "P2", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "A"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "cs"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "wait"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "P2_x<=2", 
+                    "name": "req"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "A", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P3_x = 0"
+                }, 
+                {
+                    "guard": "P3_x<=2", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "req", 
+                    "target": 0, 
+                    "target_name": "wait", 
+                    "update": "P3_x = 0, id = 3"
+                }, 
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P3_x = 0"
+                }, 
+                {
+                    "guard": "P3_x>2 && id==3", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 3, 
+                    "target_name": "cs", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "cs", 
+                    "target": 2, 
+                    "target_name": "A", 
+                    "update": "id = 0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "P3", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "A"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "cs"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "wait"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "P3_x<=2", 
+                    "name": "req"
+                }
+            ]
+        }, 
+        {
+            "committed": [], 
+            "edges": [
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 2, 
+                    "source_name": "A", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P4_x = 0"
+                }, 
+                {
+                    "guard": "P4_x<=2", 
+                    "label": "", 
+                    "source": 1, 
+                    "source_name": "req", 
+                    "target": 0, 
+                    "target_name": "wait", 
+                    "update": "P4_x = 0, id = 4"
+                }, 
+                {
+                    "guard": "id== 0", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 1, 
+                    "target_name": "req", 
+                    "update": "P4_x = 0"
+                }, 
+                {
+                    "guard": "P4_x>2 && id==4", 
+                    "label": "", 
+                    "source": 0, 
+                    "source_name": "wait", 
+                    "target": 3, 
+                    "target_name": "cs", 
+                    "update": ""
+                }, 
+                {
+                    "guard": "", 
+                    "label": "", 
+                    "source": 3, 
+                    "source_name": "cs", 
+                    "target": 2, 
+                    "target_name": "A", 
+                    "update": "id = 0"
+                }
+            ], 
+            "initial": 2, 
+            "name": "P4", 
+            "nodes": [
+                {
+                    "id": 2, 
+                    "invariant": "", 
+                    "name": "A"
+                }, 
+                {
+                    "id": 3, 
+                    "invariant": "", 
+                    "name": "cs"
+                }, 
+                {
+                    "id": 0, 
+                    "invariant": "", 
+                    "name": "wait"
+                }, 
+                {
+                    "id": 1, 
+                    "invariant": "P4_x<=2", 
+                    "name": "req"
+                }
+            ]
+        }
+    ], 
+    "clocks": "P1_x, P2_x, P3_x, P4_x", 
+    "vars": "id[-32767:32767]",
+    "formula": "P1.req --> P1.wait"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/fischer_05.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/fischer_05.muntax
@@ -0,0 +1,322 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "c"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "cs"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "a"
+                }, 
+                {
+                    "invariant": "x1<=2", 
+                    "id": 1, 
+                    "name": "b"
+                }
+            ], 
+            "initial": 0, 
+            "name": "p1", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "id=0", 
+                    "update": "x1:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "x1<=2", 
+                    "update": "x1:=0,id:=1", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=0", 
+                    "update": "x1:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=1&& x1>2", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "x1:=0,id:=0", 
+                    "target": 0, 
+                    "label": ""
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "c"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "cs"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "a"
+                }, 
+                {
+                    "invariant": "x2<=2", 
+                    "id": 1, 
+                    "name": "b"
+                }
+            ], 
+            "initial": 0, 
+            "name": "p2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "id=0", 
+                    "update": "x2:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "x2<=2", 
+                    "update": "x2:=0,id:=2", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=0", 
+                    "update": "x2:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=2&& x2>2", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "x2:=0,id:=0", 
+                    "target": 0, 
+                    "label": ""
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "c"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "cs"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "a"
+                }, 
+                {
+                    "invariant": "x3<=2", 
+                    "id": 3, 
+                    "name": "b"
+                }
+            ], 
+            "initial": 2, 
+            "name": "p3", 
+            "edges": [
+                {
+                    "source": 2, 
+                    "guard": "id=0", 
+                    "update": "x3:=0", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "x3<=2", 
+                    "update": "x3:=0,id:=3", 
+                    "target": 0, 
+                    "label": ""
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "id=0", 
+                    "update": "x3:=0", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 0, 
+                    "guard": "id=3&& x3>2", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "x3:=0,id:=0", 
+                    "target": 2, 
+                    "label": ""
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "a"
+                }, 
+                {
+                    "invariant": "x4<=2", 
+                    "id": 1, 
+                    "name": "b"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "c"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "cs"
+                }
+            ], 
+            "initial": 0, 
+            "name": "p4", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "id=0", 
+                    "update": "x4:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "x4<=2", 
+                    "update": "x4:=0,id:=4", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=0", 
+                    "update": "x4:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=4&& x4>2", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "x4:=0,id:=0", 
+                    "target": 0, 
+                    "label": ""
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 2, 
+                    "name": "c"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "cs"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "a"
+                }, 
+                {
+                    "invariant": "x5<=2", 
+                    "id": 1, 
+                    "name": "b"
+                }
+            ], 
+            "initial": 0, 
+            "name": "p5", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "id=0", 
+                    "update": "x5:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "x5<=2", 
+                    "update": "x5:=0,id:=5", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=0", 
+                    "update": "x5:=0", 
+                    "target": 1, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "id=5&& x5>2", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": ""
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "x5:=0,id:=0", 
+                    "target": 0, 
+                    "label": ""
+                }
+            ]
+        }
+    ], 
+    "clocks": "x1, x2, x3, x4, x5", 
+    "vars": "id[-1000:1000]",
+    "formula": "E<> ((p2.cs && (p1.cs)) || (p3.cs && (p1.cs || p2.cs)) || (p4.cs && (p1.cs || p2.cs || p3.cs)) || (p5.cs && (p1.cs || p2.cs || p3.cs || p4.cs)) )"
+}
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/hddi_08.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/hddi_08.muntax
@@ -0,0 +1,959 @@
+{
+    "automata": [
+        {
+            "nodes": [
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 14, 
+                    "name": "ring_to_5"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 15, 
+                    "name": "ring_5"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 12, 
+                    "name": "ring_to_4"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 13, 
+                    "name": "ring_4"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 10, 
+                    "name": "ring_to_3"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 11, 
+                    "name": "ring_3"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 8, 
+                    "name": "ring_to_2"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 9, 
+                    "name": "ring_2"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 0, 
+                    "name": "ring_to_1"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 1, 
+                    "name": "ring_1"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 2, 
+                    "name": "ring_to_6"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "ring_6"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 4, 
+                    "name": "ring_to_7"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 5, 
+                    "name": "ring_7"
+                }, 
+                {
+                    "invariant": "RING_x <= 0", 
+                    "id": 6, 
+                    "name": "ring_to_8"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 7, 
+                    "name": "ring_8"
+                }
+            ], 
+            "initial": 0, 
+            "name": "RING", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 1, 
+                    "label": "tt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 8, 
+                    "label": "rt1?"
+                }, 
+                {
+                    "source": 8, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 9, 
+                    "label": "tt2!"
+                }, 
+                {
+                    "source": 9, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 10, 
+                    "label": "rt2?"
+                }, 
+                {
+                    "source": 10, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 11, 
+                    "label": "tt3!"
+                }, 
+                {
+                    "source": 11, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 12, 
+                    "label": "rt3?"
+                }, 
+                {
+                    "source": 12, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 13, 
+                    "label": "tt4!"
+                }, 
+                {
+                    "source": 13, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 14, 
+                    "label": "rt4?"
+                }, 
+                {
+                    "source": 14, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 15, 
+                    "label": "tt5!"
+                }, 
+                {
+                    "source": 15, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 2, 
+                    "label": "rt5?"
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "tt6!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 4, 
+                    "label": "rt6?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": "tt7!"
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 6, 
+                    "label": "rt7?"
+                }, 
+                {
+                    "source": 6, 
+                    "guard": "RING_x <= 0", 
+                    "update": "", 
+                    "target": 7, 
+                    "label": "tt8!"
+                }, 
+                {
+                    "source": 7, 
+                    "guard": "", 
+                    "update": "RING_x:= 0", 
+                    "target": 0, 
+                    "label": "rt8?"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "ST1_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST1_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST1_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST1", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST1_y := 0, ST1_x:= 0", 
+                    "target": 1, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST1_x >= 20 && ST1_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST1_z := 0, ST1_x:= 0", 
+                    "target": 4, 
+                    "label": "tt1?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST1_x >= 20 && ST1_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt1!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST2_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST2_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST2_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST2", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST2_y := 0, ST2_x:= 0", 
+                    "target": 1, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST2_x >= 20 && ST2_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST2_z := 0, ST2_x:= 0", 
+                    "target": 4, 
+                    "label": "tt2?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST2_x >= 20 && ST2_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt2!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST3_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST3_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST3_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST3_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST3", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST3_y := 0, ST3_x:= 0", 
+                    "target": 1, 
+                    "label": "tt3?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST3_x >= 20 && ST3_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST3_x >= 20 && ST3_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST3_z := 0, ST3_x:= 0", 
+                    "target": 4, 
+                    "label": "tt3?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST3_x >= 20 && ST3_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt3!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST3_x >= 20 && ST3_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt3!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST4_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST4_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST4_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST4_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST4", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST4_y := 0, ST4_x:= 0", 
+                    "target": 1, 
+                    "label": "tt4?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST4_x >= 20 && ST4_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST4_x >= 20 && ST4_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST4_z := 0, ST4_x:= 0", 
+                    "target": 4, 
+                    "label": "tt4?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST4_x >= 20 && ST4_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt4!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST4_x >= 20 && ST4_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt4!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST5_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST5_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST5_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST5_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST5", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST5_y := 0, ST5_x:= 0", 
+                    "target": 1, 
+                    "label": "tt5?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST5_x >= 20 && ST5_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST5_x >= 20 && ST5_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST5_z := 0, ST5_x:= 0", 
+                    "target": 4, 
+                    "label": "tt5?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST5_x >= 20 && ST5_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt5!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST5_x >= 20 && ST5_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt5!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST6_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST6_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST6_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST6_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST6", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST6_y := 0, ST6_x:= 0", 
+                    "target": 1, 
+                    "label": "tt6?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST6_x >= 20 && ST6_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST6_x >= 20 && ST6_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST6_z := 0, ST6_x:= 0", 
+                    "target": 4, 
+                    "label": "tt6?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST6_x >= 20 && ST6_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt6!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST6_x >= 20 && ST6_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt6!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "ST7_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST7_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }, 
+                {
+                    "invariant": "ST7_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST7_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST7", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST7_y := 0, ST7_x:= 0", 
+                    "target": 1, 
+                    "label": "tt7?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST7_x >= 20 && ST7_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST7_x >= 20 && ST7_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST7_z := 0, ST7_x:= 0", 
+                    "target": 4, 
+                    "label": "tt7?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST7_x >= 20 && ST7_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt7!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST7_x >= 20 && ST7_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt7!"
+                }
+            ]
+        }, 
+        {
+            "nodes": [
+                {
+                    "invariant": "", 
+                    "id": 0, 
+                    "name": "station_z_idle"
+                }, 
+                {
+                    "invariant": "ST8_x<=20", 
+                    "id": 1, 
+                    "name": "station_z_sync"
+                }, 
+                {
+                    "invariant": "", 
+                    "id": 3, 
+                    "name": "station_y_idle"
+                }, 
+                {
+                    "invariant": "ST8_z<=420", 
+                    "id": 2, 
+                    "name": "station_z_async"
+                }, 
+                {
+                    "invariant": "ST8_y<=420", 
+                    "id": 5, 
+                    "name": "station_y_async"
+                }, 
+                {
+                    "invariant": "ST8_x<=20", 
+                    "id": 4, 
+                    "name": "station_y_sync"
+                }
+            ], 
+            "initial": 0, 
+            "name": "ST8", 
+            "edges": [
+                {
+                    "source": 0, 
+                    "guard": "", 
+                    "update": "ST8_y := 0, ST8_x:= 0", 
+                    "target": 1, 
+                    "label": "tt8?"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST8_x >= 20 && ST8_z >= 420", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 1, 
+                    "guard": "ST8_x >= 20 && ST8_z < 420", 
+                    "update": "", 
+                    "target": 2, 
+                    "label": ""
+                }, 
+                {
+                    "source": 2, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 3, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 3, 
+                    "guard": "", 
+                    "update": "ST8_z := 0, ST8_x:= 0", 
+                    "target": 4, 
+                    "label": "tt8?"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST8_x >= 20 && ST8_y >= 420", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt8!"
+                }, 
+                {
+                    "source": 4, 
+                    "guard": "ST8_x >= 20 && ST8_y < 420", 
+                    "update": "", 
+                    "target": 5, 
+                    "label": ""
+                }, 
+                {
+                    "source": 5, 
+                    "guard": "", 
+                    "update": "", 
+                    "target": 0, 
+                    "label": "rt8!"
+                }
+            ]
+        }
+    ], 
+    "clocks": "RING_x, ST1_x, ST1_y, ST1_z, ST2_x, ST2_y, ST2_z, ST3_x, ST3_y, ST3_z, ST4_x, ST4_y, ST4_z, ST5_x, ST5_y, ST5_z, ST6_x, ST6_y, ST6_z, ST7_x, ST7_y, ST7_z, ST8_x, ST8_y, ST8_z", 
+    "vars": "",
+    "formula": "E<> ((ST1.station_z_sync || ST1.station_z_async || ST1.station_y_sync || ST1.station_y_async) && (ST2.station_z_sync || ST2.station_z_async || ST2.station_y_sync || ST2.station_y_async))"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/light_switch.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/light_switch.muntax
@@ -0,0 +1,1 @@
+{"automata":[{"name":"User","initial":2,"nodes":[{"id":9,"name":"B","x":592.1171264648438,"y":196.72305297851562,"invariant":"x <= 3"},{"id":4,"name":"C","x":593.5267333984375,"y":411.2369689941406,"invariant":""},{"id":2,"name":"A","x":271.3272399902344,"y":286.3172607421875,"invariant":""}],"edges":[{"source":9,"target":4,"guard":"","label":"press!","update":""},{"source":2,"target":9,"guard":"","label":"press!","update":"x := 0"},{"source":4,"target":2,"guard":"","label":"press!","update":""}]},{"name":"Lamp","initial":5,"nodes":[{"id":8,"name":"fading","x":595.3662109375,"y":415.7825622558594,"invariant":""},{"id":7,"name":"bright","x":875.3230590820312,"y":199.18556213378906,"invariant":""},{"id":6,"name":"low","x":596.2827758789062,"y":201.7791748046875,"invariant":""},{"id":5,"name":"off","x":327.5268249511719,"y":200.4239501953125,"invariant":""}],"edges":[{"source":6,"target":8,"guard":"c >= 5","label":"press?","update":""},{"source":8,"target":5,"guard":"","label":"","update":""},{"source":7,"target":8,"guard":"","label":"press?","update":""},{"source":6,"target":7,"guard":"c < 5","label":"press?","update":""},{"source":5,"target":6,"guard":"","label":"press?","update":"c := 0"}]}],"clocks":"c, x","vars":"","formula":"A<> Lamp.bright"}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/benchmarks/simple.muntax
--- /dev/null
+++ thys/Munta_Model_Checker/benchmarks/simple.muntax
@@ -0,0 +1,93 @@
+{
+    "automata": [
+        {
+            "name": "TA2",
+            "initial": 8,
+            "nodes": [
+                {
+                    "id": 9,
+                    "name": "N",
+                    "x": 831.51513671875,
+                    "y": 294.54547119140625,
+                    "invariant": ""
+                },
+                {
+                    "id": 8,
+                    "name": "M",
+                    "x": 569.697021484375,
+                    "y": 290.9090881347656,
+                    "invariant": ""
+                }
+            ],
+            "edges": [
+                {
+                    "source": 8,
+                    "target": 9,
+                    "guard": "",
+                    "label": "",
+                    "update": ""
+                },
+                {
+                    "source": 9,
+                    "target": 9,
+                    "guard": "",
+                    "label": "",
+                    "update": ""
+                }
+            ]
+        },
+        {
+            "name": "TA1",
+            "initial": 5,
+            "nodes": [
+                {
+                    "id": 7,
+                    "name": "C",
+                    "x": 831.51513671875,
+                    "y": 269.0909118652344,
+                    "invariant": ""
+                },
+                {
+                    "id": 6,
+                    "name": "B",
+                    "x": 602,
+                    "y": 82,
+                    "invariant": ""
+                },
+                {
+                    "id": 5,
+                    "name": "A",
+                    "x": 392,
+                    "y": 268.69696044921875,
+                    "invariant": "x <= 3"
+                }
+            ],
+            "edges": [
+                {
+                    "source": 5,
+                    "target": 7,
+                    "guard": "(x > 0) && y < 2",
+                    "label": "",
+                    "update": "x := 0"
+                },
+                {
+                    "source": 7,
+                    "target": 6,
+                    "guard": "",
+                    "label": "",
+                    "update": ""
+                },
+                {
+                    "source": 6,
+                    "target": 5,
+                    "guard": "",
+                    "label": "",
+                    "update": "y := 0, x := 0"
+                }
+            ]
+        }
+    ],
+    "clocks": "x, y",
+    "vars": "",
+    "formula": "E<> TA1.C"
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/document/root.bib
--- /dev/null
+++ thys/Munta_Model_Checker/document/root.bib
@@ -0,0 +1,79 @@
+@phdthesis{wimmer-phd,
+  author       = {Simon Wimmer},
+  title        = {Trustworthy Verification of Realtime Systems},
+  school       = {Technical University of Munich, Germany},
+  year         = {2020},
+  url          = {https://nbn-resolving.org/urn:nbn:de:bvb:91-diss-20201209-1576100-1-0},
+  urn          = {urn:nbn:de:bvb:91-diss-20201209-1576100-1-0},
+  timestamp    = {Mon, 13 Dec 2021 13:21:14 +0100},
+  biburl       = {https://dblp.org/rec/phd/dnb/Wimmer20.bib},
+  bibsource    = {dblp computer science bibliography, https://dblp.org}
+}
+
+
+@inproceedings{munta-tacas,
+  author       = {Simon Wimmer and
+                  Peter Lammich},
+  title        = {Verified Model Checking of Timed Automata},
+  booktitle    = {{TACAS} {(1)}},
+  series       = {Lecture Notes in Computer Science},
+  volume       = {10805},
+  pages        = {61--78},
+  publisher    = {Springer},
+  year         = {2018}
+}
+
+@inproceedings{munta-tool-paper,
+  author       = {Simon Wimmer},
+  editor       = {{\'{E}}tienne Andr{\'{e}} and
+                  Mari{\"{e}}lle Stoelinga},
+  title        = {Munta: {A} Verified Model Checker for Timed Automata},
+  booktitle    = {Formal Modeling and Analysis of Timed Systems - 17th International
+                  Conference, {FORMATS} 2019, Amsterdam, The Netherlands, August 27-29,
+                  2019, Proceedings},
+  series       = {Lecture Notes in Computer Science},
+  volume       = {11750},
+  pages        = {236--243},
+  publisher    = {Springer},
+  year         = {2019},
+  url          = {https://doi.org/10.1007/978-3-030-29662-9\_14},
+  doi          = {10.1007/978-3-030-29662-9\_14},
+  timestamp    = {Mon, 13 Dec 2021 13:21:14 +0100},
+  biburl       = {https://dblp.org/rec/conf/formats/Wimmer19.bib},
+  bibsource    = {dblp computer science bibliography, https://dblp.org}
+}
+
+
+
+@article{timed-automata-afp,
+  author  = {Simon Wimmer},
+  title   = {Timed Automata},
+  journal = {Archive of Formal Proofs},
+  month   = {March},
+  year    = {2016},
+  note    = {\url{https://isa-afp.org/entries/Timed_Automata.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}
+
+@article{worklist-afp,
+  author  = {Simon Wimmer and Peter Lammich},
+  title   = {Worklist Algorithms},
+  journal = {Archive of Formal Proofs},
+  month   = {August},
+  year    = {2024},
+  note    = {\url{https://isa-afp.org/entries/Worklist_Algorithms.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}
+
+@article{dbm-afp,
+  author  = {Simon Wimmer and Peter Lammich},
+  title   = {Difference Bound Matrices},
+  journal = {Archive of Formal Proofs},
+  month   = {August},
+  year    = {2024},
+  note    = {\url{https://isa-afp.org/entries/Difference_Bound_Matrices.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/document/root.tex
--- /dev/null
+++ thys/Munta_Model_Checker/document/root.tex
@@ -0,0 +1,110 @@
+\documentclass[10pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+\usepackage{a4wide}
+\usepackage[english]{babel}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+\usepackage{wasysym}
+% for \<hole>
+
+% this should be the last package used
+\usepackage{pdfsetup}
+
+% urls in roman style, theory text in math-similar italics
+\urlstyle{rm}
+\isabellestyle{literal}
+
+% for uniform font size
+\renewcommand{\isastyle}{\isastyleminor}
+
+\renewcommand{\isamarkupsection}[1]{\section{#1}}
+\renewcommand{\isamarkupsubsection}[1]{\subsection{#1}}
+\renewcommand{\isamarkupsubsubsection}[1]{\paragraph{#1}}
+
+\newcommand{\munta}{\textsc{Munta}}
+
+\begin{document}
+
+\title{\munta: A Verified Model Checker for Timed Automata}
+\author{Simon Wimmer}
+\date{}
+
+\maketitle
+\begin{abstract}
+  \munta\ is a verified model checker for timed automata.
+  It has been described in detail in a PhD thesis \cite{wimmer-phd}
+  and conference publications~\cite{munta-tacas,munta-tool-paper}.
+  This entry supersedes earlier versions of \munta\ hosted on GitHub.
+\end{abstract}
+
+\setcounter{tocdepth}{2}
+\tableofcontents
+\newpage
+
+\section*{Introduction}
+
+\munta\ is a verified model checker for timed automata.
+It has been described in detail in a PhD thesis \cite{wimmer-phd}
+and presented in peer-reviewed conference publications~\cite{munta-tacas,munta-tool-paper}.
+
+This AFP entry builds upon and extends formalizations from the following previous entries:
+\begin{itemize}
+  \item \texttt{Timed\_Automata}~\cite{timed-automata-afp}
+  \item \texttt{Worklist\_Algorithms}~\cite{worklist-afp}
+  \item \texttt{Difference\_Bound\_Matrices}~\cite{dbm-afp}
+\end{itemize}
+and supersedes an earlier version of \munta\ hosted on GitHub:
+\url{https://github.com/wimmers/munta}.
+
+\paragraph{Usage}
+Theory \texttt{Simple\_Network\_Language\_Export\_Code} (page \pageref{export-munta}) describes
+how to obtain executable code for and run \munta\ from within Isabelle.
+
+This entry also provides a build setup for the \textsc{munta} executable
+using the MLton SML compiler.
+Build and run instructions are given in theory \texttt{Munta\_Compile\_MLton}
+(on page \pageref{build-mlton}).
+A number of example model files can be found in the \texttt{benchmarks} directory.
+
+\paragraph{Graphical Interface} 
+A graphical user interface for \textsc{munta} is also available:
+\url{https://munta.isabelle.systems}.
+The GUI frontend can communicate with a local instance of \munta\ behind a Python-based server
+or run \munta\ directly in the browser.
+Detailed run instructions are given on Github:
+\url{https://github.com/wimmers/munta/?tab=readme-ov-file#graphical-user-interface}.
+
+\paragraph{OCaml Code Generation} 
+While the default backend targets SML/MLton,
+\textsc{munta} in principle also supports code generation to OCaml
+(c.f.\ theory \texttt{Simple\_Network\_Language\_Export\_Code} on page \pageref{export-ocaml}).
+Among its use cases is the browser-based version of the GUI, which is compiled from the OCaml code.
+For instructions on compiling with the \texttt{dune} build system, see the \texttt{ocaml} branch:
+\url{https://github.com/wimmers/munta/tree/ocaml}.
+
+\newpage
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Abstract_Term.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Abstract_Term.thy
@@ -0,0 +1,59 @@
+theory Abstract_Term
+  imports Main
+begin
+
+ML \<open>
+  fun pull_let ctxt target name body =
+    let
+      val body1 = abstract_over (target, body);
+      val _ =
+        if body aconv body1 then
+          raise (TERM ("could not pull term; type mismatch?", [target, body]))
+        else ()
+      val r1 = Const (const_name\<open>HOL.Let\<close>, dummyT) $ target $ Abs (name, dummyT, body1);
+      val ct1 = Syntax.check_term ctxt r1;
+      val g1 =
+        Goal.prove ctxt [] [] (Logic.mk_equals (body, ct1))
+        (fn {context, ...} => EqSubst.eqsubst_tac context [0] [@{thm Let_def}] 1
+        THEN resolve_tac context [@{thm Pure.reflexive}] 1)
+    in g1 end;
+
+  fun get_rhs thm =
+    let
+      val Const\<open>Pure.eq _ for _ r\<close> = Thm.full_prop_of thm
+    in r end;
+
+  fun get_lhs thm =
+    let
+      val Const\<open>Pure.imp for Const\<open>Pure.eq _ for l _\<close> _\<close> = Thm.full_prop_of thm
+    in l end;
+
+  fun pull_tac' ctxt u name thm =
+    let
+      val l = get_lhs thm;
+      val rewr = pull_let ctxt u name l;
+    in Local_Defs.unfold_tac ctxt [rewr] thm end;
+
+  fun pull_tac ctxt u name = CHANGED (SELECT_GOAL (pull_tac' ctxt u name) 1);
+\<close>
+
+ML \<open>
+fun lift_parser p = fn (ctxt, x) => p x |> (fn (r, s) => (r, (ctxt, s)))
+\<close>
+
+method_setup abstract_let =
+  \<open>Args.term -- Scan.option (lift_parser Args.name) >> (
+    fn (t, n_opt) => fn ctxt => SIMPLE_METHOD (pull_tac ctxt t (Option.getOpt (n_opt, "_"))))
+  \<close>
+  "Abstract over a subterm and extract it into a Let-binding"
+
+text \<open>Example Usage\<close>
+
+schematic_goal "(1 :: nat) * 2 * 3 \<equiv> ?x"
+  cancel\<open>apply (abstract_let 1)\<close> \<comment> \<open>Results in error: types do not match\<close>
+  apply (abstract_let "1 :: nat" one)
+  apply (abstract_let "2 :: nat")
+  apply (abstract_let "3 :: nat" two)
+  by (rule Pure.reflexive)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Bijective_Embedding.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Bijective_Embedding.thy
@@ -0,0 +1,341 @@
+theory Bijective_Embedding
+  imports "HOL-Library.Countable_Set"
+begin
+
+text \<open>Generalization of @{thm image_vimage_eq}\<close>
+lemma inj_on_vimage_image_eq: "f -` (f ` A) \<inter> S = A" if "inj_on f S" "A \<subseteq> S"
+  using that unfolding inj_on_def by blast
+
+text \<open>Generalization of @{thm card_vimage_inj}\<close>
+theorem card_vimage_inj_on:
+  fixes f :: "'a \<Rightarrow> 'b"
+    and A :: "'b set"
+  assumes "inj_on f S"
+    and "A \<subseteq> f ` S"
+  shows "card (f -` A \<inter> S) = card A"
+  using assms
+  by (auto 4 3 simp: subset_image_iff inj_on_vimage_image_eq
+      intro: card_image[symmetric, OF subset_inj_on])
+
+context
+begin
+
+private lemma finiteI: "finite {x. x \<in> A \<and> ordX x \<le> ordX a \<and> Q x}" (is "finite ?S")
+    if "inj_on ordX A" for A a Q and ordX :: "_ \<Rightarrow> nat"
+proof -
+  have "?S \<subseteq> {x. x \<in> A \<and> ordX x \<le> ordX a}"
+    by blast
+  moreover have "finite \<dots>"
+  proof -
+    have "\<dots> = ordX -` {n. n \<le> ordX a} \<inter> A"
+      by auto
+    moreover from \<open>inj_on ordX A\<close> have "finite \<dots>"
+      by (intro finite_vimage_IntI) auto
+    ultimately show ?thesis
+      by simp
+  qed
+  ultimately show ?thesis
+    by (rule finite_subset)
+qed
+
+qualified lemma enumeration_skip_finite_set_surj:
+  "\<exists>a. a \<in> A \<and> a \<notin> S \<and> card {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> S} = n"
+  if "n > 0" "inj_on ordX A" "ordX ` A = \<nat>" "finite S" "S \<subseteq> A" for ordX :: "_ \<Rightarrow> nat"
+  using \<open>finite S\<close> \<open>n > 0\<close> \<open>S \<subseteq> A\<close>
+proof (induction arbitrary: n)
+  case empty
+  let ?a = "the_inv_into A ordX (n - 1)"
+  have "card {x \<in> A. ordX x \<le> n - 1} = n"
+  proof -
+    have "{x \<in> A. ordX x \<le> n - 1} = ordX -` {x. x \<le> n - 1} \<inter> A"
+      by auto
+    moreover have "card \<dots> = card {x. x \<le> n - 1}"
+      using \<open>inj_on ordX _\<close> \<open>ordX ` A = \<nat>\<close> by (intro card_vimage_inj_on) (auto simp: Nats_def)
+    moreover have "\<dots> = n"
+      using \<open>n > 0\<close> by auto
+    ultimately show ?thesis
+      by auto
+  qed
+  then have "\<exists> a \<in> A. card {x \<in> A. ordX x \<le> ordX a} = n"
+    using \<open>inj_on ordX A\<close> \<open>ordX ` A = \<nat>\<close>
+    unfolding Nats_def
+    by (intro bexI[where x = "the_inv_into A ordX (n - 1)"])
+       (auto intro: the_inv_into_into simp: f_the_inv_into_f[where f = ordX])
+  then show ?case
+    by auto
+next
+  case (insert x F)
+  then obtain a where a: "a \<in> A" "a \<notin> F" "card {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F} = n"
+    by blast
+  show ?case
+  proof (cases "ordX x \<le> ordX a")
+    case True
+    then have *:
+      "{xa \<in> A. ordX xa \<le> ordX a \<and> xa \<notin> insert x F} = {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F} - {x}"
+      by auto
+    let ?m = "Max (ordX ` F)"
+    show ?thesis
+    proof (cases "\<exists>b \<in> A. b \<notin> F \<and> ordX a < ordX b \<and> ordX b \<le> ?m")
+      case False
+      let ?a = "the_inv_into A ordX (max ?m (ordX a) + 1)"
+      let ?S = "{xa \<in> A. ordX xa \<le> ordX ?a \<and> xa \<notin> insert x F}"
+      have "?a \<in> A"
+        by (simp add: Nats_def \<open>inj_on ordX A\<close> \<open>ordX ` A = \<nat>\<close> the_inv_into_into)
+      have "ordX ?a = max ?m (ordX a) + 1"
+        using \<open>inj_on ordX A\<close> \<open>ordX ` A = \<nat>\<close> unfolding Nats_def
+        by (subst f_the_inv_into_f[where f = ordX]) auto
+      then have "?S = {xa \<in> A. ordX xa \<le> max ?m (ordX a) + 1 \<and> xa \<notin> insert x F}"
+        by simp
+      also have "\<dots> = {xa \<in> A. ordX xa \<le> ordX a \<and> xa \<notin> insert x F} \<union> {?a}"
+      proof -
+        have "{xa \<in> A. ordX xa = max ?m (ordX a) + 1 \<and> xa \<notin> insert x F} = {?a}"
+          unfolding \<open>ordX ?a = _\<close>
+          apply (auto simp: \<open>inj_on ordX A\<close> the_inv_into_f_eq)
+          using True \<open>ordX ?a = _\<close> \<open>?a \<in> A\<close>
+          apply -
+          apply (auto; fail)+
+          by (metis Max.boundedE Suc_eq_plus1 Suc_n_not_le_n emptyE finite_imageI imageI
+                insert.hyps(1) max.cobounded1)
+        have *: "{xa \<in> A. ordX xa \<le> ordX a \<and> xa \<notin> insert x F}
+              = {xa \<in> A. ordX xa \<le> max ?m (ordX a) \<and> xa \<notin> insert x F}"
+          using False by auto
+        show ?thesis
+          unfolding \<open>_ = {?a}\<close>[symmetric] * by auto
+      qed
+      finally have "?S = ({x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F} - {x}) \<union> {?a}"
+        by auto
+      moreover have "?a \<noteq> x"
+        using True \<open>ordX ?a = max (MAX x \<in> F. ordX x) (ordX a) + 1\<close> by auto
+      moreover have "?a \<notin> {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F}"
+        using \<open>ordX ?a = max (MAX x \<in> F. ordX x) (ordX a) + 1\<close> by auto
+      moreover have "x \<in> {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F}"
+        using \<open>ordX x \<le> _\<close> \<open>insert x _ \<subseteq> A\<close> \<open>x \<notin> F\<close> by auto
+      ultimately have "card ?S = card {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F}"
+        apply simp
+        apply (subst card_insert_disjoint)
+        subgoal
+          by (force intro: finiteI[OF \<open>inj_on ordX A\<close>, of a])
+        apply simp
+        apply (subst card.remove[symmetric])
+        subgoal
+          by (force intro: finiteI[OF \<open>inj_on ordX A\<close>, of a])
+        by auto
+      also have "\<dots> = n"
+        using \<open>_ = n\<close> by auto
+      finally show ?thesis
+        using \<open>inj_on ordX A\<close> \<open>ordX ` A = \<nat>\<close> unfolding Nats_def
+        apply -
+        apply (rule exI[where x = ?a], rule conjI)
+         apply (auto intro: the_inv_into_into)
+        using \<open>ordX x \<le> _\<close>
+          apply (subgoal_tac "ordX x = Suc (ordX a)"; simp add: f_the_inv_into_f[where f = ordX])
+         apply (subgoal_tac "ordX ?a \<le> ?m")
+        subgoal
+          using False
+          apply (simp add: f_the_inv_into_f[where f = ordX])
+          done
+         apply (rule Max_ge)
+        using \<open>finite F\<close> apply (rule finite_imageI)
+         apply (rule imageI)
+         apply (simp add: f_the_inv_into_f[where f = ordX])
+        done
+    next
+      case True
+      let ?M = "{b \<in> A. ordX a < ordX b \<and> ordX b \<le> ?m \<and> b \<notin> F}"
+      from True have "?M \<noteq> {}"
+        by auto
+      have "finite ?M"
+      proof -
+        have "?M \<subseteq> {b \<in> A. ordX b \<le> ?m}"
+          by auto
+        moreover have "finite \<dots>"
+        proof -
+          have *: "\<dots> = ordX -` {x. x \<le> ?m} \<inter> A"
+            by auto
+          from \<open>inj_on ordX A\<close> show ?thesis
+            unfolding * by (intro finite_vimage_IntI) auto
+        qed
+        ultimately show ?thesis
+          by (rule finite_subset)
+      qed
+      let ?a = "arg_min_on ordX ?M"
+      from arg_min_if_finite[OF \<open>finite ?M\<close> \<open>?M \<noteq> {}\<close>, of ordX] have a:
+        "?a \<in> ?M" "\<not> (\<exists> x \<in> ?M. ordX x < ordX ?a)"
+        by fast+
+      with \<open>ordX x \<le> ordX a\<close> have "?a \<noteq> x"
+        by auto
+      then have **: "{xa \<in> A. ordX xa \<le> ordX ?a \<and> xa \<noteq> x \<and> xa \<notin> F} =
+            {xa \<in> A. ordX xa \<le> ordX a \<and> xa \<noteq> x \<and> xa \<notin> F} \<union> {?a}"
+        using a
+        by auto
+           (smt \<open>inj_on ordX A\<close> inj_on_eq_iff le_eq_less_or_eq le_trans mem_Collect_eq not_le)
+      have "?a \<notin> {xa \<in> A. ordX xa \<le> ordX a \<and> xa \<noteq> x \<and> xa \<notin> F}"
+        using a(1) by auto
+      then have "card {xa \<in> A. ordX xa \<le> ordX ?a \<and> xa \<noteq> x \<and> xa \<notin> F} = n"
+        unfolding *[simplified] **
+        using \<open>card _ = n\<close>
+        apply simp
+        apply (subst card_insert_disjoint)
+        subgoal
+          by (auto intro: finiteI[OF \<open>inj_on ordX A\<close>])
+        apply simp
+        apply (subst card.remove[symmetric])
+        subgoal
+          by (auto intro: finiteI[OF \<open>inj_on ordX A\<close>])
+        subgoal
+          using \<open>insert x F \<subseteq> A\<close> \<open>ordX x \<le> _\<close> \<open>x \<notin> F\<close> by simp
+        apply assumption
+        done
+      then show ?thesis
+        using a(1) \<open>ordX x \<le> _\<close> by (intro exI[where x = ?a]) auto
+    qed
+  next
+    case False
+    then have
+      "{xa \<in> A. ordX xa \<le> ordX a \<and> xa \<notin> insert x F} = {x \<in> A. ordX x \<le> ordX a \<and> x \<notin> F}"
+      by auto
+    with a False show ?thesis
+      by (intro exI[where x = a]) auto
+  qed
+qed
+
+lemma bij_betw_relI:
+  assumes "\<And>x y z. x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> z \<in> B \<Longrightarrow> R x y \<Longrightarrow> R x z \<Longrightarrow> y = z"
+      and "\<And>x y z. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> z \<in> B \<Longrightarrow> R x z \<Longrightarrow> R y z \<Longrightarrow> x = y"
+      and "\<And>x. x \<in> A \<Longrightarrow> \<exists>y \<in> B. R x y" "\<And>y. y \<in> B \<Longrightarrow> \<exists>x \<in> A. R x y"
+  shows "bij_betw (\<lambda>a. SOME b. R a b \<and> b \<in> B) A B"
+  by (rule bij_betwI'; smt assms someI)
+
+lemma bijective_embedding:
+  fixes f :: "'a \<Rightarrow> 'b"
+    and A :: "'a set" and B :: "'b set"
+    and S :: "'a set"
+  assumes "inj_on f S" and "S \<subseteq> A" and "f ` S \<subseteq> B"
+    and "finite S"
+    and "countable A" and "countable B"
+    and "infinite A" and "infinite B"
+  shows "\<exists>h. bij_betw h A B \<and> (\<forall>x \<in> S. h x = f x)"
+proof -
+  obtain ordX :: "_ \<Rightarrow> nat" and ordY :: "_ \<Rightarrow> nat" where
+    "inj_on ordX A" "ordX ` A = \<nat>" and "inj_on ordY B" "ordY ` B = \<nat>"
+    using assms(5-) unfolding Nats_def
+    by (metis bij_betw_def bij_betw_from_nat_into bij_betw_imp_inj_on bij_betw_the_inv_into
+        of_nat_id surj_def)
+  define P where "P a b \<equiv> a \<in> A \<and> b \<in> B \<and> a \<notin> S \<and> b \<notin> f ` S \<and>
+    card {x. x \<in> A \<and> ordX x \<le> ordX a \<and> x \<notin> S} = card {x. x \<in> B \<and> ordY x \<le> ordY b \<and> x \<notin> f ` S}"
+    for a b
+  let ?f = "\<lambda>a. card {x. x \<in> A \<and> ordX x \<le> ordX a \<and> x \<notin> S}"
+  let ?g = "\<lambda>a. card {x. x \<in> B \<and> ordY x \<le> ordY a \<and> x \<notin> f ` S}"
+  have P_right: "\<exists> b. P a b" if "a \<in> A" "a \<notin> S" for a
+  proof -
+    have *: "\<exists>b. b \<in> B \<and> b \<notin> f ` S \<and> ?g b = n" if "n > 0" for n
+      using \<open>n > 0\<close> \<open>finite S\<close> \<open>f ` S \<subseteq> B\<close> \<open>inj_on ordY B\<close> \<open>ordY ` B = \<nat>\<close>
+      by (intro enumeration_skip_finite_set_surj) auto
+    from that have "?f a > 0"
+      unfolding card_gt_0_iff by (auto intro: finiteI[OF \<open>inj_on ordX A\<close>])
+    from *[OF this] show ?thesis
+      unfolding P_def using that by auto
+  qed
+  have P_left: "\<exists>a. P a b" if "b \<in> B" "b \<notin> f ` S" for b
+    using that \<open>finite S\<close> \<open>S \<subseteq> A\<close> \<open>inj_on ordX A\<close> \<open>inj_on ordY B\<close> \<open>ordX ` A = \<nat>\<close> unfolding P_def
+    by (auto 4 3 intro: enumeration_skip_finite_set_surj finiteI simp: card_gt_0_iff)
+  have P_surj: "a = b" if "P a c" "P b c" for a b c
+  proof -
+    from that have "?f a = ?f b" (is "card ?A = card ?B")
+      unfolding P_def by auto
+    have fin: "finite ?A" "finite ?B"
+      by (intro finiteI \<open>inj_on ordX A\<close>)+
+    have *: "a \<in> A" "b \<in> A" "a \<notin> S" "b \<notin> S"
+      using that unfolding P_def by auto
+    consider (lt) "ordX a < ordX b" | (eq) "ordX a = ordX b" | (gt) "ordX a > ordX b"
+      by force
+    then show ?thesis
+    proof cases
+      case lt
+      with * have "?f a < ?f b"
+        using leD by (intro psubset_card_mono fin) auto
+      with \<open>?f a = ?f b\<close> show ?thesis
+        by auto
+    next
+      case eq
+      then show ?thesis
+        using \<open>inj_on ordX A\<close> \<open>a \<in> A\<close> \<open>b \<in> A\<close> by (auto dest: inj_onD)
+    next
+      case gt
+      with * have "?f a > ?f b"
+        using leD by (intro psubset_card_mono fin) auto
+      with \<open>?f a = ?f b\<close> show ?thesis
+        by auto
+    qed
+  qed
+  have P_inj: "b = c" if "P a b" "P a c" for a b c
+  proof -
+    from that have "?g b = ?g c" (is "card ?A = card ?B")
+      unfolding P_def by auto
+    have fin: "finite ?A" "finite ?B"
+      by (intro finiteI \<open>inj_on ordY B\<close>)+
+    have *: "b \<in> B" "c \<in> B" "b \<notin> f ` S" "c \<notin> f ` S"
+      using that unfolding P_def by auto
+    consider (lt) "ordY b < ordY c" | (eq) "ordY b = ordY c" | (gt) "ordY b > ordY c"
+      by force
+    then show ?thesis
+    proof cases
+      case lt
+      with * have "?g b < ?g c"
+        using leD by (intro psubset_card_mono fin) (auto, blast)
+      with \<open>?g b = ?g c\<close> show ?thesis
+        by auto
+    next
+      case eq
+      then show ?thesis
+        using \<open>inj_on ordY B\<close> \<open>b \<in> B\<close> \<open>c \<in> B\<close> by (auto dest: inj_onD)
+    next
+      case gt
+      with * have "?g b > ?g c"
+        using leD by (intro psubset_card_mono fin) (auto, blast)
+      with \<open>?g b = ?g c\<close> show ?thesis
+        by auto
+    qed
+  qed
+
+  define R where "R a b \<equiv> if a \<in> S then b = f a else P a b" for a b
+  have R_A: "a \<in> A" and R_B: "b \<in> B" if "R a b" for a b
+    using that \<open>f ` S \<subseteq> B\<close> \<open>S \<subseteq> A\<close> unfolding R_def by (auto split: if_split_asm simp: P_def)
+  have R_inj: "b = c" if "R a b" "R a c" for a b c
+    using that unfolding R_def by (auto split: if_split_asm dest: P_inj)
+  moreover have R_surj: "a = b" if "R a c" "R b c" for a b c
+    using that unfolding R_def
+    by (auto split: if_split_asm dest: P_surj inj_onD[OF \<open>inj_on f S\<close>]) (auto simp: P_def)
+  moreover have R_right: "\<exists>b. R a b" if "a \<in> A" for a
+    unfolding R_def by (auto dest: P_right[OF \<open>a \<in> A\<close>])
+  moreover have R_left: "\<exists>a. R a b" if "b \<in> B" for b
+    unfolding R_def
+    by (cases "b \<in> f ` S", (auto; fail), (frule P_left[OF \<open>b \<in> B\<close>], auto simp: P_def))
+  ultimately show ?thesis
+    apply (intro exI[of _ "\<lambda>a. SOME b. R a b \<and> b \<in> B"] conjI)
+     apply (rule bij_betw_relI)
+        apply assumption+
+      apply (smt R_A R_B)+
+    using assms(3) by (subst R_def) (simp, blast)
+qed
+
+end
+
+definition extend_bij :: "('a :: countable \<Rightarrow> 'b :: countable) \<Rightarrow> 'a set \<Rightarrow> _" where
+  "extend_bij f S \<equiv> SOME h. bij h \<and> (\<forall>x \<in> S. h x = f x)"
+
+lemma
+  fixes f :: "'a  :: countable \<Rightarrow> 'b :: countable" and S :: "'a set"
+  assumes "infinite (UNIV :: 'a set)" and "infinite (UNIV :: 'b set)"
+      and "inj_on f S" and "finite S"
+    shows extend_bij_bij: "bij (extend_bij f S)"
+      and extend_bij_extends: "\<forall>x \<in> S. extend_bij f S x = f x"
+proof -
+  from bijective_embedding[OF assms(3) _ _ assms(4) _ _ assms(1,2)] obtain h where
+    "bij h \<and> (\<forall>x\<in>S. h x = f x)"
+    by auto
+  then show "bij (extend_bij f S)" "\<forall>x \<in> S. extend_bij f S x = f x"
+    unfolding atomize_conj extend_bij_def by (rule someI[where x = h])
+qed
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Error_List_Monad.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Error_List_Monad.thy
@@ -0,0 +1,65 @@
+theory Error_List_Monad
+  imports Main "HOL-Library.Monad_Syntax"
+begin
+
+
+(* Error monad *)
+type_synonym error = String.literal
+
+datatype 'a result =
+  is_result: Result 'a
+| is_error:  Error "error list"
+
+fun bind where "bind m f = (case m of
+  Result x \<Rightarrow> f x
+| Error es \<Rightarrow> Error es)"
+
+adhoc_overloading "Monad_Syntax.bind" \<rightleftharpoons> bind
+
+fun err_msg where
+  "err_msg m (Error es) = Error (m # es)"
+| "err_msg m x = x"
+
+fun make_err where
+  "make_err m (Error es) = Error (es @ [m])"
+| "make_err m _ = Error [m]"
+
+definition assert_msg where "assert_msg b m = (if b then (\<lambda>x. x) else make_err m)"
+
+definition assert where "assert b m = (if b then Result () else Error [m])"
+
+fun the_result where
+  "the_result (Result x) = x"
+
+fun the_errors where
+  "the_errors (Error es) = es"
+
+fun combine2_gen where
+  "combine2_gen comb (Error e1) (Error e2) = Error (List.append e1 e2)"
+| "combine2_gen comb (Error e)  (Result _) = Error e"
+| "combine2_gen comb (Result _) (Error e)  = Error e"
+| "combine2_gen comb (Result a) (Result b) = comb a b"
+
+(* let combine2: ('a result * 'b result -> ('a * 'b) result) = combine2_gen (fun a b -> Result (a, b)) *)
+
+definition "combine2 = combine2_gen (\<lambda>a b. Result (a, b))"
+
+notation combine2 (infixr "<|>" 60)
+
+fun combine where
+  "combine [] = Result []"
+| "combine (x # xs) = combine2_gen (\<lambda>x xs. Result (x # xs)) x (combine xs)"
+
+abbreviation (input) app (infixl "|>" 59) where "a |> b \<equiv> b a"
+
+definition "combine_map f xs = List.map f xs |> combine"
+
+fun map_errors where
+  "map_errors f (Error errors) = Error (List.map f errors)"
+| "map_errors f r = r"
+
+fun fold_error where
+  "fold_error f [] a = Result a"
+| "fold_error f (x # xs) a = do {a \<leftarrow> f x a; fold_error f xs a}"
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/ML_Util.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/ML_Util.thy
@@ -0,0 +1,89 @@
+theory ML_Util
+  imports Main
+begin
+
+(* Printing util *)
+ML \<open>
+  fun pretty_cterm ctxt ctm = Syntax.pretty_term ctxt (Thm.term_of ctm)
+  val string_of_cterm = Pretty.string_of oo pretty_cterm
+  val string_of_term = Pretty.string_of oo Syntax.pretty_term
+\<close>
+
+(* Printing util *)
+ML \<open>
+  fun string_of_thm ctxt = string_of_cterm ctxt o Thm.cprop_of
+\<close>
+
+ML \<open>
+val term_pat_setup =
+  let
+    val parser = Args.context -- Scan.lift Args.name
+      fun term_pat (ctxt, str) =
+        str
+        |> Proof_Context.read_term_pattern ctxt
+        |> ML_Syntax.print_term
+        |> ML_Syntax.atomic
+in
+  ML_Antiquotation.inline @{binding "term_pat"} (parser >> term_pat) end
+\<close>
+
+setup \<open>term_pat_setup\<close>
+
+ML \<open>
+fun match_resolve ctxt concl goal thm =
+  let
+      val concl_pat = Drule.strip_imp_concl (Thm.cprop_of thm)
+      val insts = Thm.first_order_match (concl, concl_pat)
+    in
+      (Drule.instantiate_normalize insts goal |> resolve_tac ctxt [thm] 1)
+    end handle Pattern.MATCH => Seq.empty
+val fo_assm_tac = Subgoal.FOCUS_PREMS (
+    fn {context = ctxt, concl, prems, ...} => fn goal =>
+      Seq.maps (match_resolve ctxt concl goal) (Seq.of_list prems)
+  )
+\<close>
+
+schematic_goal "c s \<Longrightarrow> c ?x"
+  apply (tactic \<open>fo_assm_tac @{context} 1\<close>)
+  done
+
+schematic_goal "c s \<Longrightarrow> ?P s"
+  apply (tactic \<open>fo_assm_tac @{context} 1\<close>)
+  done
+
+ML \<open>
+fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct)
+val refl_match_tac = Subgoal.FOCUS_PREMS (
+  fn {context = ctxt, concl, ...} =>
+    let
+      val stripped = Thm.dest_arg concl
+      val (l, r) = dest_binop stripped
+      val insts =
+        Thm.first_order_match (r, l)
+        handle Pattern.MATCH => Thm.first_order_match (l, r)
+    in
+      Seq.single o Drule.instantiate_normalize insts THEN
+      resolve_tac ctxt [@{thm HOL.refl}] 1
+    end
+    handle CTERM _ => no_tac | Pattern.MATCH => no_tac
+  )
+\<close>
+
+schematic_goal
+  "x = ?x"
+  by (tactic \<open>refl_match_tac @{context} 1\<close>)
+
+schematic_goal
+  "?x = x"
+  by (tactic \<open>refl_match_tac @{context} 1\<close>)
+
+ML \<open>fun assumption ctxt = SIMPLE_METHOD (fo_assm_tac ctxt 1)\<close>
+ML \<open>fun refl_match ctxt = SIMPLE_METHOD (refl_match_tac ctxt 1)\<close>
+
+method_setup fo_assumption =
+  \<open>Scan.succeed assumption\<close> \<open>Prove by assumption with first-order matching\<close>
+
+method_setup refl_match =
+  \<open>Scan.succeed refl_match\<close> \<open>Prove equality by reflexivity with first-order matching\<close>
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/More_Methods.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/More_Methods.thy
@@ -0,0 +1,141 @@
+theory More_Methods
+  imports Main Reordering_Quantifiers "HOL-Eisbach.Eisbach" ML_Util
+begin
+
+ML \<open>fun mini_ex ctxt = SIMPLE_METHOD (mini_ex_tac ctxt 1)\<close>
+ML \<open>fun defer_ex ctxt = SIMPLE_METHOD (defer_ex_tac ctxt 1)\<close>
+
+method_setup mini_existential =
+  \<open>Scan.succeed mini_ex\<close> \<open>Miniscope existential quantifiers\<close>
+method_setup defer_existential =
+  \<open>Scan.succeed defer_ex\<close> \<open>Rotate first conjunct under existential quantifiers to last position\<close>
+
+method mini_ex = ((simp only: ex_simps[symmetric])?, mini_existential, (simp)?)
+method defer_ex = ((simp only: ex_simps[symmetric])?, defer_existential, (simp)?)
+method defer_ex' = (defer_existential, (simp)?)
+
+method solve_triv =
+  fo_assumption | refl_match | simp; fail | assumption | (erule image_eqI[rotated], solve_triv)
+
+method solve_ex_triv2 = (((rule exI)+)?, (rule conjI)+, solve_triv)
+
+method solve_conj_triv = solve_triv | (rule conjI, solve_conj_triv)
+
+method solve_conj_triv2 =
+  (match conclusion in
+    "_ \<and> _" \<Rightarrow> \<open>rule conjI, solve_conj_triv2\<close>
+  \<bar> _ \<Rightarrow> solve_triv)
+
+method solve_ex_triv = (((rule exI)+)?, solve_conj_triv)
+
+named_theorems intros
+named_theorems elims
+
+lemmas [intros] =
+  allI ballI exI bexI[rotated] conjI impI
+and [elims] =
+  bexE exE bexE conjE impE
+
+method intros uses add  = (intro add intros)
+method elims  uses add  = (elim  add elims)
+
+text \<open>Test case\<close>
+lemma all_mp:
+  "\<forall> x. P x \<longrightarrow> R x" if "\<forall> x. P x \<longrightarrow> Q x" "\<And> x. P x \<Longrightarrow> Q x \<Longrightarrow> R x"
+  using that by (intros; elims add: allE)
+
+method rprem =
+  (match premises in R: _ \<Rightarrow> \<open>rule R\<close>)
+
+text \<open>Test case for @{method rprem}.\<close>
+lemma
+  "A \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (A \<Longrightarrow> B) \<Longrightarrow> B"
+  by rprem
+
+ML \<open>
+fun all_forward_tac ctxt n thms nsubgoal =
+  let
+    val assumes = replicate n (assume_tac ctxt nsubgoal)
+    fun forward thm = forward_tac ctxt [thm] nsubgoal :: assumes |> EVERY |> TRY
+  in EVERY (map forward thms) end;
+\<close>
+
+ML \<open>
+fun xrule_meth tac =
+  Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >>
+  (fn (n, ths) => fn ctxt => SIMPLE_METHOD (tac ctxt n ths 1));
+val all_forward_meth = xrule_meth all_forward_tac
+\<close>
+
+method_setup all_forward = \<open>all_forward_meth\<close> "Apply each rule exactly once in forward manner"
+
+named_theorems forward1
+named_theorems forward2
+named_theorems forward3
+named_theorems forward4
+
+method frule1 = changed \<open>all_forward forward1\<close>
+method frule2 = changed \<open>all_forward (1) forward2\<close>
+method frule3 = changed \<open>all_forward (2) forward3\<close>
+method frule4 = changed \<open>all_forward (3) forward4\<close>
+
+method frules = changed \<open>
+  all_forward (0) forward1,
+  all_forward (1) forward2,
+  all_forward (2) forward3,
+  all_forward (3) forward4
+\<close>
+
+ML \<open>
+fun dedup_prems_tac ctxt =
+  let
+    fun Then NONE tac = SOME tac
+      | Then (SOME tac) tac' = SOME (tac THEN' tac');
+    fun thins H (tac, n) =
+      if H then (tac, n + 1)
+      else (Then tac (rotate_tac n THEN' eresolve_tac ctxt [thin_rl]), 0);
+    fun member x = List.exists (fn y => x = y)
+    fun mark_dups _ [] = []
+    |   mark_dups passed (x :: xs) =
+        if member x passed
+        then false :: mark_dups passed xs
+        else true :: mark_dups (x :: passed) xs
+  in
+    SUBGOAL (fn (goal, i) =>
+      let
+        val Hs = Logic.strip_assums_hyp goal
+        val marked = mark_dups [] Hs
+      in
+        (case fst (fold thins marked (NONE, 0)) of
+          NONE => no_tac
+        | SOME tac => tac i)
+      end)
+  end;
+\<close>
+
+method_setup dedup_prems =
+  \<open>Scan.succeed (fn ctxt => SIMPLE_METHOD' (fn i => CHANGED (dedup_prems_tac ctxt i)))\<close>
+  "Remove duplicate premises"
+
+method repeat methods m =
+  determ \<open>(changed m, repeat m)?\<close>
+
+ML \<open>
+fun REPEAT_ROTATE tac =
+  let
+    fun repeat (trm, i) = let
+      val num_prems = Logic.strip_assums_hyp trm |> length
+      val tac2 = TRY (tac i) THEN rotate_tac 1 i
+      val tacs = replicate num_prems tac2
+    in DETERM (EVERY tacs) end
+in SUBGOAL repeat end
+\<close>
+
+method_setup repeat_rotate =
+ \<open>Method.text_closure >> (fn m => fn ctxt => fn facts =>
+   let
+     fun tac st' = method_evaluate m ctxt facts st'
+   in SIMPLE_METHOD (REPEAT_ROTATE (K tac) 1) facts end)
+\<close>
+
+end (* Theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Printing.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Printing.thy
@@ -0,0 +1,14 @@
+theory Printing
+  imports Main
+begin
+
+definition print :: "String.literal \<Rightarrow> unit" where
+"print x = ()"
+
+definition println :: "String.literal \<Rightarrow> unit" where
+  "println x = print (x + STR ''\<newline>'')"
+
+definition "print_err = print"
+definition "println_err x = print_err (x + STR ''\<newline>'')"
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Reordering_Quantifiers.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Reordering_Quantifiers.thy
@@ -0,0 +1,1051 @@
+theory Reordering_Quantifiers
+  imports Main
+begin
+
+lemma
+  assumes A: "finite A"
+  shows "finite {(a,b,c) | a b c. a \<in> A \<and> b < (n :: nat) \<and> P c}"
+  using assms [[simproc add: finite_Collect]] apply simp
+  oops
+
+lemma
+  assumes A: "finite A"
+  shows "finite {(d, c, a, b) | a b c d. d < n \<and> P c \<and> (a, b) \<in> A}"
+  using assms
+  using [[simproc add: finite_Collect]] apply simp
+  oops
+  
+lemma
+  "finite {x . x \<in> A}" if "finite A"
+  using [[simp_trace]]
+  using that by simp
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. a \<in> A \<and> c < n \<and> t = (a,c)}"
+  apply simp
+  using assms by simp
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {(a, c). a \<in> A \<and> c < n}"
+  apply simp
+  using assms by simp
+
+(* Printing util *)
+ML \<open>
+  fun pretty_cterm ctxt ctm = Syntax.pretty_term ctxt (Thm.term_of ctm)
+  val string_of_cterm = Pretty.string_of oo pretty_cterm
+  val string_of_term = Pretty.string_of oo Syntax.pretty_term
+\<close>
+
+ML_val \<open>tracing (Syntax.string_of_term @{context} @{term "a < b"})\<close>
+
+ML_val \<open>tracing (ML_Syntax.print_term @{term "a < b"})\<close>
+
+ML \<open>
+  fun strip_prop (Const (@{const_name HOL.Trueprop}, _) $ t) = t
+    | strip_prop t = t
+\<close>
+
+declare [[ML_print_depth = 50]]
+
+ML \<open>
+  signature QUANTIFIER1_DATA =
+sig
+  (*functionality*)
+  (*terms to be moved around*)
+  (*arguments: preceding quantifies, term under question, preceding terms*)
+  val move: (term * string * typ) list -> term -> term list -> bool
+  (*always move? if false then moves appear if a non-mover was encountered before*)
+  val force_move: bool
+  (*rotate quantifiers after moving*)
+  val rotate: bool
+  (*abstract syntax*)
+  val dest_eq: term -> (term * term) option
+  val dest_conj: term -> (term * term) option
+  val dest_imp: term -> (term * term) option
+  val conj: term
+  val imp: term
+  (*rules*)
+  val iff_reflection: thm (* P <-> Q ==> P == Q *)
+  val iffI: thm
+  val iff_trans: thm
+  val conjI: thm
+  val conjE: thm
+  val impI: thm
+  val mp: thm
+  val exI: thm
+  val exE: thm
+  val uncurry: thm (* P --> Q --> R ==> P & Q --> R *)
+  val iff_allI: thm (* !!x. P x <-> Q x ==> (!x. P x) = (!x. Q x) *)
+  val iff_exI: thm (* !!x. P x <-> Q x ==> (? x. P x) = (? x. Q x) *)
+  val all_comm: thm (* (!x y. P x y) = (!y x. P x y) *)
+  val ex_comm: thm (* (? x y. P x y) = (? y x. P x y) *)
+end;
+
+signature QUANTIFIER1 =
+sig
+  val prove_one_point_all_tac: Proof.context -> tactic
+  val prove_one_point_ex_tac: Proof.context -> tactic
+  val rearrange_all: Proof.context -> cterm -> thm option
+  (* XXX Need to export this ?*)
+  val rearrange_ex': Proof.context -> term -> thm option
+  val rearrange_ex: Proof.context -> cterm -> thm option
+  val rotate_ex: Proof.context -> cterm -> thm option
+  val miniscope_ex: Proof.context -> cterm -> thm option
+  val rotate_all: Proof.context -> cterm -> thm option
+  val rearrange_ball: (Proof.context -> tactic) -> Proof.context -> cterm -> thm option
+  val rearrange_bex: (Proof.context -> tactic) -> Proof.context -> cterm -> thm option
+  val rearrange_Collect: (Proof.context -> tactic) -> Proof.context -> cterm -> thm option
+end;
+
+functor Quantifier(Data: QUANTIFIER1_DATA): QUANTIFIER1 =
+struct
+
+fun extract_conj trms fst xs t =
+  (case Data.dest_conj t of
+    NONE => NONE
+  | SOME (P, Q) =>
+      let
+        val mover = Data.move xs P trms
+      in
+        if Data.force_move andalso mover then (if fst then NONE else SOME (xs, P, Q))
+        else if Data.force_move andalso Data.move xs Q (P :: trms) then SOME (xs, Q, P)
+        else if mover andalso not fst then SOME (xs, P, Q)
+        else if
+          not Data.force_move andalso (not mover orelse not fst) andalso Data.move xs Q (P :: trms)
+          then SOME (xs, Q, P)
+        else
+          (case extract_conj trms (if Data.force_move then false else fst) xs P of
+            SOME (xs, eq, P') => SOME (xs, eq, Data.conj $ P' $ Q)
+          | NONE =>
+              (case extract_conj (P :: trms)
+                    (if Data.force_move then false else (fst andalso mover)) xs Q
+               of
+                SOME (xs, eq, Q') => SOME (xs, eq, Data.conj $ P $ Q')
+              | NONE => NONE))
+      end);
+(* XXX This is not regularized with respect to term context *)
+fun extract_imp fst xs t =
+  (case Data.dest_imp t of
+    NONE => NONE
+  | SOME (P, Q) =>
+      if Data.move xs P [] then (if fst then NONE else SOME (xs, P, Q))
+      else
+        (case extract_conj [] false xs P of
+          SOME (xs, eq, P') => SOME (xs, eq, Data.imp $ P' $ Q)
+        | NONE =>
+            (case extract_imp false xs Q of
+              NONE => NONE
+            | SOME (xs, eq, Q') => SOME (xs, eq, Data.imp $ P $ Q'))));
+
+fun extract_quant extract q =
+  let
+    fun exqu xs ((qC as Const (qa, _)) $ Abs (x, T, Q)) =
+          if qa = q then exqu ((qC, x, T) :: xs) Q else NONE
+      | exqu xs P = extract (if Data.force_move then null xs else true) xs P
+  in exqu [] end;
+
+fun prove_conv ctxt tu tac =
+  let
+    val (goal, ctxt') =
+      yield_singleton (Variable.import_terms true) (Logic.mk_equals tu) ctxt;
+    val thm =
+      Goal.prove ctxt' [] [] goal
+        (fn {context = ctxt'', ...} =>
+          resolve_tac ctxt'' [Data.iff_reflection] 1 THEN tac ctxt'');
+  in singleton (Variable.export ctxt' ctxt) thm end;
+
+fun maybe_tac tac = if Data.rotate then tac else K all_tac;
+
+fun qcomm_tac ctxt qcomm qI i =
+  REPEAT_DETERM (maybe_tac (resolve_tac ctxt [qcomm]) i THEN resolve_tac ctxt [qI] i);
+
+(* Proves (? x0..xn. ... & x0 = t & ...) = (? x1..xn x0. x0 = t & ... & ...)
+   Better: instantiate exI
+*)
+local
+  val excomm = Data.ex_comm RS Data.iff_trans;
+in
+  fun prove_rotate_ex_tac ctxt i = qcomm_tac ctxt excomm Data.iff_exI i
+  fun prove_one_point_ex_tac ctxt =
+    prove_rotate_ex_tac ctxt 1 THEN resolve_tac ctxt [Data.iffI] 1 THEN
+    ALLGOALS
+      (EVERY' [maybe_tac (eresolve_tac ctxt [Data.exE]),
+        REPEAT_DETERM o eresolve_tac ctxt [Data.conjE],
+        maybe_tac (resolve_tac ctxt [Data.exI]),
+        DEPTH_SOLVE_1 o ares_tac ctxt [Data.conjI]])
+end;
+
+(* Proves (! x0..xn. (... & x0 = t & ...) --> P x0) =
+          (! x1..xn x0. x0 = t --> (... & ...) --> P x0)
+*)
+local
+  fun tac ctxt =
+    SELECT_GOAL
+      (EVERY1 [REPEAT o dresolve_tac ctxt [Data.uncurry],
+        REPEAT o resolve_tac ctxt [Data.impI],
+        eresolve_tac ctxt [Data.mp],
+        REPEAT o eresolve_tac ctxt [Data.conjE],
+        REPEAT o ares_tac ctxt [Data.conjI]]);
+  val allcomm = Data.all_comm RS Data.iff_trans;
+in
+  fun prove_one_point_all_tac ctxt =
+    EVERY1 [qcomm_tac ctxt allcomm Data.iff_allI,
+      resolve_tac ctxt [Data.iff_allI],
+      resolve_tac ctxt [Data.iffI], tac ctxt, tac ctxt];
+end
+
+(* Proves (! x0..xn. (... & x0 = t & ...) --> P x0) =
+          (! x1..xn x0. x0 = t --> (... & ...) --> P x0)
+*)
+local
+  val allcomm = Data.all_comm RS Data.iff_trans;
+in
+  fun prove_one_point_all_tac2 ctxt =
+    EVERY1 [qcomm_tac ctxt allcomm Data.iff_allI,
+      resolve_tac ctxt [Data.iff_allI],
+      resolve_tac ctxt [Data.iffI], blast_tac ctxt, blast_tac ctxt];
+end
+
+fun renumber l u (Bound i) =
+      Bound (if i < l orelse i > u then i else if i = u then l else i + 1)
+  | renumber l u (s $ t) = renumber l u s $ renumber l u t
+  | renumber l u (Abs (x, T, t)) = Abs (x, T, renumber (l + 1) (u + 1) t)
+  | renumber _ _ atom = atom;
+
+fun quantify qC x T xs P =
+  let
+    fun quant [] P = P
+      | quant ((qC, x, T) :: xs) P = quant xs (qC $ Abs (x, T, P));
+    val n = length xs;
+    val Q = if n = 0 then P else renumber 0 n P;
+  in if Data.rotate then quant xs (qC $ Abs (x, T, Q)) else qC $ Abs (x, T, quant xs P) end;
+
+fun rearrange_all ctxt ct =
+  (case Thm.term_of ct of
+    F as (all as Const (q, _)) $ Abs (x, T, P) =>
+      (case extract_quant extract_imp q P of
+        NONE => NONE
+      | SOME (xs, eq, Q) =>
+          let val R = quantify all x T xs (Data.imp $ eq $ Q)
+          in SOME (prove_conv ctxt (F, R) prove_one_point_all_tac) end)
+  | _ => NONE);
+
+fun rotate_all ctxt ct =
+  let
+    fun extract fst xs P =
+      if fst then NONE else SOME (xs, P, P)
+    in
+  (case strip_prop (Thm.term_of ct) of
+    F as (ex as Const (q, _)) $ Abs (x, T, P) =>
+      (case extract_quant extract q P of
+        NONE => NONE
+      | SOME (xs, _, Q) =>
+          let val R = quantify ex x T xs Q
+          in SOME (prove_conv ctxt (F, R) prove_one_point_all_tac2) end)
+  | _ => NONE) end;
+
+fun rearrange_ball tac ctxt ct =
+  (case Thm.term_of ct of
+    F as Ball $ A $ Abs (x, T, P) =>
+      (case extract_imp true [] P of
+        NONE => NONE
+      | SOME (xs, eq, Q) =>
+          if not (null xs) then NONE
+          else
+            let val R = Data.imp $ eq $ Q
+            in SOME (prove_conv ctxt (F, Ball $ A $ Abs (x, T, R)) tac) end)
+  | _ => NONE);
+
+fun rearrange_ex' ctxt trm =
+  (case strip_prop trm of
+    F as (ex as Const (q, _)) $ Abs (x, T, P) =>
+      (case extract_quant (extract_conj []) q P of
+        NONE => NONE
+      | SOME (xs, eq, Q) =>
+          let val R = quantify ex x T xs (Data.conj $ eq $ Q)
+          in SOME (prove_conv ctxt (F, R) prove_one_point_ex_tac) end)
+  | _ => NONE);
+
+fun rearrange_ex ctxt = rearrange_ex' ctxt o Thm.term_of
+
+fun rotate_ex ctxt ct =
+  let
+    fun extract fst xs P =
+      if fst then NONE else SOME (xs, P, P)
+    in
+  (case strip_prop (Thm.term_of ct) of
+    F as (ex as Const (q, _)) $ Abs (x, T, P) =>
+      (case extract_quant extract q P of
+        NONE => NONE
+      | SOME (xs, _, Q) =>
+          let val R = quantify ex x T xs Q
+          in SOME (prove_conv ctxt (F, R) prove_one_point_ex_tac) end)
+  | _ => NONE) end;
+
+fun miniscope_ex ctxt ct =
+  let
+    fun extract fst xs t =
+      case Data.dest_conj t of
+        SOME (P, _) => if Data.move xs P [] andalso not fst then SOME (xs, t, t) else NONE
+      | NONE => NONE
+    in
+  (case strip_prop (Thm.term_of ct) of
+    F as (ex as Const (q, _)) $ Abs (x, T, P) =>
+      (case extract_quant extract q P of
+        NONE => NONE
+      | SOME (xs, _, Q) =>
+          let val R = quantify ex x T xs Q
+          in SOME (prove_conv ctxt (F, R) prove_one_point_ex_tac) end)
+  | _ => NONE) end;
+
+fun rearrange_bex tac ctxt ct =
+  (case Thm.term_of ct of
+    F as Bex $ A $ Abs (x, T, P) =>
+      (case extract_conj [] true [] P of
+        NONE => NONE
+      | SOME (xs, eq, Q) =>
+          if not (null xs) then NONE
+          else SOME (prove_conv ctxt (F, Bex $ A $ Abs (x, T, Data.conj $ eq $ Q)) tac))
+  | _ => NONE);
+
+fun rearrange_Collect tac ctxt ct =
+  (case Thm.term_of ct of
+    F as Collect $ Abs (x, T, P) =>
+      (case extract_conj [] true [] P of
+        NONE => NONE
+      | SOME (_, eq, Q) =>
+          let val R = Collect $ Abs (x, T, Data.conj $ eq $ Q)
+          in SOME (prove_conv ctxt (F, R) tac) end)
+  | _ => NONE);
+
+end;
+
+structure Quantifier1 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move xs eq _ =
+  (case dest_eq eq of
+    SOME (s, t) =>
+      let val n = length xs in
+        s = Bound n andalso not (loose_bvar1 (t, n)) orelse
+        t = Bound n andalso not (loose_bvar1 (s, n))
+      end
+  | NONE => false);
+  val force_move = true
+  val rotate = true
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+(* loose_bvar2(t,k) iff t contains a 'loose' bound variable referring to
+   a level below k. *)
+fun loose_bvar2(Bound i,k) = i < k
+  | loose_bvar2(f$t, k) = loose_bvar2(f,k) orelse loose_bvar2(t,k)
+  | loose_bvar2(Abs(_,_,t),k) = loose_bvar2(t,k+1)
+  | loose_bvar2 _ = false;
+
+structure Quantifier2 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move xs t _ =  
+      let val n = length xs in
+        loose_bvar1 (t, n) andalso not (loose_bvar2 (t, n))
+      end
+  val force_move = false
+  val rotate = false
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+structure Quantifier3 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move xs t _ =  
+      let val n = length xs in
+        loose_bvar1 (t, n) andalso not (loose_bvar (t, n + 1))
+      end
+  val force_move = false
+  val rotate = false
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+signature Int_Param =
+  sig
+    val x : int
+  end;
+
+fun is_conj (Const(@{const_name HOL.conj},_) $ _ $ _) = true
+    | is_conj _ = false;
+
+functor Quantifier4 (to_move: Int_Param) = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move _ P trms = length trms + 1 = to_move.x andalso not (is_conj P)
+  val force_move = true
+  val rotate = false
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+structure Quantifier5 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move _ t _ = is_conj t
+  val force_move = true
+  val rotate = false
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+structure Quantifier6 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move xs t _ =  
+      let val n = length xs in
+        not (loose_bvar1 (t, n))
+      end
+  val force_move = true
+  val rotate = true
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+structure Quantifier7 = Quantifier
+(
+  (*abstract syntax*)
+  fun dest_eq (Const(@{const_name HOL.eq},_) $ s $ t) = SOME (s, t)
+    | dest_eq _ = NONE;
+  fun dest_conj (Const(@{const_name HOL.conj},_) $ s $ t) = SOME (s, t)
+    | dest_conj _ = NONE;
+  fun dest_imp (Const(@{const_name HOL.implies},_) $ s $ t) = SOME (s, t)
+    | dest_imp _ = NONE;
+  val conj = HOLogic.conj
+  val imp  = HOLogic.imp
+  fun move xs t _ =  
+      let val n = length xs in
+        not (loose_bvar1 (t, n))
+      end
+  val force_move = true
+  val rotate = true
+  (*rules*)
+  val iff_reflection = @{thm eq_reflection}
+  val iffI = @{thm iffI}
+  val iff_trans = @{thm trans}
+  val conjI= @{thm conjI}
+  val conjE= @{thm conjE}
+  val impI = @{thm impI}
+  val mp   = @{thm mp}
+  val uncurry = @{thm uncurry}
+  val exI  = @{thm exI}
+  val exE  = @{thm exE}
+  val iff_allI = @{thm iff_allI}
+  val iff_exI = @{thm iff_exI}
+  val all_comm = @{thm all_comm}
+  val ex_comm = @{thm ex_comm}
+);
+
+\<close>
+
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a \<in> A"}\<close>
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a = b"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. a = b \<and> c < n"}\<close>
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<exists> a c. a = b \<and> c < n"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a = b"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. a < n \<and> a = b"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. a < n \<and> c < n \<and> a = b"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a > c"}\<close>
+ML_val \<open>Quantifier3.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a > c"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> a > b"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "\<exists> a c. c < n \<and> (P a c \<and> a > b) \<and> Q c"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "finite {(a, c) | a c. c < n \<and> a \<in> A}"}\<close>
+ML_val \<open>Quantifier2.rearrange_ex @{context} @{cterm "finite {t. \<exists> a c. a \<in> A \<and> c < n \<and> t = (a,c)}"}\<close>
+ML_val \<open>Quantifier1.rotate_ex @{context} @{cterm "\<exists> a c. c < n \<and> a > b"}\<close>
+ML_val \<open>Quantifier1.rotate_ex @{context} @{cterm "\<exists> a c d. c < n \<and> a > b \<and> P d"}\<close>
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<exists> a c. a < n \<and> c = b"}\<close>
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<forall> a. \<exists> c. a < n \<and> c = b"}\<close>
+ML_val \<open>Quantifier1.rearrange_ex @{context} @{cterm "\<forall> a c. a < n \<and> c = b"}\<close>
+ML_val \<open>Quantifier6.rearrange_ex @{context} @{cterm "\<exists> a b c. a < n \<and> b < 3 \<and> b > c"}\<close>
+ML_val \<open>Quantifier6.rearrange_ex @{context} @{cterm "\<exists>b c. a < n \<and> b < 3 \<and> b > c"}\<close>
+ML_val \<open>Quantifier7.miniscope_ex @{context} @{cterm "\<exists> a b c. a < n \<and> b < 3 \<and> b > c"}\<close>
+ML_val \<open>Quantifier7.miniscope_ex @{context} @{cterm "\<exists>b c. a < n \<and> b < 3 \<and> b > c"}\<close>
+
+simproc_setup ex_reorder ("\<exists>x. P x") = \<open>fn _ => Quantifier2.rearrange_ex\<close>
+declare [[simproc del: ex_reorder]]
+simproc_setup ex_reorder2 ("\<exists>x. P x") = \<open>fn _ => Quantifier3.rearrange_ex\<close>
+declare [[simproc del: ex_reorder2]]
+simproc_setup ex_reorder3 ("\<exists>x. P x") = \<open>fn _ => Quantifier6.rearrange_ex\<close>
+declare [[simproc del: ex_reorder3]]
+simproc_setup ex_reorder4 ("\<exists>x. P x") = \<open>fn _ => Quantifier7.miniscope_ex\<close>
+declare [[simproc del: ex_reorder4]]
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {(a, c). c < n \<and> a \<in> A}"
+  using assms
+    using [[simproc add: finite_Collect]]
+    by simp
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {(a, c) | a c. c < n \<and> a \<in> A}"
+  using assms
+  using [[simproc add: ex_reorder]]
+    by simp
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. a \<in> A \<and> c < n \<and> t = (a,c)}"
+  apply simp
+  using assms apply simp
+  oops
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. (t = (a,c) \<and> c < n) \<and> a \<in> A}"
+      using [[simproc add: ex_reorder]]
+  using [[simp_trace]] apply simp
+  using assms by simp
+
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. (t = (a,c) \<and> a \<in> A) \<and> c < n}"
+  using [[simp_trace]] apply (simp del: Product_Type.Collect_case_prod)
+  using assms by simp
+  
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. (a \<in> A \<and> t = (a,c)) \<and> c < n}"
+  using [[simp_trace]] apply simp
+  using assms by simp
+  
+lemma
+  fixes n :: nat
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. a \<in> A \<and> c < n \<and> t = (a,c)}"
+  using [[simp_trace]] apply simp
+  using assms by simp
+  
+lemma
+  assumes A: "finite A"
+  shows "finite {t. \<exists> a c. a \<in> A \<and> P c \<and> t = (a,c)}"
+  using [[simp_trace]] apply simp
+  using assms apply simp
+  oops
+
+ML \<open>
+  fun rotate_quant reorder_thm n ctxt =
+    let
+      fun subst j =
+        if j > n then K all_tac else
+          (
+            EqSubst.eqsubst_tac ctxt [j] [reorder_thm]
+          ) THEN' subst (j + 1)
+    in subst 1 end;
+\<close>
+
+ML \<open>
+  fun rotate_ex_tac ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        case Quantifier1.rotate_ex ctxt concl of
+          NONE => no_tac
+        | SOME thm => rewrite_goals_tac ctxt [thm]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+\<close>
+
+ML \<open>
+  fun rotate_all_tac ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        case Quantifier1.rotate_all ctxt concl of
+          NONE => no_tac
+        | SOME thm => rewrite_goals_tac ctxt [thm]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+\<close>
+
+ML \<open>
+  fun rearrange_ex_tac ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        case Quantifier2.rearrange_ex ctxt concl of
+          NONE => no_tac
+        | SOME thm => rewrite_goals_tac ctxt [thm]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+\<close>
+
+ML \<open>
+  fun rearrange_ex_tac2 ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        case Quantifier3.rearrange_ex ctxt concl of
+          NONE => no_tac
+        | SOME thm => rewrite_goals_tac ctxt [thm]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+\<close>
+
+ML \<open>
+
+  fun strip_fin (Const (@{const_name "finite"}, _) $ (Const (@{const_name "Collect"}, _) $ t)) = t
+    | strip_fin t = t
+
+  fun wrap_fin tac ctxt = tac ctxt o strip_fin
+
+  structure Quant2 = Quantifier4(val x = 2);
+  structure Quant3 = Quantifier4(val x = 3);
+  structure Quant4 = Quantifier4(val x = 4);
+  structure Quant5 = Quantifier4(val x = 5);
+
+  fun rearrange_ex_fixed_n rearrange_n ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        case rearrange_n ctxt concl of
+          NONE => no_tac
+        | SOME thm => rewrite_goals_tac ctxt [thm, @{thm HOL.conj_assoc} RS @{thm HOL.eq_reflection}]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+  
+  val rearrange_ex_fixed_2 = rearrange_ex_fixed_n Quant2.rearrange_ex;
+  val rearrange_ex_fixed_3 = rearrange_ex_fixed_n Quant3.rearrange_ex;
+  val rearrange_ex_fixed_4 = rearrange_ex_fixed_n Quant4.rearrange_ex;
+  val rearrange_ex_fixed_5 = rearrange_ex_fixed_n Quant5.rearrange_ex;
+
+  (* val defer_ex = rearrange_ex_fixed_n (wrap_fin Quantifier5.rearrange_ex); *)
+
+  fun CONV conv ctxt =
+    let
+      fun foc_tac {concl, ...} =
+        rewrite_goals_tac ctxt [conv ctxt concl]
+    in
+      Subgoal.FOCUS foc_tac ctxt
+    end;
+
+  fun mk_conv f ctxt ct =
+    case (f ctxt ct) of
+      SOME thm => thm
+    | _ => raise CTERM ("no conversion", [])
+
+  fun success_conv cv ct =
+    let
+      val eq = cv ct
+    in
+      if Thm.is_reflexive eq then raise CTERM ("no conversion", []) else eq
+    end
+
+  fun mk_conv' f ctxt ct = the_default (Thm.reflexive ct) (f ctxt ct)
+  val assoc_conv = Conv.rewr_conv (@{thm HOL.conj_assoc} RS @{thm HOL.eq_reflection})
+  val comm_conv = Conv.rewr_conv (@{thm HOL.conj_commute} RS @{thm HOL.eq_reflection})
+  fun wrap_conv f ctxt =
+    success_conv (
+      Conv.top_sweep_conv (fn ctxt => mk_conv f ctxt then_conv Conv.repeat_conv assoc_conv) ctxt
+    )
+  fun mk_tac conv ctxt = CONVERSION (Conv.concl_conv ~1 (Object_Logic.judgment_conv ctxt (conv ctxt)))
+
+  val defer_conv = mk_conv Quantifier5.rearrange_ex
+  val conv = wrap_conv Quantifier5.rearrange_ex
+  fun defer_ex_tac ctxt i =
+    CHANGED (mk_tac (fn ctxt => wrap_conv Quantifier5.rearrange_ex ctxt else_conv Conv.top_sweep_conv (K comm_conv) ctxt) ctxt i)
+  val mini_ex_tac = mk_tac (wrap_conv Quantifier6.rearrange_ex)
+  val mini_ex_tac2 = mk_tac (wrap_conv Quantifier7.miniscope_ex)
+
+  val rearrange_ex_fixed_2 = mk_tac (wrap_conv Quant2.rearrange_ex);
+  val rearrange_ex_fixed_3 = mk_tac (wrap_conv Quant3.rearrange_ex);
+  val rearrange_ex_fixed_4 = mk_tac (wrap_conv Quant4.rearrange_ex);
+  val rearrange_ex_fixed_5 = mk_tac (wrap_conv Quant5.rearrange_ex);
+\<close>
+
+ML_val \<open>defer_conv @{context} @{cterm "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>assoc_conv @{cterm "(a < 1 \<and> b < 2) \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>Conv.binder_conv (K assoc_conv) @{context} @{cterm "\<exists> a. (a < 1 \<and> b < 2) \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>Conv.top_sweep_conv (K assoc_conv) @{context} @{cterm "\<exists> a b c d. (a < 1 \<and> b < 2) \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>Conv.bottom_conv (K (Conv.try_conv assoc_conv)) @{context} @{cterm "\<exists> a b c d. (a < 1 \<and> b < 2) \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>Conv.every_conv [defer_conv @{context}, Conv.try_conv assoc_conv] @{cterm "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>conv @{context} @{cterm "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"}\<close>
+ML_val \<open>conv @{context} @{cterm "finite {t. \<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4}"}\<close>
+ML_val \<open>conv @{context} @{cterm "\<exists>a b c d. d = 4 \<and> c = 3 \<and> b < 2 \<and> a < 1"}\<close>
+
+ML_val \<open>Quantifier1.rotate_ex @{context} @{cterm "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"}\<close>
+
+ML_val \<open>Quantifier1.rotate_all @{context} @{cterm "\<forall> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"}\<close>
+
+lemma
+  "\<forall> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  apply (tactic \<open>rotate_all_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_all_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_all_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_all_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_all_tac @{context} 1\<close>)
+  oops
+
+lemma
+  "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c = 3 \<and> d = 4"
+  apply (tactic \<open>rearrange_ex_fixed_2 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_3 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_4 @{context} 1\<close>)
+  apply (tactic \<open>defer_ex_tac @{context} 1\<close>)
+  apply (subst conj_assoc)+
+  oops
+
+lemma
+  "finite {t. \<exists> a b c d. a < 1 \<and> b < 2 \<and> c = 3 \<and> d = 4}"
+  apply (tactic \<open>rearrange_ex_fixed_2 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_3 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_4 @{context} 1\<close>)
+  apply (tactic \<open>defer_ex_tac @{context} 1\<close>)
+  oops
+  
+lemma
+  "finite S \<Longrightarrow> finite {t. \<exists> a b c d. a < 1 \<and> b < 2 \<and> c = 3 \<and> d = 4}"
+  apply (tactic \<open>rearrange_ex_fixed_2 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_3 @{context} 1\<close>)
+  apply (tactic \<open>rearrange_ex_fixed_4 @{context} 1\<close>)
+  apply (tactic \<open>defer_ex_tac @{context} 1\<close>, simp only: conj_assoc)
+  oops
+
+lemma
+  "finite S \<Longrightarrow> finite {t. \<exists> a b c d. P a b d \<and> c > 3}"
+  apply (tactic \<open>defer_ex_tac @{context} 1\<close>)
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply (simp only: ex_simps)
+  oops
+
+lemma
+  "\<exists> a b c d. d < 4 \<and> a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  using [[simproc add: ex_reorder3]]
+  apply simp
+  oops
+
+lemma
+  "\<exists> a b c d. d < 4 \<and> a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  using [[simproc add: ex_reorder4]]
+  apply simp
+  oops
+
+lemma
+  "\<exists> a b c d. d < 4 \<and> a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  oops
+
+lemma
+  "\<exists> a b c d. d < 4 \<and> a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  apply (tactic \<open>mini_ex_tac2 @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac2 @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac2 @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  apply (tactic \<open>mini_ex_tac @{context} 1\<close>)
+  apply simp
+  oops
+  
+lemma
+  "\<exists> a c b d. d < 4 \<and> a < 1 \<and> b < 2 \<and> c < 3"
+  apply simp
+  oops
+  
+lemma
+  "\<exists> a b c d. a < 1 \<and> b < 2 \<and> c < 3 \<and> d < 4"
+  apply (tactic \<open>rotate_ex_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_ex_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_ex_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_ex_tac @{context} 1\<close>)
+  apply (tactic \<open>rotate_ex_tac @{context} 1\<close>)
+  oops
+
+lemma
+  "\<exists> a b c d. b < 2 \<and> c < 3 \<and> d < 4 \<and> a < 1"
+  apply (tactic \<open>rearrange_ex_tac @{context} 1\<close>)
+  oops
+
+lemma
+  "\<exists> a b c d. b < 2 \<and> c < 3 \<and> d < 4 \<and> a < c"
+  apply (tactic \<open>rearrange_ex_tac2 @{context} 1\<close>; simp only: conj_assoc)+
+  oops
+
+lemma
+  "\<exists> a b c d. b < 2 \<and> c < 3 \<and> d < 4 \<and> a < c"
+  apply (tactic \<open>rearrange_ex_tac2 @{context} 1\<close>; simp del: ex_simps)+
+  oops
+
+lemma
+  "\<exists> a b c d. b < 2 \<and> c < 3 \<and> d < 4 \<and> a < c"
+  apply (tactic \<open>rearrange_ex_tac2 @{context} 1\<close>)
+  apply (simp del: ex_simps)
+  apply (tactic \<open>rearrange_ex_tac2 @{context} 1\<close>)
+  apply (simp del: ex_simps)
+  apply (tactic \<open>rearrange_ex_tac2 @{context} 1\<close>)
+  using [[simp_trace]]
+  apply (simp del: ex_simps)
+  oops
+
+lemma finite_Collect_bounded_ex_4:
+  assumes "finite {(a,b,c,d) . P a b c d}"
+  shows
+    "finite {x. \<exists>a b c d. P a b c d \<and> Q x a b c d}
+    \<longleftrightarrow> (\<forall> a b c d. P a b c d \<longrightarrow> finite {x. Q x a b c d})"
+proof -
+  have *:
+    "{x. \<exists>a b c d. P a b c d \<and> Q x a b c d}
+    = {x. \<exists> t. t \<in> {(a,b,c,d). P a b c d} \<and> (\<exists>a b c d. t = (a, b, c, d) \<and> Q x a b c d)}"
+    by simp
+  show ?thesis apply (subst *)
+    apply (subst finite_Collect_bounded_ex)
+    using assms by simp+
+oops
+  
+lemma finite_Collect_bounded_ex_4':
+  assumes "finite {(a,b,c,d) | a b c d. P a b c d}"
+  shows
+    "finite {x. \<exists>a b c d. P a b c d \<and> Q x a b c d}
+    \<longleftrightarrow> (\<forall> a b c d. P a b c d \<longrightarrow> finite {x. Q x a b c d})"
+proof -
+  have *:
+    "{x. \<exists>a b c d. P a b c d \<and> Q x a b c d}
+    = {x. \<exists> t. t \<in> {(a,b,c,d) | a b c d. P a b c d} \<and> (\<exists>a b c d. t = (a, b, c, d) \<and> Q x a b c d)}"
+    by simp
+  show ?thesis apply (subst *)
+    apply (subst finite_Collect_bounded_ex)
+    using assms by simp+
+qed
+
+lemma finite_Collect_bounded_ex_2 [simp]:
+  assumes "finite {(a,b). P a b}"
+  shows
+    "finite {x. \<exists>a b. P a b \<and> Q x a b}
+    \<longleftrightarrow> (\<forall> a b. P a b \<longrightarrow> finite {x. Q x a b})"
+  using assms finite_Collect_bounded_ex[OF assms, where Q = "\<lambda> x. \<lambda> (a, b). Q x a b"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_3 [simp]:
+  assumes "finite {(a,b,c) . P a b c}"
+  shows
+    "finite {x. \<exists>a b c. P a b c \<and> Q x a b c}
+    \<longleftrightarrow> (\<forall> a b c. P a b c \<longrightarrow> finite {x. Q x a b c})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c). Q x a b c"]
+  by clarsimp
+
+lemma finite_Collect_bounded_ex_4 [simp]:
+  assumes "finite {(a,b,c,d) . P a b c d}"
+  shows
+    "finite {x. \<exists>a b c d. P a b c d \<and> Q x a b c d}
+    \<longleftrightarrow> (\<forall> a b c d. P a b c d \<longrightarrow> finite {x. Q x a b c d})"
+  using assms finite_Collect_bounded_ex[OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d). Q x a b c d"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_5 [simp]:
+  assumes "finite {(a,b,c,d,e) . P a b c d e}"
+  shows
+    "finite {x. \<exists>a b c d e. P a b c d e \<and> Q x a b c d e}
+    \<longleftrightarrow> (\<forall> a b c d e. P a b c d e \<longrightarrow> finite {x. Q x a b c d e})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e). Q x a b c d e"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_6 [simp]:
+  assumes "finite {(a,b,c,d,e,f) . P a b c d e f}"
+  shows
+    "finite {x. \<exists>a b c d e f. P a b c d e f \<and> Q x a b c d e f}
+    \<longleftrightarrow> (\<forall> a b c d e f. P a b c d e f \<longrightarrow> finite {x. Q x a b c d e f})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e, f). Q x a b c d e f"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_7 [simp]:
+  assumes "finite {(a,b,c,d,e,f,g) . P a b c d e f g}"
+  shows
+    "finite {x. \<exists>a b c d e f g. P a b c d e f g \<and> Q x a b c d e f g}
+    \<longleftrightarrow> (\<forall> a b c d e f g. P a b c d e f g \<longrightarrow> finite {x. Q x a b c d e f g})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e, f, g). Q x a b c d e f g"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_8 [simp]:
+  assumes "finite {(a,b,c,d,e,f,g,h) . P a b c d e f g h}"
+  shows
+    "finite {x. \<exists>a b c d e f g h. P a b c d e f g h \<and> Q x a b c d e f g h}
+    \<longleftrightarrow> (\<forall> a b c d e f g h. P a b c d e f g h \<longrightarrow> finite {x. Q x a b c d e f g h})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e, f, g, h). Q x a b c d e f g h"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_9 [simp]:
+  assumes "finite {(a,b,c,d,e,f,g,h,i) . P a b c d e f g h i}"
+  shows
+    "finite {x. \<exists>a b c d e f g h i. P a b c d e f g h i \<and> Q x a b c d e f g h i}
+    \<longleftrightarrow> (\<forall> a b c d e f g h i. P a b c d e f g h i \<longrightarrow> finite {x. Q x a b c d e f g h i})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e, f, g, h, i). Q x a b c d e f g h i"]
+  by clarsimp (* force, simp *)
+
+lemma finite_Collect_bounded_ex_10 [simp]:
+  assumes "finite {(a,b,c,d,e,f,g,h,i,j) . P a b c d e f g h i j}"
+  shows
+    "finite {x. \<exists>a b c d e f g h i j. P a b c d e f g h i j \<and> Q x a b c d e f g h i j}
+    \<longleftrightarrow> (\<forall> a b c d e f g h i j. P a b c d e f g h i j \<longrightarrow> finite {x. Q x a b c d e f g h i j})"
+  using assms finite_Collect_bounded_ex
+    [OF assms, where Q = "\<lambda> x. \<lambda> (a, b, c, d, e, f, g, h, i, j). Q x a b c d e f g h i j"]
+  by clarsimp (* force, simp *)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Subsumption_Graphs.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Subsumption_Graphs.thy
@@ -0,0 +1,1226 @@
+theory Subsumption_Graphs
+  imports
+    Timed_Automata.Graphs
+    Timed_Automata.More_List1
+begin
+
+section \<open>Subsumption Graphs\<close>
+
+subsection \<open>Preliminaries\<close>
+
+subsubsection \<open>Transitive Closure\<close>
+
+context
+  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes R_trans[intro]: "\<And> x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
+begin
+
+lemma rtranclp_transitive_compress1: "R a c" if "R a b" "R** b c"
+  using that(2,1) by induction auto
+
+lemma rtranclp_transitive_compress2: "R a c" if "R** a b" "R b c"
+  using that by induction auto
+
+end (* Transitivity *)
+
+(* XXX Move *)
+lemma rtranclp_ev_induct[consumes 1, case_names irrefl trans step]:
+  fixes P :: "'a \<Rightarrow> bool" and R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes reachable_finite: "finite {x. R** a x}"
+  assumes R_irrefl: "\<And> x. \<not> R x x" and R_trans[intro]: "\<And> x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
+  assumes step: "\<And> x. R** a x \<Longrightarrow> P x \<or> (\<exists> y. R x y)"
+  shows "\<exists> x. P x \<and> R** a x"
+proof -
+  let ?S = "{y. R** a y}"
+  from reachable_finite have "finite ?S"
+    by auto
+  then have "\<exists> x \<in> ?S. P x"
+    using step
+  proof (induction ?S arbitrary: a rule: finite_psubset_induct)
+    case psubset
+    let ?S = "{y. R** a y}"
+    from psubset have "finite ?S" by auto
+    show ?case
+    proof (cases "?S = {}")
+      case True
+      then show ?thesis by auto
+    next
+      case False
+      then obtain y where "R** a y"
+        by auto
+      from psubset(3)[OF this] show ?thesis
+      proof
+        assume "P y"
+        with \<open>R** a y\<close> show ?thesis by auto
+      next
+        assume "\<exists> z. R y z"
+        then obtain z where "R y z" by safe
+        let ?T = "{y. R** z y}"
+        from \<open>R y z\<close> \<open>R** a y\<close> have "\<not> R** z a"
+          by (auto simp: R_irrefl dest!: rtranclp_transitive_compress2[of R, rotated])
+        then have "a \<notin> ?T" by auto
+        moreover have "?T \<subseteq> ?S"
+          using \<open>R** a y\<close> \<open>R y z\<close> by auto
+        ultimately have "?T \<subset> ?S"
+          by auto
+        have "P x \<or> Ex (R x)" if "R** z x" for x
+          using that \<open>R y z\<close> \<open>R** a y\<close> by (auto intro!: psubset.prems)
+        from psubset.hyps(2)[OF \<open>?T \<subset> ?S\<close> this] psubset.prems \<open>R y z\<close> \<open>R** a y\<close> obtain w
+          where "R** z w" "P w" by auto
+        with \<open>R** a y\<close> \<open>R y z\<close> have "R** a w" by auto
+        with \<open>P w\<close> show ?thesis by auto
+      qed
+    qed
+  qed
+  then show ?thesis by auto
+qed
+
+(* XXX Move *)
+lemma rtranclp_ev_induct2[consumes 2, case_names irrefl trans step]:
+  fixes P Q :: "'a \<Rightarrow> bool"
+  assumes Q_finite: "finite {x. Q x}" and Q_witness: "Q a"
+  assumes R_irrefl: "\<And> x. \<not> R x x" and R_trans[intro]: "\<And> x y z. R x y \<Longrightarrow> R y z \<Longrightarrow> R x z"
+  assumes step: "\<And> x. Q x \<Longrightarrow> P x \<or> (\<exists> y. R x y \<and> Q y)"
+  shows "\<exists> x. P x \<and> Q x \<and> R** a x"
+proof -
+  let ?R = "\<lambda> x y. R x y \<and> Q x \<and> Q y"
+  have [intro]: "R** a x" if "?R** a x" for x
+    using that by induction auto
+  have [intro]: "Q x" if "?R** a x" for x
+    using that \<open>Q a\<close> by (auto elim: rtranclp.cases)
+  have "{x. ?R** a x} \<subseteq> {x. Q x}" by auto
+  with \<open>finite _\<close> have "finite {x. ?R** a x}" by - (rule finite_subset)
+  then have "\<exists>x. P x \<and> ?R** a x"
+  proof (induction rule: rtranclp_ev_induct)
+    case prems: (step x)
+    with step[of x] show ?case by auto
+  qed (auto simp: R_irrefl)
+  then show ?thesis by auto
+qed
+
+
+subsection \<open>Definitions\<close>
+
+locale Subsumption_Graph_Pre_Defs =
+  ord less_eq less for less_eq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>" 50) and less (infix "\<prec>" 50) +
+  fixes E ::  "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>The full edge set\<close>
+begin
+
+sublocale Graph_Defs E .
+
+end
+
+
+locale Subsumption_Graph_Pre_Nodes_Defs = Subsumption_Graph_Pre_Defs +
+  fixes V :: "'a \<Rightarrow> bool"
+begin
+
+sublocale Subgraph_Node_Defs_Notation .
+
+end  (* Subsumption Graph Pre Nodes Defs *)
+
+
+(* XXX Merge with Worklist locales *)
+locale Subsumption_Graph_Defs = Subsumption_Graph_Pre_Defs +
+  fixes s0 :: 'a \<comment> \<open>Start state\<close>
+  fixes RE :: "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>Subgraph of the graph given by the full edge set\<close>
+begin
+
+sublocale Graph_Start_Defs E s0 .
+
+sublocale G: Graph_Start_Defs RE s0 .
+
+sublocale G': Graph_Start_Defs "\<lambda> x y. RE x y \<or> (x \<prec> y \<and> G.reachable y)" s0 .
+
+abbreviation G'_E    ("_ \<rightarrow>G' _" [100, 100] 40) where
+  "G'_E x y \<equiv> RE x y \<or> (x \<prec> y \<and> G.reachable y)"
+
+notation RE          ("_ \<rightarrow>G _"   [100, 100] 40)
+
+notation G.reaches   ("_ \<rightarrow>G* _"  [100, 100] 40)
+
+notation G.reaches1  ("_ \<rightarrow>G+ _"  [100, 100] 40)
+
+notation G'.reaches  ("_ \<rightarrow>G*'' _" [100, 100] 40)
+
+notation G'.reaches1 ("_ \<rightarrow>G+'' _" [100, 100] 40)
+
+end (* Subsumption Graph Defs *)
+
+locale Subsumption_Graph_Pre = Subsumption_Graph_Defs + preorder less_eq less +
+  assumes mono:
+    "a \<preceq> b \<Longrightarrow> E a a' \<Longrightarrow> reachable a \<Longrightarrow> reachable b \<Longrightarrow> \<exists> b'. E b b' \<and> a' \<preceq> b'"
+begin
+
+lemmas preorder_intros = order_trans less_trans less_imp_le
+
+end (* Subsumption Graph Pre *)
+
+
+locale Subsumption_Graph_Pre_Nodes = Subsumption_Graph_Pre_Nodes_Defs + preorder less_eq less +
+  assumes mono:
+    "a \<preceq> b \<Longrightarrow> a \<rightarrow> a' \<Longrightarrow> V a \<Longrightarrow> V b \<Longrightarrow> \<exists> b'. b \<rightarrow> b' \<and> a' \<preceq> b'"
+begin
+
+lemmas preorder_intros = order_trans less_trans less_imp_le
+
+end (* Subsumption Graph Pre Nodes *)
+
+text \<open>
+  This is sufficient to show that if \<open>\<rightarrow>G\<close> cannot reach an accepting state,
+  then \<open>\<rightarrow>\<close> cannot either.
+\<close>
+locale Reachability_Compatible_Subsumption_Graph_Pre =
+  Subsumption_Graph_Defs + preorder less_eq less +
+  assumes mono:
+    "a \<preceq> b \<Longrightarrow> E a a' \<Longrightarrow> reachable a \<or> G.reachable a \<Longrightarrow> reachable b \<or> G.reachable b
+    \<Longrightarrow> \<exists> b'. E b b' \<and> a' \<preceq> b'"
+  assumes reachability_compatible:
+    "\<forall> s. G.reachable s \<longrightarrow> (\<forall> s'. E s s' \<longrightarrow> RE s s') \<or> (\<exists> t. s \<prec> t \<and> G.reachable t)"
+  assumes finite_reachable: "finite {a. G.reachable a}"
+
+locale Reachability_Compatible_Subsumption_Graph =
+  Subsumption_Graph_Defs + Subsumption_Graph_Pre +
+  assumes reachability_compatible:
+    "\<forall> s. G.reachable s \<longrightarrow> (\<forall> s'. E s s' \<longrightarrow> RE s s') \<or> (\<exists> t. s \<prec> t \<and> G.reachable t)"
+  assumes subgraph: "\<forall> s s'. RE s s' \<longrightarrow> E s s'"
+  assumes finite_reachable: "finite {a. G.reachable a}"
+
+locale Subsumption_Graph_View_Defs = Subsumption_Graph_Defs +
+  fixes SE ::  "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>Subsumption edges\<close>
+    and covered :: "'a \<Rightarrow> bool"
+
+locale Reachability_Compatible_Subsumption_Graph_View =
+  Subsumption_Graph_View_Defs + Subsumption_Graph_Pre +
+  assumes reachability_compatible:
+    "\<forall> s. G.reachable s \<longrightarrow>
+      (if covered s then (\<exists> t. SE s t \<and> G.reachable t) else (\<forall> s'. E s s' \<longrightarrow> RE s s'))"
+  assumes subsumption: "\<forall> s'. SE s s' \<longrightarrow> s \<prec> s'"
+  assumes subgraph: "\<forall> s s'. RE s s' \<longrightarrow> E s s'"
+  assumes finite_reachable: "finite {a. G.reachable a}"
+begin
+
+sublocale Reachability_Compatible_Subsumption_Graph "(\<preceq>)" "(\<prec>)" E s0 RE
+proof unfold_locales
+  have "(\<forall>s'. E s s' \<longrightarrow> RE s s') \<or> (\<exists>t. s \<prec> t \<and> G.reachable t)" if "G.reachable s" for s
+    using that reachability_compatible subsumption by (cases "covered s"; fastforce)
+  then show "\<forall>s. G.reachable s \<longrightarrow> (\<forall>s'. E s s' \<longrightarrow> RE s s') \<or> (\<exists>t. s \<prec> t \<and> G.reachable t)"
+    by auto
+qed (use subgraph in \<open>auto intro: finite_reachable mono\<close>)
+
+end (* Reachability Compatible Subsumption Graph View *)
+
+locale Subsumption_Graph_Closure_View_Defs =
+  ord less_eq less for less_eq :: "'b \<Rightarrow> 'b \<Rightarrow> bool" (infix "\<preceq>" 50) and less (infix "\<prec>" 50) +
+  fixes E ::  "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>The full edge set\<close>
+    and s0 :: 'a                 \<comment> \<open>Start state\<close>
+  fixes RE :: "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>Subgraph of the graph given by the full edge set\<close>
+  fixes SE ::  "'a \<Rightarrow> 'a \<Rightarrow> bool" \<comment> \<open>Subsumption edges\<close>
+    and covered :: "'a \<Rightarrow> bool"
+  fixes closure :: "'a \<Rightarrow> 'b"
+  fixes P :: "'a \<Rightarrow> bool"
+  fixes Q :: "'a \<Rightarrow> bool"
+begin
+
+sublocale Graph_Start_Defs E s0 .
+
+sublocale G: Graph_Start_Defs RE s0 .
+
+end (* Subsumption Graph Closure View Defs *)
+
+locale Reachability_Compatible_Subsumption_Graph_Closure_View =
+  Subsumption_Graph_Closure_View_Defs +
+  preorder less_eq less +
+  assumes mono:
+    "closure a \<preceq> closure b \<Longrightarrow> E a a' \<Longrightarrow> P a \<Longrightarrow> P b \<Longrightarrow> \<exists> b'. E b b' \<and> closure a' \<preceq> closure b'"
+  assumes closure_eq:
+    "closure a = closure b \<Longrightarrow> E a a' \<Longrightarrow> P a \<Longrightarrow> P b \<Longrightarrow> \<exists> b'. E b b' \<and> closure a' = closure b'"
+  assumes reachability_compatible:
+    "\<forall> s. Q s \<longrightarrow> (if covered s then (\<exists> t. SE s t \<and> G.reachable t) else (\<forall> s'. E s s' \<longrightarrow> RE s s'))"
+  assumes subsumption: "\<forall> s'. SE s s' \<longrightarrow> closure s \<prec> closure s'"
+  assumes subgraph: "\<forall> s s'. RE s s' \<longrightarrow> E s s'"
+  assumes finite_closure: "finite (closure ` UNIV)"
+  assumes P_post: "a \<rightarrow> b \<Longrightarrow> P b"
+  assumes P_pre: "a \<rightarrow> b \<Longrightarrow> P a"
+  assumes P_s0: "P s0"
+  assumes Q_post: "RE a b \<Longrightarrow> Q b"
+  assumes Q_s0: "Q s0"
+begin
+
+definition close where "close e a b = (\<exists> x y. e x y \<and> a = closure x \<and> b = closure y)"
+
+lemma Simulation_close:
+  "Simulation A (close A) (\<lambda> a b. b = closure a)"
+  unfolding close_def by standard auto
+
+sublocale view: Reachability_Compatible_Subsumption_Graph
+  "(\<preceq>)" "(\<prec>)" "close E" "closure s0" "close RE"
+  supply [simp] = close_def
+  supply [intro] = P_pre P_post Q_post
+proof (standard, goal_cases)
+  case prems: (1 a b a')
+  then obtain x y where [simp]: "x \<rightarrow> y" "a = closure x" "a' = closure y"
+    by auto
+  then have "P x" "P y"
+    by blast+
+  from prems(4) P_s0 obtain x' where [simp]: "b = closure x'" "P x'"
+    unfolding Graph_Start_Defs.reachable_def by cases auto
+  from mono[OF \<open>_ \<preceq> _\<close>[simplified] \<open>x \<rightarrow> y\<close> \<open>P x\<close> \<open>P x'\<close>] obtain b' where
+    "x' \<rightarrow> b'" "closure y \<preceq> closure b'"
+    by auto
+  then show ?case
+    by auto
+next
+  case 2
+  interpret Simulation RE "close RE" "\<lambda> a b. b = closure a"
+    by (rule Simulation_close)
+  { fix x assume "Graph_Start_Defs.reachable (close RE) (closure s0) x"
+    then obtain x' where [simp]: "x = closure x'" "Q x'" "P x'"
+      using Q_s0 P_s0 subgraph unfolding Graph_Start_Defs.reachable_def by cases auto
+    have "(\<forall>s'. close E x s' \<longrightarrow> close RE x s')
+        \<or> (\<exists>t. x \<prec> t \<and> Graph_Start_Defs.reachable (close RE) (closure s0) t)"
+    proof (cases "covered x'")
+      case True
+      with reachability_compatible \<open>Q x'\<close> obtain t where "SE x' t" "G.reachable t"
+        by fastforce
+      then show ?thesis
+        using subsumption
+        by - (rule disjI2, auto dest: simulation_reaches simp: Graph_Start_Defs.reachable_def)
+    next
+      case False
+      with reachability_compatible \<open>Q x'\<close> have "\<forall>s'. x' \<rightarrow> s' \<longrightarrow> RE x' s'"
+        by auto
+      then show ?thesis
+        unfolding close_def using closure_eq[OF _ _ _ \<open>P x'\<close>] by - (rule disjI1, force)
+    qed
+  }
+  then show ?case
+    by (intro allI impI)
+next
+  case 3
+  then show ?case
+    using subgraph by auto
+next
+  case 4
+  have "{a. Graph_Start_Defs.reachable (close RE) (closure s0) a} \<subseteq> closure ` UNIV"
+    by (smt Graph_Start_Defs.reachable_induct close_def full_SetCompr_eq mem_Collect_eq subsetI)
+  also have "finite \<dots>"
+    by (rule finite_closure)
+  finally show ?case .
+qed
+
+end (* Reachability Compatible Subsumption Graph Closure View *)
+
+locale Reachability_Compatible_Subsumption_Graph_Final = Reachability_Compatible_Subsumption_Graph +
+  fixes F :: "'a \<Rightarrow> bool" \<comment> \<open>Final states\<close>
+  assumes F_mono[intro]: "F a \<Longrightarrow> a \<preceq> b \<Longrightarrow> F b"
+
+locale Liveness_Compatible_Subsumption_Graph = Reachability_Compatible_Subsumption_Graph_Final +
+  assumes no_subsumption_cycle:
+    "G'.reachable x \<Longrightarrow> x \<rightarrow>G+' x \<Longrightarrow> x \<rightarrow>G+ x"
+
+subsection \<open>Reachability\<close>
+
+context Subsumption_Graph_Defs
+begin
+
+text \<open>Setup for automation\<close>
+context
+  includes graph_automation
+begin
+
+lemma G'_reachable_G_reachable[intro]:
+  "G.reachable a" if "G'.reachable a"
+  using that by (induction; blast)
+
+lemma G_reachable_G'_reachable[intro]:
+  "G'.reachable a" if "G.reachable a"
+  using that by (induction; blast)
+
+lemma G_G'_reachable_iff:
+  "G.reachable a \<longleftrightarrow> G'.reachable a"
+  by blast
+
+end (* Automation *)
+
+end (* Subsumption Graph Defs *)
+
+
+context Reachability_Compatible_Subsumption_Graph_Pre
+begin
+
+lemmas preorder_intros = order_trans less_trans less_imp_le
+
+lemma G'_finite_reachable: "finite {a. G'.reachable a}"
+  by (blast intro: finite_subset[OF _ finite_reachable])
+
+lemma G_reachable_has_surrogate:
+  "\<exists> t. G.reachable t \<and> s \<preceq> t \<and> (\<forall> s'. E t s' \<longrightarrow> RE t s')" if "G.reachable s"
+proof -
+  note [intro] = preorder_intros
+  from finite_reachable \<open>G.reachable s\<close> obtain x where
+    "\<forall>s'. E x s' \<longrightarrow> RE x s'" "G.reachable x" "((\<prec>)**) s x"
+    apply atomize_elim
+    apply (induction rule: rtranclp_ev_induct2)
+    using reachability_compatible by auto
+  moreover from \<open>((\<prec>)**) s x\<close> have "s \<prec> x \<or> s = x"
+    by induction auto
+  ultimately show ?thesis
+    by auto
+qed
+
+lemma reachable_has_surrogate:
+  "\<exists> t. G.reachable t \<and> s \<preceq> t \<and> (\<forall> s'. E t s' \<longrightarrow> RE t s')" if "reachable s"
+  using that
+proof induction
+  case start
+  have "G.reachable s0"
+    by auto
+  then show ?case
+    by (rule G_reachable_has_surrogate)
+next
+  case (step a b)
+  then obtain t where *: "G.reachable t" "a \<preceq> t" "(\<forall>s'. t \<rightarrow> s' \<longrightarrow> t \<rightarrow>G s')"
+    by auto
+  from mono[OF \<open>a \<preceq> t\<close> \<open>a \<rightarrow> b\<close>] \<open>reachable a\<close> \<open>G.reachable t\<close> obtain b' where
+    "t \<rightarrow> b'" "b \<preceq> b'"
+    by auto
+  with G_reachable_has_surrogate[of b'] * show ?case
+    by (auto intro: preorder_intros G.reachable_step)
+qed
+
+context
+  fixes F :: "'a \<Rightarrow> bool" \<comment> \<open>Final states\<close>
+  assumes F_mono[intro]: "F a \<Longrightarrow> a \<preceq> b \<Longrightarrow> F b"
+begin
+
+corollary reachability_correct:
+  "\<nexists> s'. reachable s' \<and> F s'" if "\<nexists> s'. G.reachable s' \<and> F s'"
+  using that by (auto dest!: reachable_has_surrogate)
+
+end (* Context for property *)
+
+end (* Reachability Compatible Subsumption Graph Pre *)
+
+
+context Reachability_Compatible_Subsumption_Graph
+begin
+
+text \<open>Setup for automation\<close>
+context
+  includes graph_automation
+begin
+
+lemma subgraph'[intro]:
+  "E s s'" if "RE s s'"
+  using that subgraph by blast
+
+lemma G_reachability_sound[intro]:
+  "reachable a" if "G.reachable a"
+  using that by (induction; blast)
+
+lemma G_steps_sound[intro]:
+  "steps xs" if "G.steps xs"
+  using that by (induction; blast)
+
+lemma G_run_sound[intro]:
+  "run xs" if "G.run xs"
+  using that by (coinduction arbitrary: xs) (auto 4 3 elim: G.run.cases)
+
+lemma G'_reachability_sound[intro]:
+  "reachable a" if "G'.reachable a"
+  using that by (induction; blast)
+
+lemma G'_finite_reachable: "finite {a. G'.reachable a}"
+  by (blast intro: finite_subset[OF _ finite_reachable])
+
+lemma G_steps_G'_steps[intro]:
+  "G'.steps as" if "G.steps as"
+  using that by induction auto
+
+lemma reachable_has_surrogate:
+  "\<exists> t. G.reachable t \<and> s \<preceq> t \<and> (\<forall> s'. E t s' \<longrightarrow> RE t s')" if "G.reachable s"
+proof -
+  note [intro] = preorder_intros
+  from finite_reachable \<open>G.reachable s\<close> obtain x where
+    "\<forall>s'. E x s' \<longrightarrow> RE x s'" "G.reachable x" "((\<prec>)**) s x"
+    apply atomize_elim
+    apply (induction rule: rtranclp_ev_induct2)
+    using reachability_compatible by auto
+  moreover from \<open>((\<prec>)**) s x\<close> have "s \<prec> x \<or> s = x"
+    by induction auto
+  ultimately show ?thesis
+    by auto
+qed
+
+lemma reachable_has_surrogate':
+  "\<exists> t. s \<preceq> t \<and> s \<rightarrow>G*' t \<and> (\<forall> s'. E t s' \<longrightarrow> RE t s')" if "G.reachable s"
+proof -
+  note [intro] = preorder_intros
+  from \<open>G.reachable s\<close> have \<open>G.reachable s\<close> by auto
+  from finite_reachable this obtain x where
+    real_edges: "\<forall>s'. E x s' \<longrightarrow> RE x s'" and "G.reachable x" "((\<prec>)**) s x"
+    apply atomize_elim
+    apply (induction rule: rtranclp_ev_induct2)
+    using reachability_compatible by auto
+  from \<open>((\<prec>)**) s x\<close> have "s \<prec> x \<or> s = x"
+    by induction auto
+  then show ?thesis
+  proof
+    assume "s \<prec> x"
+    with real_edges \<open>G.reachable x\<close> show ?thesis
+      by (inst_existentials "x") auto
+  next
+    assume "s = x"
+    with real_edges show ?thesis
+      by (inst_existentials "s") auto
+  qed
+qed
+
+lemma subsumption_step:
+  "\<exists> a'' b'. a' \<preceq> a'' \<and> b \<preceq> b' \<and> a'' \<rightarrow>G b' \<and> G.reachable a''" if
+  "reachable a" "E a b" "G.reachable a'" "a \<preceq> a'"
+proof -
+  note [intro] = preorder_intros
+  from mono[OF \<open>a \<preceq> a'\<close> \<open>E a b\<close> \<open>reachable a\<close>] \<open>G.reachable a'\<close> obtain b' where "E a' b'" "b \<preceq> b'"
+    by auto
+  from reachable_has_surrogate[OF \<open>G.reachable a'\<close>] obtain a''
+    where "a' \<preceq> a''" "G.reachable a''" and *: "\<forall> s'. E a'' s' \<longrightarrow> RE a'' s'"
+    by auto
+  from mono[OF \<open>a' \<preceq> a''\<close> \<open>E a' b'\<close>] \<open>G.reachable a'\<close> \<open>G.reachable a''\<close> obtain b'' where
+    "E a'' b''" "b' \<preceq> b''"
+    by auto
+  with * \<open>a' \<preceq> a''\<close> \<open>b \<preceq> b'\<close> \<open>G.reachable a''\<close> show ?thesis
+    by auto
+qed
+
+lemma subsumption_step':
+  "\<exists> b'. b \<preceq> b' \<and> a' \<rightarrow>G+' b'" if "reachable a" "a \<rightarrow> b" "G'.reachable a'" "a \<preceq> a'"
+proof -
+  note [intro] = preorder_intros
+  from mono[OF \<open>a \<preceq> a'\<close> \<open>E a b\<close> \<open>reachable a\<close>] \<open>G'.reachable a'\<close> obtain b' where
+    "b \<preceq> b'" "a' \<rightarrow> b'"
+    by auto
+  from reachable_has_surrogate'[of a'] \<open>G'.reachable a'\<close> obtain a'' where *:
+    "a' \<preceq> a''" "a' \<rightarrow>G*' a''" "\<forall>s'. a'' \<rightarrow> s' \<longrightarrow> a'' \<rightarrow>G s'"
+    by auto
+  with \<open>G'.reachable a'\<close> have "G'.reachable a''"
+    by blast
+  with mono[OF \<open>a' \<preceq> a''\<close> \<open>E a' b'\<close>] \<open>G'.reachable a'\<close> obtain b'' where
+    "b' \<preceq> b''" "a'' \<rightarrow> b''"
+    by auto
+  with * \<open>b \<preceq> b'\<close> \<open>b' \<preceq> b''\<close> \<open>G'.reachable a''\<close> show ?thesis
+    by (auto simp: G'.reaches1_reaches_iff2) (* XXX *)
+qed
+
+theorem reachability_complete':
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "a \<rightarrow>* s" "G.reachable a"
+  using that
+proof (induction)
+  case base
+  then show ?case by auto
+next
+  case (step s t)
+  then obtain s' where "s \<preceq> s'" "G.reachable s'"
+    by auto
+  with step(4) have "reachable a" "G.reachable s'"
+    by auto
+  with step(1) have "reachable s"
+    by auto
+  from subsumption_step[OF \<open>reachable s\<close> \<open>E s t\<close> \<open>G.reachable s'\<close> \<open>s \<preceq> s'\<close>] obtain s'' t' where
+    "s' \<preceq> s''" "t \<preceq> t'" "s'' \<rightarrow>G t'" "G.reachable s''"
+    by atomize_elim
+  with \<open>G.reachable s'\<close> show ?case
+    by auto
+qed
+
+theorem steps_complete':
+  "\<exists> ys. list_all2 (\<preceq>) xs ys \<and> G.steps (a # ys)" if
+  "steps (a # xs)" "G.reachable a"
+  using that
+proof (induction "a # xs" arbitrary: a xs rule: steps_alt_induct)
+  case (Single x)
+  then show ?case by auto
+oops
+
+theorem steps_complete':
+  "\<exists> c ys. list_all2 (\<preceq>) xs ys \<and> G.steps (c # ys) \<and> b \<preceq> c" if
+  "steps (a # xs)" "reachable a" "a \<preceq> b" "G.reachable b"
+oops
+
+(* XXX Does this hold? *)
+theorem run_complete':
+  "\<exists> ys. stream_all2 (\<preceq>) xs ys \<and> G.run (a ## ys)" if "run (a ## xs)" "G.reachable a"
+proof -
+  define f where "f = (\<lambda> x b. SOME y. x \<preceq> y \<and> RE b y)"
+  define gen where "gen a xs = sscan f xs a" for a xs
+  have gen_ctr: "gen x xs = f (shd xs) x ## gen (f (shd xs) x) (stl xs)" for x xs
+    unfolding gen_def by (subst sscan.ctr) (rule HOL.refl)
+  from that have "G.run (gen a xs)"
+  proof (coinduction arbitrary: a xs)
+    case run
+    then show ?case
+      apply (cases xs)
+      apply auto
+      apply (subst gen_ctr)
+      apply simp
+      apply (subst gen_ctr)
+      apply simp
+      apply rule
+oops
+
+corollary reachability_complete:
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "reachable s"
+  using reachability_complete'[of s0 s] that unfolding reachable_def by auto
+
+corollary reachability_correct:
+  "(\<exists> s'. s \<preceq> s' \<and> reachable s') \<longleftrightarrow> (\<exists> s'. s \<preceq> s' \<and> G.reachable s')"
+  by (blast dest: reachability_complete intro: preorder_intros)
+
+lemma steps_G'_steps:
+  "\<exists> ys ns. list_all2 (\<preceq>) xs (nths ys ns) \<and> G'.steps (b # ys)" if
+  "steps (a # xs)" "reachable a" "a \<preceq> b" "G'.reachable b"
+  using that
+proof (induction "a # xs" arbitrary: a b xs)
+  case (Single)
+  then show ?case by force
+next
+  case (Cons x y xs)
+  from subsumption_step'[OF \<open>reachable x\<close> \<open>E x y\<close> _ \<open>x \<preceq> b\<close>] \<open>G'.reachable b\<close> obtain b' where
+    "y \<preceq> b'" "b \<rightarrow>G+' b'"
+    by auto
+  with \<open>reachable x\<close> Cons.hyps(1) Cons.prems(3) obtain ys ns where
+    "list_all2 (\<preceq>) xs (nths ys ns)" "G'.steps (b' # ys)"
+    by atomize_elim (blast intro: Cons.hyps(3)[OF _ \<open>y \<preceq> b'\<close>] intro: graphI_aggressive)
+  from  \<open>b \<rightarrow>G+' b'\<close> this(2) obtain as where
+    "G'.steps (b # as @ b' # ys)"
+    by (fastforce intro: G'.graphI_aggressive1)
+  with \<open>y \<preceq> b'\<close> show ?case
+    apply (inst_existentials "as @ b' # ys" "{length as} \<union> {n + length as + 1 | n. n \<in> ns}")
+    subgoal
+      apply (subst nths_split, force)
+      apply (subst nths_nth, (simp; fail))
+      apply simp
+      apply (subst nths_shift, force)
+      subgoal premises prems
+      proof -
+        have
+          "{x - length as |x. x \<in> {Suc (n + length as) |n. n \<in> ns}} = {n + 1 | n. n \<in> ns}"
+          by force
+        with \<open>list_all2 _ _ _\<close> show ?thesis
+          by (simp add: nths_Cons)
+      qed
+      done
+    by assumption
+qed
+
+lemma cycle_G'_cycle'':
+  assumes "steps (s0 # ws @ x # xs @ [x])"
+  shows "\<exists> x' xs' ys'. x \<preceq> x' \<and> G'.steps (s0 # xs' @ x' # ys' @ [x'])"
+proof -
+  let ?n  = "card {x. G'.reachable x} + 1"
+  let ?xs = "x # concat (replicate ?n (xs @ [x]))"
+  from assms(1) have "steps (x # xs @ [x])"
+    by (auto intro: graphI_aggressive2)
+  with steps_replicate[of "x # xs @ [x]" ?n] have "steps ?xs"
+    by auto
+  have "steps (s0 # ws @ ?xs)"
+  proof -
+    from assms have "steps (s0 # ws @ [x])" (* XXX *)
+      by (auto intro: graphI_aggressive2)
+    with \<open>steps ?xs\<close> show ?thesis
+      by (fastforce intro: graphI_aggressive1)
+  qed
+  from steps_G'_steps[OF this, of s0] obtain ys ns where ys:
+    "list_all2 (\<preceq>) (ws @ x # concat (replicate ?n (xs @ [x]))) (nths ys ns)"
+    "G'.steps (s0 # ys)"
+    by auto
+  then obtain x' ys' ns' ws' where ys':
+    "G'.steps (x' # ys')" "G'.steps (s0 # ws' @ [x'])"
+    "list_all2 (\<preceq>) (concat (replicate ?n (xs @ [x]))) (nths ys' ns')"
+    apply atomize_elim
+    apply auto
+    apply (subst (asm) list_all2_append1)
+    apply safe
+    apply (subst (asm) list_all2_Cons1)
+    apply safe
+    apply (drule nths_eq_appendD)
+    apply safe
+    apply (drule nths_eq_ConsD)
+    apply safe
+    subgoal for ys1 ys2 z ys3 ys4 ys5 ys6 ys7 i
+      apply (inst_existentials z ys7)
+      subgoal premises prems
+        using prems(1) by (auto intro: G'.graphI_aggressive2)
+      subgoal premises prems
+      proof -
+        from prems have "G'.steps ((s0 # ys4 @ ys6 @ [z]) @ ys7)"
+          by auto
+        moreover then have "G'.steps (s0 # ys4 @ ys6 @ [z])"
+          by (auto intro: G'.graphI_aggressive2)
+        ultimately show ?thesis
+          by (inst_existentials "ys4 @ ys6") auto
+      qed
+      by force
+    done
+  let ?ys = "filter ((\<preceq>) x) ys'"
+  have "length ?ys \<ge> ?n"
+    using list_all2_replicate_elem_filter[OF ys'(3), of x]
+    using filter_nths_length[of "((\<preceq>) x)" ys' ns']
+    by auto
+  from \<open>G'.steps (s0 # ws' @ [x'])\<close> have "G'.reachable x'"
+    by - (rule G'.reachable_reaches, auto)
+  have "set ?ys \<subseteq> set ys'"
+    by auto
+  also have "\<dots> \<subseteq> {x. G'.reachable x}"
+    using \<open>G'.steps (x' # _)\<close> \<open>G'.reachable x'\<close>
+    by clarsimp (rule G'.reachable_steps_elem[rotated], assumption, auto)
+  finally have "\<not> distinct ?ys"
+    using distinct_card[of ?ys] \<open>_ >= ?n\<close>
+    by - (rule ccontr; drule distinct_length_le[OF G'_finite_reachable]; simp)
+  from not_distinct_decomp[OF this] obtain as y bs cs where "?ys = as @ [y] @ bs @ [y] @ cs"
+    by auto
+  then obtain as' bs' cs' where
+    "ys' = as' @ [y] @ bs' @ [y] @ cs'"
+    apply atomize_elim
+    apply simp
+    apply (drule filter_eq_appendD filter_eq_ConsD filter_eq_appendD[OF sym], clarify)+
+    apply clarsimp
+    subgoal for as1 as2 bs1 bs2 cs'
+      by (inst_existentials "as1 @ as2" "bs1 @ bs2") simp
+    done
+  have "G'.steps (y # bs' @ [y])"
+  proof -
+    (* XXX Decision procedure? *)
+    from \<open>G'.steps (x' # _)\<close> \<open>ys' = _\<close> show ?thesis
+      by (force intro: G'.graphI_aggressive2)
+  qed
+  moreover have "G'.steps (s0 # ws' @ x' # as' @ [y])"
+  proof -
+    (* XXX Decision procedure? *)
+    from \<open>G'.steps (x' # ys')\<close> \<open>ys' = _\<close> have "G'.steps (x' # as' @ [y])"
+      by (force intro: G'.graphI_aggressive2)
+    with \<open>G'.steps (s0 # ws' @ [x'])\<close> show ?thesis
+      by (fastforce intro: G'.graphI_aggressive1)
+  qed
+  moreover from \<open>?ys = _\<close> have "x \<preceq> y"
+  proof -
+    from \<open>?ys = _\<close> have "y \<in> set ?ys" by auto
+    then show ?thesis by auto
+  qed
+  ultimately show ?thesis
+    by (inst_existentials y "ws' @ x' # as'" bs'; fastforce intro: G'.graphI_aggressive1)
+qed
+
+lemma cycle_G'_cycle':
+  assumes "steps (s0 # ws @ x # xs @ [x])"
+  shows "\<exists> y ys. x \<preceq> y \<and> G'.steps (y # ys @ [y]) \<and> G'.reachable y"
+proof -
+  from cycle_G'_cycle''[OF assms] obtain x' xs' ys' where
+    "x \<preceq> x'" "G'.steps (s0 # xs' @ x' # ys' @ [x'])"
+    by auto
+  then show ?thesis
+    apply (inst_existentials x' ys')
+    subgoal by assumption
+    subgoal by (auto intro: G'.graphI_aggressive2)
+    by (rule G'.reachable_reaches, auto intro: G'.graphI_aggressive2)
+qed
+
+lemma cycle_G'_cycle:
+  assumes "reachable x" "x \<rightarrow>+ x"
+  shows "\<exists> y ys. x \<preceq> y \<and> G'.reachable y \<and> y \<rightarrow>G+' y"
+proof -
+  from assms(2) obtain xs where *: "steps (x # xs @ x # xs @ [x])"
+    by (fastforce intro: graphI_aggressive1)
+  from reachable_steps[of x] assms(1) obtain ws where "steps ws" "hd ws = s0" "last ws = x"
+    by auto
+  with * obtain us where "steps (s0 # (us @ xs) @ x # xs @ [x])"
+    by (cases ws; force intro: graphI_aggressive1) (* slow *)
+  from cycle_G'_cycle'[OF this] show ?thesis
+    by (auto intro: G'.graphI_aggressive2)
+qed
+
+corollary G'_reachability_complete:
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "G'.reachable s"
+  using reachability_complete that by auto
+
+end (* Subsumption *)
+
+end (* Reachability Compatible Subsumption Graph *)
+
+corollary (in Reachability_Compatible_Subsumption_Graph_Final) reachability_correct:
+  "(\<exists> s'. reachable s' \<and> F s') \<longleftrightarrow> (\<exists> s'. G.reachable s' \<and> F s')"
+  using reachability_complete by blast
+
+
+subsection \<open>Liveness\<close>
+
+theorem (in Liveness_Compatible_Subsumption_Graph) cycle_iff:
+  "(\<exists> x. x \<rightarrow>+ x \<and> reachable x \<and> F x) \<longleftrightarrow> (\<exists> x. x \<rightarrow>G+ x \<and> G.reachable x \<and> F x)"
+  by (auto 4 4 intro: no_subsumption_cycle steps_reaches1 dest: cycle_G'_cycle G.graphD)
+
+subsection \<open>Appendix\<close>
+
+context Subsumption_Graph_Pre_Nodes
+begin
+
+text \<open>Setup for automation\<close>
+context
+  includes graph_automation
+begin
+
+lemma steps_mono:
+  assumes "G'.steps (x # xs)" "x \<preceq> y" "V x" "V y"
+  shows "\<exists> ys. G'.steps (y # ys) \<and> list_all2 (\<preceq>) xs ys"
+  using assms including subgraph_automation
+proof (induction "x # xs" arbitrary: x y xs)
+  case (Single x)
+  then show ?case by auto
+next
+  case (Cons x y xs x')
+  from mono[OF \<open>x \<preceq> x'\<close>] \<open>x \<rightarrow> y\<close> Cons.prems obtain y' where "x' \<rightarrow> y'" "y \<preceq> y'"
+    by auto
+  with Cons.hyps(3)[OF \<open>y \<preceq> y'\<close>] \<open>x \<rightarrow> y\<close> Cons.prems obtain ys where
+    "G'.steps (y' # ys)" "list_all2 (\<preceq>) xs ys"
+    by auto
+  with \<open>x' \<rightarrow> y'\<close> \<open>y \<preceq> y'\<close> show ?case
+    by auto
+qed
+
+lemma steps_append_subsumption:
+  assumes "G'.steps (x # xs)" "G'.steps (y # ys)" "y \<preceq> last (x # xs)" "V x" "V y"
+  shows "\<exists> ys'. G'.steps (x # xs @ ys') \<and> list_all2 (\<preceq>) ys ys'"
+proof -
+  from assms have "V (last (x # xs))"
+    by - (rule G'_steps_V_last, auto)
+  from steps_mono[OF \<open>G'.steps (y # ys)\<close> \<open>y \<preceq> _\<close> \<open>V y\<close> this] obtain ys' where
+    "G'.steps (last (x # xs) # ys')" "list_all2 (\<preceq>) ys ys'"
+    by auto
+  with G'.steps_append[OF \<open>G'.steps (x # xs)\<close> this(1)] show ?thesis
+    by auto
+qed
+
+lemma steps_replicate_subsumption:
+  assumes "x \<preceq> last (x # xs)" "G'.steps (x # xs)" "n > 0" "V x"
+  notes [intro] = preorder_intros
+  shows "\<exists> ys. G'.steps (x # ys) \<and> list_all2 (\<preceq>) (concat (replicate n xs)) ys"
+  using assms
+proof (induction n)
+  case 0
+  then show ?case by simp
+next
+  case (Suc n)
+  show ?case
+  proof (cases n)
+    case 0
+    with Suc.prems show ?thesis
+      by (inst_existentials xs) (auto intro: list_all2_refl)
+  next
+    case prems: (Suc n')
+    with Suc \<open>n = _\<close> obtain ys where ys:
+      "list_all2 (\<preceq>) (concat (replicate n xs)) ys" "G'.steps (x # ys)"
+      by auto
+    with \<open>n = _\<close> have "list_all2 (\<preceq>) (concat (replicate n' xs) @ xs) ys"
+      by (metis append_Nil2 concat.simps(1,2) concat_append replicate_Suc replicate_append_same)
+    with \<open>x \<preceq> _\<close> have "x \<preceq> last (x # ys)"
+      by (cases xs; auto 4 3 dest: list_all2_last split: if_split_asm)
+    from steps_append_subsumption[OF \<open>G'.steps (x # ys)\<close> \<open>G'.steps (x # xs)\<close> this] \<open>V x\<close> obtain
+      ys' where "G'.steps (x # ys @ ys')" "list_all2 (\<preceq>) xs ys'"
+      by auto
+    with ys(1) \<open>n = _\<close> show ?thesis
+      apply (inst_existentials "ys @ ys'")
+      by auto
+        (metis
+          append_Nil2 concat.simps(1,2) concat_append list_all2_appendI replicate_Suc
+          replicate_append_same
+        )
+  qed
+qed
+
+context
+  assumes finite_V: "finite {x. V x}"
+begin
+
+(* XXX Unused *)
+lemma wf_less_on_reachable_set:
+  assumes antisym: "\<And> x y. x \<preceq> y \<Longrightarrow> y \<preceq> x \<Longrightarrow> x = y"
+  shows "wf {(x, y). y \<prec> x \<and> V x \<and> V y}" (is "wf ?S")
+proof (rule finite_acyclic_wf)
+  have "?S \<subseteq> {(x, y). V x \<and> V y}"
+    by auto
+  also have "finite \<dots>"
+    using finite_V by auto
+  finally show "finite ?S" .
+next
+  interpret order by unfold_locales (rule antisym)
+  show "acyclicP (\<lambda>x y. y \<prec> x \<and> V x \<and> V y)"
+    by (rule acyclicI_order[where f = id]) auto
+qed
+
+text \<open>
+  This shows that looking for cycles and pre-cycles is equivalent in monotone subsumption graphs.
+\<close>
+(* XXX Duplication -- cycle_G'_cycle'' *)
+lemma pre_cycle_cycle':
+  (* XXX Move to different locale *)
+  assumes A: "x \<preceq> x'" "G'.steps (x # xs @ [x'])" "V x"
+  shows "\<exists> x'' ys. x' \<preceq> x'' \<and> G'.steps (x'' # ys @ [x'']) \<and> V x''"
+proof -
+  let ?n  = "card {x. V x} + 1"
+  let ?xs = "concat (replicate ?n (xs @ [x']))"
+  from steps_replicate_subsumption[OF _ \<open>G'.steps _\<close>, of ?n] \<open>V x\<close> \<open>x \<preceq> x'\<close> obtain ys where
+    "G'.steps (x # ys)" "list_all2 (\<preceq>) ?xs ys"
+    by auto
+  let ?ys = "filter ((\<preceq>) x') ys"
+  have "length ?ys \<ge> ?n"
+    using list_all2_replicate_elem_filter[OF \<open>list_all2 (\<preceq>) ?xs ys\<close>, of x']
+    by auto
+  have "set ?ys \<subseteq> set ys"
+    by auto
+  also have "\<dots> \<subseteq> {x. V x}"
+    using G'_steps_V_all[OF \<open>G'.steps (x # ys)\<close>] \<open>V x\<close> unfolding list_all_iff by auto
+  finally have "\<not> distinct ?ys"
+    using distinct_card[of ?ys] \<open>_ >= ?n\<close>
+    by - (rule ccontr, drule distinct_length_le[OF finite_V], auto)
+  from not_distinct_decomp[OF this] obtain as y bs cs where "?ys = as @ [y] @ bs @ [y] @ cs"
+    by auto
+  then obtain as' bs' cs' where
+    "ys = as' @ [y] @ bs' @ [y] @ cs'"
+    apply atomize_elim
+    apply simp
+    apply (drule filter_eq_appendD filter_eq_ConsD filter_eq_appendD[OF sym], clarify)+
+    apply clarsimp
+    subgoal for as1 as2 bs1 bs2 cs'
+      by (inst_existentials "as1 @ as2" "bs1 @ bs2") simp
+    done
+  have "G'.steps (y # bs' @ [y])"
+  proof -
+    (* XXX Decision procedure? *)
+    from \<open>G'.steps (x # ys)\<close> \<open>ys = _\<close> have "G'.steps (x # as' @ (y # bs' @ [y]) @ cs')"
+      by auto
+    then show ?thesis
+      by - ((simp; fail) | drule G'.stepsD)+
+  qed
+  moreover have "V y"
+  proof -
+    from \<open>G'.steps (x # ys)\<close> \<open>ys = _\<close> have "G'.steps ((x # as' @ [y]) @ (bs' @ y # cs'))" (* XXX *)
+      by simp
+    then have "G'.steps (x # as' @ [y])"
+      by (blast dest: G'.stepsD)
+    with \<open>V x\<close> show ?thesis
+      by (auto dest: G'_steps_V_last)
+  qed
+  moreover from \<open>?ys = _\<close> have "x' \<preceq> y"
+  proof -
+    from \<open>?ys = _\<close> have "y \<in> set ?ys" by auto
+    then show ?thesis by auto
+  qed
+  ultimately show ?thesis
+    by auto
+qed
+
+lemma pre_cycle_cycle:
+  "(\<exists> x x'. V x \<and> x \<rightarrow>+ x' \<and> x \<preceq> x') \<longleftrightarrow> (\<exists> x. V x \<and> x \<rightarrow>+ x)"
+  including reaches_steps_iff by (force dest: pre_cycle_cycle')
+
+(* XXX Generalize subgraph properties *)
+lemma pre_cycle_cycle_reachable:
+  "(\<exists> x x'. a0 \<rightarrow>* x \<and> V x \<and> x \<rightarrow>+ x' \<and> x \<preceq> x') \<longleftrightarrow> (\<exists> x. a0 \<rightarrow>* x \<and> V x \<and> x \<rightarrow>+ x)"
+proof -
+  interpret interp: Subsumption_Graph_Pre_Nodes _ _ E "\<lambda> x. a0 \<rightarrow>* x \<and> V x"
+    including graph_automation_aggressive
+    by standard (drule mono, auto 4 3 simp: Subgraph_Node_Defs.E'_def E'_def)
+  interpret start: Graph_Start_Defs E' a0 .
+  have *: "start.reachable_subgraph.E' = interp.E'"
+    unfolding interp.E'_def start.reachable_subgraph.E'_def
+    unfolding start.reachable_def E'_def
+    by auto
+  have *: "start.reachable_subgraph.G'.reaches1 = interp.G'.reaches1"
+    unfolding tranclp_def * ..
+  have *: "interp.G'.reaches1 x y \<longleftrightarrow> x \<rightarrow>+ y" if "a0 \<rightarrow>* x" for x y
+    using start.reachable_reaches1_equiv[of x y] that unfolding * by (simp add: start.reachable_def)
+  from interp.pre_cycle_cycle finite_V show ?thesis
+    by (auto simp: *)
+qed
+
+end (* Automation *)
+
+end (* Finite Subgraph *)
+
+end (* Subsumption Graph Pre Nodes *)
+
+
+(* XXX Obsolete *)
+context Subsumption_Graph_Pre
+begin
+
+text \<open>Setup for automation\<close>
+context
+  includes graph_automation
+begin
+
+interpretation Subsumption_Graph_Pre_Nodes _ _ E reachable
+  apply standard
+  apply (drule mono)
+     apply (simp_all add: Subgraph_Node_Defs.E'_def)
+     apply force
+  by auto
+
+lemma steps_mono:
+  assumes "steps (x # xs)" "x \<preceq> y" "reachable x" "reachable y"
+  shows "\<exists> ys. steps (y # ys) \<and> list_all2 (\<preceq>) xs ys"
+  using assms steps_mono by (simp add: reachable_steps_equiv)
+
+lemma steps_append_subsumption:
+  assumes "steps (x # xs)" "steps (y # ys)" "y \<preceq> last (x # xs)" "reachable x" "reachable y"
+  shows "\<exists> ys'. steps (x # xs @ ys') \<and> list_all2 (\<preceq>) ys ys'"
+  using assms steps_append_subsumption by (simp add: reachable_steps_equiv)
+
+lemma steps_replicate_subsumption:
+  assumes "x \<preceq> last (x # xs)" "steps (x # xs)" "n > 0" "reachable x"
+  notes [intro] = preorder_intros
+  shows "\<exists> ys. steps (x # ys) \<and> list_all2 (\<preceq>) (concat (replicate n xs)) ys"
+  using assms steps_replicate_subsumption by (simp add: reachable_steps_equiv)
+
+context
+  assumes finite_reachable: "finite {x. reachable x}"
+begin
+
+(* XXX Unused *)
+lemma wf_less_on_reachable_set:
+  assumes antisym: "\<And> x y. x \<preceq> y \<Longrightarrow> y \<preceq> x \<Longrightarrow> x = y"
+  shows "wf {(x, y). y \<prec> x \<and> reachable x \<and> reachable y}" (is "wf ?S")
+proof (rule finite_acyclic_wf)
+  have "?S \<subseteq> {(x, y). reachable x \<and> reachable y}"
+    by auto
+  also have "finite \<dots>"
+    using finite_reachable by auto
+  finally show "finite ?S" .
+next
+  interpret order by standard (rule antisym)
+  show "acyclicP (\<lambda>x y. y \<prec> x \<and> reachable x \<and> reachable y)"
+    by (rule acyclicI_order[where f = id]) auto
+qed
+
+text \<open>
+  This shows that looking for cycles and pre-cycles is equivalent in monotone subsumption graphs.
+\<close>
+(* XXX Duplication -- cycle_G'_cycle'' *)
+lemma pre_cycle_cycle':
+  (* XXX Move to different locale *)
+  assumes A: "x \<preceq> x'" "steps (x # xs @ [x'])" "reachable x"
+  shows "\<exists> x'' ys. x' \<preceq> x'' \<and> steps (x'' # ys @ [x'']) \<and> reachable x''"
+  using assms pre_cycle_cycle'[OF finite_reachable] reachable_steps_equiv by meson
+
+lemma pre_cycle_cycle:
+  "(\<exists> x x'. reachable x \<and> reaches x x' \<and> x \<preceq> x') \<longleftrightarrow> (\<exists> x. reachable x \<and> reaches x x)"
+  including reaches_steps_iff by (force dest: pre_cycle_cycle')
+
+end (* Automation *)
+
+end (* Finite Reachable Subgraph *)
+
+end (* Subsumption Graph Pre *)
+
+
+context Subsumption_Graph_Defs
+begin
+
+sublocale G'': Graph_Start_Defs "\<lambda> x y. \<exists> z. G.reachable z \<and> x \<preceq> z \<and> RE z y" s0 .
+
+lemma G''_reachable_G'[intro]:
+  "G'.reachable x" if "G''.reachable x"
+  using that
+  unfolding G'.reachable_def G''.reachable_def G_G'_reachable_iff Graph_Start_Defs.reachable_def
+proof (induction)
+  case base
+  then show ?case
+    by blast
+next
+  case (step y z)
+  then obtain z' where
+    "RE** s0 z'" "y \<preceq> z'" "RE z' z"
+    by auto
+  from this(1) have "(\<lambda>x y. RE x y \<or> x \<prec> y \<and> RE** s0 y)** s0 z'"
+    by (induction; blast intro: rtranclp.intros(2))
+  with \<open>RE z' z\<close> show ?case
+    by (blast intro: rtranclp.intros(2))
+qed
+
+end (* Subsumption Graph Defs *)
+
+locale Reachability_Compatible_Subsumption_Graph_Total = Reachability_Compatible_Subsumption_Graph +
+  assumes total: "reachable a \<Longrightarrow> reachable b \<Longrightarrow> a \<preceq> b \<or> b \<preceq> a"
+begin
+
+sublocale G''_pre: Subsumption_Graph_Pre "(\<preceq>)" "(\<prec>)" "\<lambda> x y. \<exists> z. G.reachable z \<and> x \<preceq> z \<and> RE z y"
+proof (standard, safe, goal_cases)
+  case prems: (1 a b a' z)
+  show ?case
+  proof (cases "b \<preceq> z")
+    case True
+    with prems show ?thesis
+      by auto
+  next
+    case False
+    with total[of b z] prems have "z \<preceq> b"
+      by auto
+    with subsumption_step[of z a' b] prems obtain a'' b' where
+      "b \<preceq> a''" "a' \<preceq> b'" "RE a'' b'" "G.reachable a''"
+      by auto
+    then show ?thesis
+      by (inst_existentials b' a'') auto
+  qed
+qed
+
+end (* Reachability Compatible Subsumption Graph Total *)
+
+subsection \<open>Old Material\<close>
+
+locale Reachability_Compatible_Subsumption_Graph' = Subsumption_Graph_Defs + order "(\<preceq>)" "(\<prec>)" +
+  assumes reachability_compatible:
+    "\<forall> s. G.reachable s \<longrightarrow> (\<forall> s'. E s s' \<longrightarrow> RE s s') \<or> (\<exists> t. s \<prec> t \<and> G.reachable t)"
+  assumes subgraph: "\<forall> s s'. RE s s' \<longrightarrow> E s s'"
+  assumes finite_reachable: "finite {a. G.reachable a}"
+  assumes mono:
+    "a \<preceq> b \<Longrightarrow> E a a' \<Longrightarrow> reachable a \<Longrightarrow> G.reachable b \<Longrightarrow> \<exists> b'. E b b' \<and> a' \<preceq> b'"
+begin
+
+text \<open>Setup for automation\<close>
+context
+  includes graph_automation
+  notes [intro] = order.trans
+begin
+
+lemma subgraph'[intro]:
+  "E s s'" if "RE s s'"
+  using that subgraph by blast
+
+lemma G_reachability_sound[intro]:
+  "reachable a" if "G.reachable a"
+  using that unfolding reachable_def G.reachable_def by (induction; blast intro: rtranclp.intros(2))
+
+lemma G_steps_sound[intro]:
+  "steps xs" if "G.steps xs"
+  using that by induction auto
+
+lemma G_run_sound[intro]:
+  "run xs" if "G.run xs"
+  using that by (coinduction arbitrary: xs) (auto 4 3 elim: G.run.cases)
+
+lemma reachable_has_surrogate:
+  "\<exists> t. G.reachable t \<and> s \<preceq> t \<and> (\<forall> s'. E t s' \<longrightarrow> RE t s')" if "G.reachable s"
+  using that
+proof -
+  from finite_reachable \<open>G.reachable s\<close> obtain x where
+    "\<forall>s'. E x s' \<longrightarrow> RE x s'" "G.reachable x" "((\<prec>)**) s x"
+    apply atomize_elim
+    apply (induction rule: rtranclp_ev_induct2)
+    using reachability_compatible by auto
+  moreover from \<open>((\<prec>)**) s x\<close> have "s \<prec> x \<or> s = x"
+    by induction auto
+  ultimately show ?thesis by auto
+qed
+
+lemma subsumption_step:
+  "\<exists> a'' b'. a' \<preceq> a'' \<and> b \<preceq> b' \<and> RE a'' b' \<and> G.reachable a''" if
+  "reachable a" "E a b" "G.reachable a'" "a \<preceq> a'"
+proof -
+  from mono[OF \<open>a \<preceq> a'\<close> \<open>E a b\<close> \<open>reachable a\<close> \<open>G.reachable a'\<close>] obtain b' where "E a' b'" "b \<preceq> b'"
+    by auto
+  from reachable_has_surrogate[OF \<open>G.reachable a'\<close>] obtain a''
+    where "a' \<preceq> a''" "G.reachable a''" and *: "\<forall> s'. E a'' s' \<longrightarrow> RE a'' s'"
+    by auto
+  from mono[OF \<open>a' \<preceq> a''\<close> \<open>E a' b'\<close>] \<open>G.reachable a'\<close> \<open>G.reachable a''\<close> obtain b'' where
+    "E a'' b''" "b' \<preceq> b''"
+    by auto
+  with * \<open>a' \<preceq> a''\<close> \<open>b \<preceq> b'\<close> \<open>G.reachable a''\<close> show ?thesis by auto
+qed
+
+theorem reachability_complete':
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "E** a s" "G.reachable a"
+  using that
+proof (induction)
+  case base
+  then show ?case by auto
+next
+  case (step s t)
+  then obtain s' where "s \<preceq> s'" "G.reachable s'"
+    by auto
+  with step(4) have "reachable a" "G.reachable s'"
+    by auto
+  with step(1) have "reachable s"
+    by (auto simp: reachable_def)
+  from subsumption_step[OF \<open>reachable s\<close> \<open>E s t\<close> \<open>G.reachable s'\<close> \<open>s \<preceq> s'\<close>] obtain s'' t' where
+    "s' \<preceq> s''" "t \<preceq> t'" "s'' \<rightarrow>G t'" "G.reachable s''"
+    by atomize_elim
+  with \<open>G.reachable s'\<close> show ?case
+    by (auto simp: reachable_def)
+qed
+
+theorem steps_complete':
+  "\<exists> ys. list_all2 (\<preceq>) xs ys \<and> G.steps (a # ys)" if
+  "steps (a # xs)" "G.reachable a"
+  using that
+proof (induction "a # xs" arbitrary: a xs rule: steps_alt_induct)
+  case (Single x)
+  then show ?case by auto
+oops
+
+theorem steps_complete':
+  "\<exists> c ys. list_all2 (\<preceq>) xs ys \<and> G.steps (c # ys) \<and> b \<preceq> c" if
+  "steps (a # xs)" "reachable a" "a \<preceq> b" "G.reachable b"
+  using that
+proof (induction "a # xs" arbitrary: a b xs)
+  case (Single x)
+  then show ?case by auto
+next
+  case (Cons x y xs)
+  from subsumption_step[OF \<open>reachable x\<close> \<open>E _ _\<close> \<open>G.reachable b\<close> \<open>x \<preceq> b\<close>] obtain b' y' where
+    "b \<preceq> b'" "y \<preceq> y'" "b' \<rightarrow>G y'" "G.reachable b'"
+    by atomize_elim
+  with Cons obtain y'' ys where "list_all2 (\<preceq>) xs ys" "G.steps (y'' # ys)" "y' \<preceq> y''"
+    oops
+
+(* XXX Does this hold? *)
+theorem run_complete':
+  "\<exists> ys. stream_all2 (\<preceq>) xs ys \<and> G.run (a ## ys)" if "run (a ## xs)" "G.reachable a"
+proof -
+  define f where "f = (\<lambda> x b. SOME y. x \<preceq> y \<and> RE b y)"
+  define gen where "gen a xs = sscan f xs a" for a xs
+  have gen_ctr: "gen x xs = f (shd xs) x ## gen (f (shd xs) x) (stl xs)" for x xs
+    unfolding gen_def by (subst sscan.ctr) (rule HOL.refl)
+  from that have "G.run (gen a xs)"
+  proof (coinduction arbitrary: a xs)
+    case run
+    then show ?case
+      apply (cases xs)
+      apply auto
+      apply (subst gen_ctr)
+      apply simp
+      apply (subst gen_ctr)
+      apply simp
+      apply rule
+oops
+
+corollary reachability_complete:
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "reachable s"
+  using reachability_complete'[of s0 s] that unfolding reachable_def by auto
+
+corollary reachability_correct:
+  "(\<exists> s'. s \<preceq> s' \<and> reachable s') \<longleftrightarrow> (\<exists> s'. s \<preceq> s' \<and> G.reachable s')"
+  using reachability_complete by blast
+
+lemma G'_reachability_sound[intro]:
+  "reachable a" if "G'.reachable a"
+  using that by induction auto
+
+corollary G'_reachability_complete:
+  "\<exists> s'. s \<preceq> s' \<and> G.reachable s'" if "G'.reachable s"
+  using reachability_complete that by auto
+
+end (* Automation *)
+
+end (* Reachability Compatible Subsumption Graph' *)
+
+end (* Theory *)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/TA_Syntax_Bundles.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/TA_Syntax_Bundles.thy
@@ -0,0 +1,60 @@
+theory TA_Syntax_Bundles
+  imports
+    "HOL-Library.Stream"
+    "HOL-Imperative_HOL.Imperative_HOL"
+    "Word_Lib.Bit_Shifts_Infix_Syntax"
+    "Word_Lib.Syntax_Bundles"
+    Automatic_Refinement.Relators
+    "HOL-Library.Extended_Nat"
+begin
+
+bundle bit_syntax begin
+
+unbundle bit_projection_infix_syntax
+
+notation
+      ring_bit_operations_class.not ("NOT")
+  and semiring_bit_operations_class.and (infixr "AND" 64)
+  and semiring_bit_operations_class.or  (infixr "OR"  59)
+  and semiring_bit_operations_class.xor (infixr "XOR" 59)
+  and shiftl (infixl "<<" 55)
+  and shiftr (infixl ">>" 55)
+
+end
+
+bundle no_bit_syntax begin
+no_notation
+      ring_bit_operations_class.not ("NOT")
+  and semiring_bit_operations_class.and (infixr "AND" 64)
+  and semiring_bit_operations_class.or  (infixr "OR"  59)
+  and semiring_bit_operations_class.xor (infixr "XOR" 59)
+  and shiftl (infixl "<<" 55)
+  and shiftr (infixl ">>" 55)
+  and bit (infixl "!!" 100)
+end
+
+bundle imperative_hol_syntax begin
+notation
+  Ref.noteq (infix "=!=" 70) and
+  Ref.lookup ("!_" 61) and
+  Ref.update ("_ := _" 62) and
+  Array.noteq (infix "=!!=" 70)
+end
+
+bundle no_imperative_hol_syntax begin
+no_notation
+  Ref.noteq (infix "=!=" 70) and
+  Ref.lookup ("!_" 61) and
+  Ref.update ("_ := _" 62) and
+  Array.noteq (infix "=!!=" 70)
+end
+
+bundle no_library_syntax begin
+no_notation Stream.snth (infixl "!!" 100)
+no_notation fun_rel_syn (infixr "\<rightarrow>" 60)
+no_notation Extended_Nat.infinity ("\<infinity>")
+unbundle no_imperative_hol_syntax
+unbundle no_bit_syntax
+end
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Temporal_Logics.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Temporal_Logics.thy
@@ -0,0 +1,495 @@
+theory Temporal_Logics
+  imports Timed_Automata.CTL LTL.LTL LTL_Master_Theorem.Omega_Words_Fun_Stream
+begin
+
+lemmas [simp] = holds.simps
+
+lemma suffix_stl:
+  "suffix (Suc 0) (snth xs) = snth (stl xs)"
+  unfolding suffix_def by auto
+
+lemma suntil_iff_sdrop:
+  "(\<phi> suntil \<psi>) xs \<longleftrightarrow> (\<exists>i. \<psi> (sdrop i xs) \<and> (\<forall>j<i. \<phi> (sdrop j xs)))"
+proof safe
+  show "\<exists>i. \<psi> (sdrop i xs) \<and> (\<forall>j<i. \<phi> (sdrop j xs))" if "(\<phi> suntil \<psi>) xs"
+    using that
+  proof (induction rule: suntil.induct)
+    case (base \<omega>)
+    then show ?case
+      by force
+  next
+    case (step \<omega>)
+    then obtain i where "\<psi> (sdrop i (stl \<omega>))" "\<forall>j<i. \<phi> (sdrop j (stl \<omega>))"
+      by safe
+    with \<open>\<phi> \<omega>\<close> have "\<phi> (sdrop j \<omega>)" if "j<i + 1" for j
+      using that by (cases j) auto
+    with \<open>\<psi> (sdrop i (stl \<omega>))\<close> show ?case
+      by (inst_existentials "i + 1") auto
+  qed
+  show "(\<phi> suntil \<psi>) xs" if "\<psi> (sdrop i xs)" "\<forall>j<i. \<phi> (sdrop j xs)" for i
+    using that by (induction i arbitrary: xs; fastforce intro: suntil.intros)
+qed
+
+lemma to_stream_suffix_Suc:
+  "to_stream (suffix (Suc k) xs) = stl (to_stream (suffix k xs))"
+  by (metis add.right_neutral add_Suc_right suffix_stl suffix_suffix
+        to_omega_def to_omega_to_stream to_stream_to_omega)
+
+lemma to_stream_stake:
+  "sdrop k (to_stream w) = to_stream (suffix k w)"
+  by (induction k arbitrary: w) (auto simp: sdrop_stl to_stream_suffix_Suc)
+
+lemma to_omega_suffix[simp]:
+  "suffix k (to_omega s) = to_omega (sdrop k s)"
+  by auto
+
+primrec semantics_ltlc' :: "['a word, ('a \<Rightarrow> bool) ltlc] \<Rightarrow> bool" ("_ \<Turnstile>c'' _" [80,80] 80)
+where
+  "\<xi> \<Turnstile>c' truec = True"
+| "\<xi> \<Turnstile>c' falsec = False"
+| "\<xi> \<Turnstile>c' propc(q) = (q (\<xi> 0))"
+| "\<xi> \<Turnstile>c' notc \<phi> = (\<not> \<xi> \<Turnstile>c' \<phi>)"
+| "\<xi> \<Turnstile>c' \<phi> andc \<psi> = (\<xi> \<Turnstile>c' \<phi> \<and> \<xi> \<Turnstile>c' \<psi>)"
+| "\<xi> \<Turnstile>c' \<phi> orc \<psi> = (\<xi> \<Turnstile>c' \<phi> \<or> \<xi> \<Turnstile>c' \<psi>)"
+| "\<xi> \<Turnstile>c' \<phi> impliesc \<psi> = (\<xi> \<Turnstile>c' \<phi> \<longrightarrow> \<xi> \<Turnstile>c' \<psi>)"
+| "\<xi> \<Turnstile>c' Xc \<phi> = (suffix 1 \<xi> \<Turnstile>c' \<phi>)"
+| "\<xi> \<Turnstile>c' Fc \<phi> = (\<exists>i. suffix i \<xi> \<Turnstile>c' \<phi>)"
+| "\<xi> \<Turnstile>c' Gc \<phi> = (\<forall>i. suffix i \<xi> \<Turnstile>c' \<phi>)"
+| "\<xi> \<Turnstile>c' \<phi> Uc \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>c' \<psi> \<and> (\<forall>j<i. suffix j \<xi> \<Turnstile>c' \<phi>))"
+| "\<xi> \<Turnstile>c' \<phi> Rc \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>c' \<psi> \<or> (\<exists>j<i. suffix j \<xi> \<Turnstile>c' \<phi>))"
+| "\<xi> \<Turnstile>c' \<phi> Wc \<psi> = (\<forall>i. suffix i \<xi> \<Turnstile>c' \<phi> \<or> (\<exists>j\<le>i. suffix j \<xi> \<Turnstile>c' \<psi>))"
+| "\<xi> \<Turnstile>c' \<phi> Mc \<psi> = (\<exists>i. suffix i \<xi> \<Turnstile>c' \<phi> \<and> (\<forall>j\<le>i. suffix j \<xi> \<Turnstile>c' \<psi>))"
+
+lemma semantics_ltlc'_sugar [simp]:
+  "\<xi> \<Turnstile>c' \<phi> iffc \<psi> = (\<xi> \<Turnstile>c' \<phi> \<longleftrightarrow> \<xi> \<Turnstile>c' \<psi>)"
+  "\<xi> \<Turnstile>c' Fc \<phi> = \<xi> \<Turnstile>c' (truec Uc \<phi>)"
+  "\<xi> \<Turnstile>c' Gc \<phi> = \<xi> \<Turnstile>c' (falsec Rc \<phi>)"
+  by (auto simp add: Iff_ltlc_def)
+
+definition "language_ltlc' \<phi> \<equiv> {\<xi>. \<xi> \<Turnstile>c' \<phi>}"
+
+fun ltlc_to_pltl :: "('a \<Rightarrow> bool) ltlc \<Rightarrow> 'a pltl"
+where
+  "ltlc_to_pltl truec = truep"
+| "ltlc_to_pltl falsec = falsep"
+| "ltlc_to_pltl (propc(q)) = atomp(q)"
+| "ltlc_to_pltl (notc \<phi>) = notp (ltlc_to_pltl \<phi>)"
+| "ltlc_to_pltl (\<phi> andc \<psi>) = (ltlc_to_pltl \<phi>) andp (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (\<phi> orc \<psi>) = (ltlc_to_pltl \<phi>) orp (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (\<phi> impliesc \<psi>) = (ltlc_to_pltl \<phi>) impliesp (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (Xc \<phi>) = Xp (ltlc_to_pltl \<phi>)"
+| "ltlc_to_pltl (Fc \<phi>) = Fp (ltlc_to_pltl \<phi>)"
+| "ltlc_to_pltl (Gc \<phi>) = Gp (ltlc_to_pltl \<phi>)"
+| "ltlc_to_pltl (\<phi> Uc \<psi>) = (ltlc_to_pltl \<phi>) Up (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (\<phi> Rc \<psi>) = (ltlc_to_pltl \<phi>) Rp (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (\<phi> Wc \<psi>) = (ltlc_to_pltl \<phi>) Wp (ltlc_to_pltl \<psi>)"
+| "ltlc_to_pltl (\<phi> Mc \<psi>) = (ltlc_to_pltl \<phi>) Mp (ltlc_to_pltl \<psi>)"
+
+lemma ltlc_to_pltl_semantics [simp]:
+  "w \<Turnstile>p ltlc_to_pltl \<phi> \<longleftrightarrow> w \<Turnstile>c' \<phi>"
+  by (induction \<phi> arbitrary: w) simp_all
+
+lemma semantics_ltlc'_semantics_ltlc_atoms_iff:
+  "w \<Turnstile>c' \<phi> \<longleftrightarrow> (\<lambda>i. {a \<in> atoms_ltlc \<phi>. a (w i)}) \<Turnstile>c \<phi>"
+proof -
+  have *:
+    "(\<lambda>i. {a. (a \<in> atoms_ltlc \<phi>1 \<or> a \<in> atoms_ltlc \<phi>2) \<and> a (w i)}) \<Turnstile>c \<phi>1
+    \<longleftrightarrow> ((\<lambda>i. {a \<in> atoms_ltlc \<phi>1. a (w i)}) \<Turnstile>c \<phi>1)" for w :: "nat \<Rightarrow> 'a" and \<phi>1 \<phi>2
+    by (rule ltlc_eq_on) (auto simp: pw_eq_on_def)
+  have **:
+    "(\<lambda>i. {a. (a \<in> atoms_ltlc \<phi>1 \<or> a \<in> atoms_ltlc \<phi>2) \<and> a (w i)}) \<Turnstile>c \<phi>2
+    \<longleftrightarrow> ((\<lambda>i. {a \<in> atoms_ltlc \<phi>2. a (w i)}) \<Turnstile>c \<phi>2)"  for w :: "nat \<Rightarrow> 'a" and \<phi>1 \<phi>2
+      by (rule ltlc_eq_on) (auto simp: pw_eq_on_def)
+  show ?thesis
+    by (induction \<phi> arbitrary: w) (simp_all add: suffix_def * **)
+qed
+
+lemma semantics_ltlc'_semantics_ltlc_atoms_iff':
+  "w \<Turnstile>c' \<phi> \<longleftrightarrow> ((\<lambda>x. {a \<in> atoms_ltlc \<phi>. a x}) o w) \<Turnstile>c \<phi>"
+  unfolding comp_def by (rule semantics_ltlc'_semantics_ltlc_atoms_iff)
+
+lemma map_semantics_ltlc_inj:
+  assumes "inj f"
+  shows "w \<Turnstile>c \<phi> \<longleftrightarrow> (image f o w) \<Turnstile>c map_ltlc f \<phi>"
+  using assms unfolding comp_def by (intro map_semantics_ltlc_aux) auto
+
+lemma semantics_ltlc'_abstract:
+  assumes "inj abstr" "\<And>x. label x = abstr ` {a \<in> atoms_ltlc \<phi>. a x}"
+  shows "w \<Turnstile>c' \<phi> \<longleftrightarrow> (label o w) \<Turnstile>c map_ltlc abstr \<phi>"
+  by (subst semantics_ltlc'_semantics_ltlc_atoms_iff')
+     (simp add: comp_def assms(2) map_semantics_ltlc_inj[OF \<open>inj abstr\<close>])
+
+context
+  includes lifting_syntax
+begin
+
+lemma holds_transfer:
+  "((R ===> (=)) ===> stream_all2 R ===> (=)) holds holds"
+  apply (intro rel_funI)
+  subgoal for \<phi> \<psi> xs ys
+    by (cases xs; cases ys) (auto dest: rel_funD)
+  done
+
+lemma alw_mono1:
+  "alw \<phi> xs" if "(stream_all2 R ===> (=)) \<phi> \<psi>" "stream_all2 R xs ys" "alw \<psi> ys"
+  using that(2,3)
+  by (coinduction arbitrary: xs ys) (use that(1) stream.rel_sel in \<open>auto dest!: rel_funD\<close>)
+
+lemma alw_mono2:
+  "alw \<psi> ys" if "(stream_all2 R ===> (=)) \<phi> \<psi>" "stream_all2 R xs ys" "alw \<phi> xs"
+  using that(2,3)
+  by (coinduction arbitrary: xs ys) (use that(1) stream.rel_sel in \<open>auto dest!: rel_funD\<close>)
+
+lemma alw_transfer':
+  "alw x xs = alw y ys" if "(stream_all2 R ===> (=)) x y" "stream_all2 R xs ys"
+  using alw_mono1 alw_mono2 that by blast
+
+lemma alw_transfer:
+  "((stream_all2 R ===> (=)) ===> stream_all2 R ===> (=)) alw alw"
+  by (intro rel_funI) (rule alw_transfer')
+
+lemma nxt_transfer:
+  "((stream_all2 R ===> (=)) ===> stream_all2 R ===> (=)) nxt nxt"
+  by (intro rel_funI) (simp, meson rel_funD stream.rel_sel)
+
+lemma suntil_mono1:
+  "(\<phi> suntil \<psi>) xs"
+  if "(stream_all2 R ===> (=)) \<phi> \<phi>'" "(stream_all2 R ===> (=)) \<psi> \<psi>'" "stream_all2 R xs ys"
+     "(\<phi>' suntil \<psi>') ys"
+  using that(4,3)
+proof (induction arbitrary: xs)
+  case (base \<omega>)
+  then show ?case
+    using that(1,2) by (auto dest!: rel_funD intro: suntil.base)
+next
+  case (step \<omega>)
+  from \<open>stream_all2 R xs \<omega>\<close> have "stream_all2 R (stl xs) (stl \<omega>)"
+    using stream.rel_sel by auto
+  from \<open>\<phi>' \<omega>\<close> \<open>stream_all2 R xs \<omega>\<close> have "\<phi> xs"
+    using that(1,2) by (auto dest!: rel_funD)
+  moreover from step.IH \<open>stream_all2 R xs \<omega>\<close> have "(\<phi> suntil \<psi>) (stl xs)"
+    using stream.rel_sel by auto
+  ultimately show ?case ..
+qed
+
+lemma suntil_mono2:
+  "(\<phi>' suntil \<psi>') ys"
+  if "(stream_all2 R ===> (=)) \<phi> \<phi>'" "(stream_all2 R ===> (=)) \<psi> \<psi>'" "stream_all2 R xs ys"
+     "(\<phi> suntil \<psi>) xs"
+  using that(4,3)
+proof (induction arbitrary: ys)
+  case (base \<omega>)
+  then show ?case
+    using that(1,2) by (auto dest!: rel_funD intro: suntil.base)
+next
+  case (step xs \<omega>)
+  from \<open>stream_all2 R xs \<omega>\<close> have "stream_all2 R (stl xs) (stl \<omega>)"
+    using stream.rel_sel by auto
+  from \<open>\<phi> xs\<close> \<open>stream_all2 R xs \<omega>\<close> have "\<phi>' \<omega>"
+    using that(1,2) by (auto dest!: rel_funD)
+  moreover from step.IH \<open>stream_all2 R xs \<omega>\<close> have "(\<phi>' suntil \<psi>') (stl \<omega>)"
+    using stream.rel_sel by auto
+  ultimately show ?case ..
+qed
+
+lemma suntil_transfer':
+  "(\<phi> suntil \<psi>) xs = (\<phi>' suntil \<psi>') ys"
+  if "(stream_all2 R ===> (=)) \<phi> \<phi>'" "(stream_all2 R ===> (=)) \<psi> \<psi>'" "stream_all2 R xs ys"
+  using suntil_mono1 suntil_mono2 that by metis
+
+lemma suntil_transfer:
+  "((stream_all2 R ===> (=)) ===> (stream_all2 R ===> (=)) ===> stream_all2 R ===> (=))
+    Linear_Temporal_Logic_on_Streams.suntil Linear_Temporal_Logic_on_Streams.suntil"
+  by (intro rel_funI) (rule suntil_transfer')
+
+lemma ev_mono1:
+  "ev \<phi> xs" if "(stream_all2 R ===> (=)) \<phi> \<psi>" "stream_all2 R xs ys" "ev \<psi> ys"
+  using that(3,2) by (induction arbitrary: xs; use that(1) stream.rel_sel in \<open>blast dest: rel_funD\<close>)
+
+lemma ev_mono2:
+  "ev \<psi> ys" if "(stream_all2 R ===> (=)) \<phi> \<psi>" "stream_all2 R xs ys" "ev \<phi> xs"
+  using that(3,2) by (induction arbitrary: ys; use that(1) stream.rel_sel in \<open>blast dest: rel_funD\<close>)
+
+lemma ev_transfer':
+  "ev x xs = ev y ys" if "(stream_all2 R ===> (=)) x y" "stream_all2 R xs ys"
+  using ev_mono1 ev_mono2 that by blast
+
+lemma ev_transfer:
+  "((stream_all2 R ===> (=)) ===> (stream_all2 R ===> (=))) ev ev"
+  by (intro rel_funI) (rule ev_transfer')
+
+end
+
+datatype 'a ctl_formula =
+  AG "'a ctl_formula" | AX "'a ctl_formula" | EG "'a ctl_formula" | EX "'a ctl_formula" | PropC 'a |
+  ImpliesC "'a ctl_formula" "'a ctl_formula" | NotC "'a ctl_formula"
+
+datatype 'a state_formula =
+  All "'a path_formula" | Ex "'a path_formula"
+| ImpliesS "'a state_formula" "'a state_formula" | NotS "'a state_formula" | PropS 'a
+and 'a path_formula =
+  G "'a path_formula" | F "'a path_formula"
+| X "'a path_formula" | Until "'a path_formula" "'a path_formula"
+| ImpliesP "'a path_formula" "'a path_formula" | NotP "'a path_formula" | State "'a state_formula"
+| FalseP
+
+fun ctl_to_state where
+  "ctl_to_state (AG \<phi>) = All (G (State (ctl_to_state \<phi>)))"
+| "ctl_to_state (AX \<phi>) = All (F (State (ctl_to_state \<phi>)))"
+| "ctl_to_state (EG \<phi>) = Ex  (G (State (ctl_to_state \<phi>)))"
+| "ctl_to_state (ctl_formula.EX \<phi>) = Ex (F (State (ctl_to_state \<phi>)))"
+| "ctl_to_state (PropC \<phi>) = PropS \<phi>"
+| "ctl_to_state (ImpliesC \<phi> \<psi>) = ImpliesS (ctl_to_state \<phi>) (ctl_to_state \<psi>)"
+| "ctl_to_state (NotC \<phi>) = NotS (ctl_to_state \<phi>)"
+
+fun ltlp_to_path where
+  "ltlp_to_path falsep = FalseP"
+| "ltlp_to_path (atomp(\<phi>)) = State (PropS \<phi>)"
+| "ltlp_to_path (\<phi> impliesp \<psi>) = ImpliesP (ltlp_to_path \<phi>) (ltlp_to_path \<psi>)"
+| "ltlp_to_path (Xp \<phi>) = X (ltlp_to_path \<phi>)"
+| "ltlp_to_path (\<phi> Up \<psi>) = Until (ltlp_to_path \<phi>) (ltlp_to_path \<psi>)"
+
+fun rel_pltl where
+  "rel_pltl R falsep falsep = True"
+| "rel_pltl R atomp(x) atomp(y) = R x y"
+| "rel_pltl R (x impliesp y) (x' impliesp y') \<longleftrightarrow> rel_pltl R x x' \<and> rel_pltl R y y'"
+| "rel_pltl R (x Up y) (x' Up y') \<longleftrightarrow> rel_pltl R x x' \<and> rel_pltl R y y'"
+| "rel_pltl R (Xp x) (Xp y) \<longleftrightarrow> rel_pltl R x y"
+| "rel_pltl _ _ _ = False"
+
+lemma rel_ltlp_to_path:
+  "rel_pltl R \<phi> \<psi> \<longleftrightarrow> rel_path_formula R (ltlp_to_path \<phi>) (ltlp_to_path \<psi>)"
+  by (induction R \<phi> \<psi> rule: rel_pltl.induct) auto
+
+lemma [simp]:
+  "falsep = truep \<longleftrightarrow> False"
+  unfolding True_ltlp_def Not_ltlp_def by auto
+
+lemmas ltlp_defs =
+  True_ltlp_def Not_ltlp_def And_ltlp_def Or_ltlp_def
+  Eventually_ltlp_def Always_ltlp_def Release_ltlp_def WeakUntil_ltlp_def StrongRelease_ltlp_def
+
+text \<open>The converse does not hold!\<close>
+lemma rel_ltlc_to_pltl:
+  "rel_pltl R (ltlc_to_pltl \<phi>) (ltlc_to_pltl \<psi>)" if "rel_ltlc R \<phi> \<psi>"
+  using that by (induction rule: ltlc.rel_induct) (auto simp: ltlp_defs)
+
+context Graph_Defs
+begin
+
+fun models_state and models_path where
+  "models_state (PropS P) x = P x"
+| "models_state (All \<phi>) x = (\<forall>xs. run (x ## xs) \<longrightarrow> models_path \<phi> (x ## xs))"
+| "models_state (Ex \<phi>) x = (\<exists>xs. run (x ## xs) \<and> models_path \<phi> (x ## xs))"
+| "models_state (ImpliesS \<psi> \<psi>') xs = (models_state \<psi> xs \<longrightarrow> models_state \<psi>' xs)"
+| "models_state (NotS \<psi>) xs = (\<not> models_state \<psi> xs)"
+| "models_path  (State \<psi>) = holds (models_state \<psi>)"
+| "models_path  (G \<psi>) = alw (models_path \<psi>)"
+| "models_path  (F \<psi>) = ev (models_path \<psi>)"
+| "models_path  (X \<psi>) = nxt (models_path \<psi>)"
+| "models_path  (Until \<psi> \<psi>') = models_path \<psi> suntil models_path \<psi>'"
+| "models_path  FalseP = (\<lambda>_. False)"
+| "models_path  (ImpliesP \<psi> \<psi>') = (\<lambda>xs. models_path \<psi> xs \<longrightarrow> models_path \<psi>' xs)"
+| "models_path  (NotP \<psi>) = (\<lambda>xs. \<not> models_path \<psi> xs)"
+
+fun models_ctl where
+  "models_ctl (PropC P) = P"
+| "models_ctl (AG \<phi>) = Alw_alw (models_ctl \<phi>)"
+| "models_ctl (AX \<phi>) = Alw_ev (models_ctl \<phi>)"
+| "models_ctl (EG \<phi>) = Ex_alw (models_ctl \<phi>)"
+| "models_ctl (ctl_formula.EX \<phi>) = Ex_ev (models_ctl \<phi>)"
+| "models_ctl (ImpliesC \<phi> \<psi>) = (\<lambda>x. models_ctl \<phi> x \<longrightarrow> models_ctl \<psi> x)"
+| "models_ctl (NotC \<phi>) = (\<lambda>x. \<not> models_ctl \<phi> x)"
+
+fun models_ltlp where
+  "models_ltlp falsep = (\<lambda>_. False)"
+| "models_ltlp (atomp(P)) = holds P"
+| "models_ltlp (\<phi> impliesp \<psi>) = (\<lambda>x. models_ltlp \<phi> x \<longrightarrow> models_ltlp \<psi> x)"
+| "models_ltlp (\<phi> Up \<psi>) = models_ltlp \<phi> suntil models_ltlp \<psi>"
+| "models_ltlp (Xp \<phi>) = nxt (models_ltlp \<phi>)"
+
+lemma models_ltlp_correct:
+  "models_ltlp \<phi> xs \<longleftrightarrow> to_omega xs \<Turnstile>p \<phi>"
+  by (induction \<phi> arbitrary: xs; simp add: suntil_iff_sdrop)
+
+definition
+  "models_ltlc \<phi> xs = to_omega xs \<Turnstile>c' \<phi>"
+
+lemma models_ltlc_alt_def:
+  "models_ltlc \<phi> = models_ltlp (ltlc_to_pltl \<phi>)"
+  unfolding models_ltlc_def models_ltlp_correct by simp
+
+theorem ctl_to_state_correct:
+  "models_ctl \<phi> = models_state (ctl_to_state \<phi>)"
+  by (induction \<phi>) (simp add: Alw_alw_def Alw_ev_def Ex_ev_def Ex_alw_def)+
+
+theorem ltlp_to_path_correct:
+  "models_ltlp \<phi> = models_path (ltlp_to_path \<phi>)"
+  by (induction \<phi>; simp)
+
+end
+
+context Bisimulation_Invariant
+begin
+
+context
+  includes lifting_syntax
+begin
+
+abbreviation compatible where
+  "compatible \<equiv> A_B.equiv' ===> (=)"
+
+abbreviation (input)
+  "compatible_op \<equiv> compatible ===> compatible"
+
+abbreviation
+  "compatible_path \<equiv> stream_all2 A_B.equiv' ===> (=)"
+
+named_theorems compatible
+
+lemma compatible_opI:
+  assumes "\<And>\<phi> \<psi>. compatible \<phi> \<psi> \<Longrightarrow> compatible (\<Phi> \<phi>) (\<Psi> \<psi>)"
+  shows "compatible_op \<Phi> \<Psi>"
+  using assms unfolding rel_fun_def by auto
+
+lemma compatible_opE:
+  assumes "compatible_op \<Phi> \<Psi>" "compatible \<phi> \<psi>"
+  shows "compatible (\<Phi> \<phi>) (\<Psi> \<psi>)"
+  using assms unfolding rel_fun_def by auto
+
+lemma Ex_ev_compatible[compatible, transfer_rule]:
+  "compatible_op A.Ex_ev B.Ex_ev"
+  using Ex_ev_iff unfolding rel_fun_def by blast
+
+lemma Alw_ev_compatible[compatible, transfer_rule]:
+  "compatible_op A.Alw_ev B.Alw_ev"
+  using Alw_ev_iff unfolding rel_fun_def by blast
+
+lemma Not_compatible[compatible]:
+  "compatible_op ((\<circ>) Not) ((\<circ>) Not)"
+  unfolding rel_fun_def by simp
+
+lemma compatible_op_compI:
+  assumes [transfer_rule]: "compatible_op \<Phi> \<Psi>" "compatible_op \<Phi>' \<Psi>'"
+  shows "compatible_op (\<Phi> o \<Phi>') (\<Psi> o \<Psi>')"
+  by transfer_prover
+
+lemma compatible_op_NotI:
+  assumes "compatible_op \<Phi> \<Psi>"
+  shows "compatible_op (\<lambda>\<phi> x. \<not> \<Phi> \<phi> x) (\<lambda>\<psi> x. \<not> \<Psi> \<psi> x)"
+  by (rule compatible_op_compI[unfolded comp_def],
+      rule Not_compatible[unfolded comp_def], rule assms)
+
+lemma
+  shows Alw_alw_compatible[compatible, transfer_rule]: "compatible_op A.Alw_alw B.Alw_alw"
+    and Ex_alw_compatible[compatible, transfer_rule]:  "compatible_op A.Ex_alw B.Ex_alw"
+  unfolding Graph_Defs.Alw_alw_iff Graph_Defs.Ex_alw_iff
+  by (rule compatible_op_NotI, rule compatible_op_compI[unfolded comp_def]; (rule compatible))+
+
+lemma conj_compatible:
+  "(compatible ===> compatible ===> compatible) (\<lambda>\<phi> \<phi>' x. \<phi> x \<and> \<phi>' x) (\<lambda>\<psi> \<psi>' x. \<psi> x \<and> \<psi>' x)"
+  by transfer_prover
+
+lemma implies_compatible:
+  "(compatible ===> compatible ===> compatible) (\<lambda>\<phi> \<phi>' x. \<phi> x \<longrightarrow> \<phi>' x) (\<lambda>\<psi> \<psi>' x. \<psi> x \<longrightarrow> \<psi>' x)"
+  by transfer_prover
+
+lemma disj_compatible:
+  "(compatible ===> compatible ===> compatible) (\<lambda>\<phi> \<phi>' x. \<phi> x \<or> \<phi>' x) (\<lambda>\<psi> \<psi>' x. \<psi> x \<or> \<psi>' x)"
+  by transfer_prover
+
+lemma Leadsto_compatible:
+  "(compatible ===> compatible ===> compatible) A.leadsto B.leadsto"
+  unfolding A.leadsto_def B.leadsto_def by transfer_prover
+
+lemmas [compatible] = implies_compatible[THEN rel_funD]
+
+definition
+  "protect x = x"
+
+lemma protect_cong[cong]:
+  "protect x = protect x"
+  unfolding protect_def ..
+
+lemmas [compatible] = Not_compatible[unfolded comp_def]
+
+lemma CTL_compatible:
+  assumes "rel_ctl_formula compatible \<phi> \<psi>"
+  shows "compatible (A.models_ctl \<phi>) (B.models_ctl \<psi>)"
+  using assms by induction (simp; rule compatible_opE, rule compatible, assumption)+
+
+lemma ctl_star_compatible_aux:
+  "(rel_state_formula compatible \<phi> \<phi>' \<longrightarrow> compatible (A.models_state \<phi>) (B.models_state \<phi>'))
+\<and> (rel_path_formula compatible \<psi> \<psi>' \<longrightarrow> compatible_path (A.models_path \<psi>) (B.models_path \<psi>'))"
+proof (induction rule: state_formula_path_formula.rel_induct)
+  case (All a b) \<comment> \<open>State\<close>
+  then show ?case
+    by - (drule holds_transfer[THEN rel_funD], unfold A.models_path.simps B.models_path.simps)
+next
+  case (ImpliesS a b) \<comment> \<open>All\<close>
+  then show ?case
+    apply simp
+    apply (intro rel_funI allI iffI impI)
+     apply (auto 4 4
+        dest!: B_A.simulation_run stream_all2_rotate_1 dest: rel_funD elim: equiv'_rotate_1; fail)
+    by (auto 4 4 dest!: A_B.simulation_run dest: rel_funD)
+next
+  case (NotS a12 b12) \<comment> \<open>Ex\<close>
+  then show ?case
+    apply simp
+    apply (intro rel_funI allI iffI impI)
+     apply (smt A_B.simulation_run rel_fun_def stream.rel_sel stream.sel(1) stream.sel(2))
+    by (smt B_A.simulation_run equiv'_rotate_1 rel_fun_def stream.rel_inject stream_all2_rotate_1)
+next
+  case (Until a22 b22) \<comment> \<open>F\<close>
+  then show ?case
+    by - (drule ev_transfer[THEN rel_funD], unfold A.models_path.simps B.models_path.simps)
+next
+  case (X a b) \<comment> \<open>G\<close>
+  then show ?case
+    by - (drule alw_transfer[THEN rel_funD], unfold A.models_path.simps B.models_path.simps)
+next
+  case (ImpliesP a b) \<comment> \<open>X\<close>
+  then show ?case
+    by - (drule nxt_transfer[THEN rel_funD], unfold A.models_path.simps B.models_path.simps)
+next
+  case (NotP a22 b22) \<comment> \<open>Until\<close>
+  then show ?case
+    by - (drule suntil_transfer[THEN rel_funD, THEN rel_funD],
+          unfold A.models_path.simps B.models_path.simps)
+qed (simp add: rel_fun_def; fail | auto; fail)+
+
+lemmas models_state_compatible = ctl_star_compatible_aux[THEN conjunct1, rule_format]
+   and models_path_compatible  = ctl_star_compatible_aux[THEN conjunct2, rule_format]
+
+lemma
+  "(compatible_path ===> compatible_path) alw alw"
+  by (rule alw_transfer)
+
+lemma
+  "(compatible_path ===> compatible_path) ev ev"
+  by (rule ev_transfer)
+
+lemma holds_compatible:
+  "(compatible ===> compatible_path) holds holds"
+  by (rule holds_transfer)
+
+lemma models_ltlp_compatible:
+  assumes "rel_pltl compatible \<psi> \<psi>'"
+  shows "compatible_path (A.models_ltlp \<psi>) (B.models_ltlp \<psi>')"
+  by (metis assms
+      A.ltlp_to_path_correct B.ltlp_to_path_correct models_path_compatible rel_ltlp_to_path)
+
+lemma models_ltlc_compatible:
+  assumes "rel_ltlc compatible \<psi> \<psi>'"
+  shows "compatible_path (A.models_ltlc \<psi>) (B.models_ltlc \<psi>')"
+  using assms unfolding A.models_ltlc_alt_def
+  by (intro models_ltlp_compatible) (simp only: rel_ltlc_to_pltl)
+
+end (* Transfer Syntax *)
+
+end (* Bisimulation Invariant *)
+
+lemmas [simp del] = holds.simps
+
+end (* Theory *)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Trace_Timing.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Trace_Timing.thy
@@ -0,0 +1,107 @@
+theory Trace_Timing
+  imports Refine_Imperative_HOL.Sepref
+begin
+
+datatype time = Time int
+
+definition now :: "unit \<Rightarrow> time" where
+  "now _ = undefined"
+
+instantiation time :: minus
+begin
+
+definition "minus_time (_ :: time) (_ :: time) = (undefined :: time)"
+
+instance ..
+
+end
+
+definition time_to_string :: "time \<Rightarrow> String.literal" where
+  "time_to_string _ = undefined"
+
+code_printing
+  constant "now" \<rightharpoonup> (SML) "Time.now _"
+
+code_printing
+  constant "time_to_string" \<rightharpoonup> (SML) "(fn x => Time.toString x) _"
+
+code_printing
+  constant "(-) :: time \<Rightarrow> time \<Rightarrow> time" \<rightharpoonup> (SML) "(fn x => fn y => Time.- (x, y)) _ _"
+
+code_printing code_module "Timing" \<rightharpoonup> (SML)
+\<open>
+structure Timing : sig
+  val start_timer: unit -> unit
+  val save_time: string -> unit
+  val get_timings: unit -> (string * Time.time) list
+end = struct
+  val t = Unsynchronized.ref Time.zeroTime;
+  val timings = Unsynchronized.ref [];
+  fun start_timer () = (t := Time.now ());
+  fun get_elapsed () = Time.- (Time.now (), !t);
+  fun save_time s = (timings := ((s, get_elapsed ()) :: !timings));
+  fun get_timings () = !timings;
+end
+\<close>
+
+definition start_timer :: "unit \<Rightarrow> unit" where
+  "start_timer _ = ()"
+
+definition save_time :: "String.literal \<Rightarrow> unit" where
+  "save_time s = ()"
+
+code_reserved (SML) Timing
+
+code_printing
+  constant start_timer \<rightharpoonup> (SML) "Timing.start'_timer _"
+
+code_printing
+  constant save_time \<rightharpoonup> (SML) "Timing.save'_time _"
+
+definition "test \<equiv> let _ = start_timer (); _ = save_time STR ''test'' in ()"
+
+export_code test checking SML
+
+hide_const test
+
+paragraph \<open>Setup for Sepref\<close>
+
+lemma [sepref_import_param]:
+  "(start_timer, start_timer) \<in> Id \<rightarrow> Id"
+  "(save_time, save_time) \<in> Id \<rightarrow> Id"
+  by simp+
+
+definition
+  "START_TIMER = RETURN o start_timer"
+
+definition
+  "start_timer_impl = return o start_timer"
+
+definition
+  "SAVE_TIME = RETURN o save_time"
+
+definition
+  "save_time_impl = return o save_time"
+
+lemma start_timer_hnr:
+  "hn_refine (hn_val Id x' x) (start_timer_impl x) (hn_val Id x' x) unit_assn (START_TIMER $ x')"
+  unfolding START_TIMER_def start_timer_impl_def by sepref_to_hoare sep_auto
+
+lemma save_time_hnr:
+  "(save_time_impl, SAVE_TIME) \<in> id_assnk \<rightarrow>a unit_assn"
+  unfolding SAVE_TIME_def save_time_impl_def by sepref_to_hoare sep_auto
+
+sepref_register START_TIMER SAVE_TIME
+
+lemmas [sepref_fr_rules] = start_timer_hnr save_time_hnr
+
+definition
+  "time_it s f = (let _ = start_timer (); r = f (); _ = save_time s in r)"
+
+lemma time_it:
+  "e = time_it s (\<lambda>_. e)"
+  unfolding Let_def time_it_def ..
+
+method time_it for s :: String.literal and t = (subst time_it[where s = s and e = t])
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Munta_Model_Checker/library/Tracing.thy
--- /dev/null
+++ thys/Munta_Model_Checker/library/Tracing.thy
@@ -0,0 +1,88 @@
+theory Tracing
+  imports Main Refine_Imperative_HOL.Sepref
+begin
+
+datatype message = ExploredState
+
+definition write_msg :: "message \<Rightarrow> unit" where
+  "write_msg m = ()"
+
+code_printing code_module "Tracing" \<rightharpoonup> (SML)
+\<open>
+structure Tracing : sig
+  val count_up : unit -> unit
+  val get_count : unit -> int
+end = struct
+  val counter = Unsynchronized.ref 0;
+  fun count_up () = (counter := !counter + 1);
+  fun get_count () = !counter;
+end
+\<close> and (OCaml)
+\<open>
+module Tracing : sig
+  val count_up : unit -> unit
+  val get_count : unit -> int
+end = struct
+  let counter = ref 0
+  let count_up () = (counter := !counter + 1)
+  let get_count () = !counter
+end
+\<close>
+
+code_reserved (SML) Tracing
+
+code_reserved (OCaml) Tracing
+
+code_printing
+  constant write_msg \<rightharpoonup> (SML) "(fn x => Tracing.count'_up ()) _"
+       and            (OCaml) "(fun x -> Tracing.count'_up ()) _"
+
+definition trace where
+  "trace m x = (let a = write_msg m in x)"
+
+lemma trace_alt_def[simp]:
+  "trace m x = (\<lambda> _. x) (write_msg x)"
+  unfolding trace_def by simp
+
+definition
+  "test m = trace ExploredState ((3 :: int) + 1)"
+
+definition "TRACE m = RETURN (trace m ())"
+
+lemma TRACE_bind[simp]:
+  "do {TRACE m; c} = c"
+  unfolding TRACE_def by simp
+
+lemma [sepref_import_param]:
+  "(trace, trace) \<in> \<langle>Id,\<langle>Id,Id\<rangle>fun_rel\<rangle>fun_rel"
+  by simp
+
+sepref_definition TRACE_impl is
+  "TRACE" :: "id_assnk \<rightarrow>a unit_assn"
+  unfolding TRACE_def by sepref
+
+lemmas [sepref_fr_rules] = TRACE_impl.refine
+
+
+text \<open>Somehow Sepref does not want to pick up TRACE as it is, so we use the following workaround:\<close>
+
+definition "TRACE' = TRACE ExploredState"
+
+definition "trace' = trace ExploredState"
+
+lemma TRACE'_alt_def:
+  "TRACE' = RETURN (trace' ())"
+  unfolding TRACE_def TRACE'_def trace'_def ..
+
+lemma [sepref_import_param]:
+  "(trace', trace') \<in> \<langle>Id,Id\<rangle>fun_rel"
+  by simp
+
+sepref_definition TRACE'_impl is
+  "uncurry0 TRACE'" :: "unit_assnk \<rightarrow>a unit_assn"
+  unfolding TRACE'_alt_def
+  by sepref
+
+lemmas [sepref_fr_rules] = TRACE'_impl.refine
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/ROOTS
--- thys/ROOTS
+++ thys/ROOTS
@@ -160,6 +160,7 @@
 Constructive_Cryptography_CM
 Constructor_Funs
 Containers
+Context_Free_Grammar
 Continued_Fractions
 Cook_Levin
 Coproduct_Measure
@@ -229,6 +230,7 @@
 Echelon_Form
 EdmondsKarp_Maxflow
 Edwards_Elliptic_Curves_Group
+Elementary_Ultrametric_Spaces
 Efficient-Mergesort
 Efficient_Weighted_Path_Order
 Elimination_Of_Repeated_Factors
@@ -264,6 +266,7 @@
 Factor_Algebraic_Polynomial
 Factored_Transition_System_Bounding
 Falling_Factorial_Sum
+Farey_Sequences
 Farkas
 FeatherweightJava
 Featherweight_OCL
@@ -341,6 +344,7 @@
 HOL-CSP
 HOL-CSPM
 HOL-CSP_OpSem
+HOL-CSP_RS
 HOLCF-Prelude
 HRB-Slicing
 Hahn_Jordan_Decomposition
@@ -543,6 +547,7 @@
 Multiset_Ordering_NPC
 Multitape_To_Singletape_TM
 Myhill-Nerode
+Munta_Model_Checker
 Name_Carrying_Type_Inference
 Nano_JSON
 Nash_Williams
@@ -710,6 +715,9 @@
 ResiduatedTransitionSystem2
 Residuated_Lattices
 Resolution_FOL
+Restriction_Spaces
+Restriction_Spaces-Examples
+Restriction_Spaces-Ultrametric
 Rewrite_Properties_Reduction
 Rewriting_Z
 Ribbon_Proofs
@@ -752,6 +760,7 @@
 Serializable
 Shadow_DOM
 Shadow_SC_DOM
+Shallow_Expressions
 Shivers-CFA
 Schoenhage_Strassen
 ShortestPath
@@ -771,6 +780,7 @@
 Smith_Normal_Form
 Smooth_Manifolds
 Solidity
+Sophie_Germain
 Sophomores_Dream
 Sort_Encodings
 Sorted_Terms
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/HOLCF/Restriction_Spaces-HOLCF.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/HOLCF/Restriction_Spaces-HOLCF.thy
@@ -0,0 +1,226 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+(*<*)
+theory "Restriction_Spaces-HOLCF"
+  imports HOLCF Restriction_Spaces
+begin
+
+default_sort type
+  (*>*)
+
+
+section \<open>Definitions\<close>
+
+interpretation belowRS : Restriction \<open>(\<down>)\<close> \<open>(\<sqsubseteq>)\<close> by unfold_locales
+    \<comment> \<open>Just to recover const\<open>belowRS.restriction_related_set\<close> and
+   const\<open>belowRS.restriction_not_related_set\<close>.\<close>
+
+class po_restriction_space = restriction + po +
+  assumes restriction_0_below [simp] : \<open>x \<down> 0 \<sqsubseteq> y \<down> 0\<close>
+    and mono_restriction_below     : \<open>x \<sqsubseteq> y \<Longrightarrow> x \<down> n \<sqsubseteq> y \<down> n\<close>
+    and ex_not_restriction_below   : \<open>\<not> x \<sqsubseteq> y \<Longrightarrow> \<exists>n. \<not> x \<down> n \<sqsubseteq> y \<down> n\<close>
+
+interpretation belowRS : PreorderRestrictionSpace \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a :: po_restriction_space\<close> \<open>(\<sqsubseteq>)\<close>
+proof unfold_locales
+  show \<open>x \<down> 0 \<sqsubseteq> y \<down> 0\<close> for x y :: 'a by simp
+next
+  show \<open>x \<sqsubseteq> y \<Longrightarrow> x \<down> n \<sqsubseteq> y \<down> n\<close> for x y :: 'a and n
+    by (fact mono_restriction_below)
+next
+  show \<open>\<not> x \<sqsubseteq> y \<Longrightarrow> \<exists>n. \<not> x \<down> n \<sqsubseteq> y \<down> n\<close> for x y :: 'a
+    by (simp add: ex_not_restriction_below)
+next
+  show \<open>x \<sqsubseteq> y \<Longrightarrow> y \<sqsubseteq> z \<Longrightarrow> x \<sqsubseteq> z\<close> for x y z :: 'a by (fact below_trans)
+qed
+
+
+subclass (in po_restriction_space) restriction_space
+proof unfold_locales
+  show \<open>x \<down> 0 = y \<down> 0\<close> for x y :: 'a by (rule below_antisym) simp_all
+next
+  show \<open>x \<noteq> y \<Longrightarrow> \<exists>n. x \<down> n \<noteq> y \<down> n\<close> for x y :: 'a
+    by (metis ex_not_restriction_below po_eq_conv)
+qed
+
+
+
+
+class cpo_restriction_space = po_restriction_space +
+  assumes cpo : \<open>chain S \<Longrightarrow> \<exists>x. range S <<| x\<close>
+
+
+subclass (in cpo_restriction_space) cpo
+  by unfold_locales (fact cpo)
+
+
+class pcpo_restriction_space = cpo_restriction_space +
+  assumes least : \<open>\<exists>x. \<forall>y. x \<sqsubseteq> y\<close>
+
+subclass (in pcpo_restriction_space) pcpo
+  by unfold_locales (fact least)
+
+
+interpretation belowRS : Restriction_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'a :: {restriction, below} \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<sqsubseteq>)\<close>
+  \<open>(\<down>) :: 'b :: po_restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<sqsubseteq>)\<close> ..
+
+text \<open>With this we recover constants like const\<open>less_eqRS.restriction_shift_on\<close>.\<close>
+
+interpretation belowRS : PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'a :: po_restriction_space \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<sqsubseteq>)\<close>
+  \<open>(\<down>) :: 'b :: po_restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<sqsubseteq>)\<close> ..
+
+text \<open>With that we recover theorems like @{thm belowRS.constructive_restriction_restriction}.\<close>
+
+interpretation belowRS : Restriction_2_PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'a :: {restriction, below} \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<sqsubseteq>)\<close>
+  \<open>(\<down>) :: 'b :: po_restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<sqsubseteq>)\<close>
+  \<open>(\<down>) :: 'c :: po_restriction_space \<Rightarrow> nat \<Rightarrow> 'c\<close> \<open>(\<sqsubseteq>)\<close>..
+
+text \<open>And with that we recover theorems like @{thm less_eqRS.constructive_on_comp_non_destructive_on}.\<close>
+
+
+
+context fixes f :: \<open>'a :: restriction \<Rightarrow> 'b :: po_restriction_space\<close> begin
+text \<open>From @{thm below_antisym} we can obtain stronger lemmas.\<close>
+
+corollary below_restriction_shift_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow>
+             f x \<down> nat (int n + k) \<sqsubseteq> f y \<down> nat (int n + k))
+   \<Longrightarrow> restriction_shift_on f k A\<close>
+  by (simp add: below_antisym restriction_shift_onI)
+
+corollary below_restriction_shiftI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow>
+             f x \<down> nat (int n + k) \<sqsubseteq> f y \<down> nat (int n + k))
+   \<Longrightarrow> restriction_shift f k\<close>
+  by (simp add: below_antisym restriction_shiftI)
+
+corollary below_non_too_destructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> Suc n = y \<down> Suc n\<rbrakk> \<Longrightarrow>
+             f x \<down> n \<sqsubseteq> f y \<down> n)
+   \<Longrightarrow> non_too_destructive_on f A\<close>
+  by (simp add: below_antisym non_too_destructive_onI)
+
+corollary below_non_too_destructiveI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> Suc n = y \<down> Suc n\<rbrakk> \<Longrightarrow> f x \<down> n \<sqsubseteq> f y \<down> n)
+   \<Longrightarrow> non_too_destructive f\<close>
+  by (simp add: below_antisym non_too_destructiveI)
+
+corollary below_non_destructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>n \<noteq> 0; x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> n \<sqsubseteq> f y \<down> n)
+   \<Longrightarrow> non_destructive_on f A\<close>
+  by (simp add: below_antisym non_destructive_onI)
+
+corollary below_non_destructiveI :
+  \<open>(\<And>x y n. \<lbrakk>n \<noteq> 0; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> n \<sqsubseteq> f y \<down> n)
+   \<Longrightarrow> non_destructive f\<close>
+  by (simp add: below_antisym non_destructiveI)
+
+corollary below_constructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> Suc n \<sqsubseteq> f y \<down> Suc n)
+   \<Longrightarrow> constructive_on f A\<close>
+  by (simp add: below_antisym constructive_onI)
+
+corollary below_constructiveI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> Suc n \<sqsubseteq> f y \<down> Suc n)
+   \<Longrightarrow> constructive f\<close>
+  by (simp add: below_antisym constructiveI)
+
+end
+
+
+
+section \<open>Equality of Fixed-Point Operators\<close>
+
+lemma restriction_fix_is_fix :
+  \<open>(\<upsilon> X. f X) = (\<mu> X. f X)\<close> if \<open>cont f\<close> \<open>constructive f\<close>
+for f :: \<open>'a :: {pcpo_restriction_space, complete_restriction_space} \<Rightarrow> 'a\<close>
+proof (rule restriction_fix_unique)
+  show \<open>constructive f\<close> by (fact \<open>constructive f\<close>)
+next
+  show \<open>f (\<mu> x. f x) = (\<mu> x. f x)\<close> by (metis def_cont_fix_eq \<open>cont f\<close>)
+qed
+
+
+
+section \<open>Product\<close>
+
+instance prod :: (po_restriction_space, po_restriction_space) po_restriction_space
+proof intro_classes
+  show \<open>p \<down> 0 \<sqsubseteq> q \<down> 0\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (metis below_refl restriction_0_related)
+next
+  show \<open>p \<sqsubseteq> q \<Longrightarrow> p \<down> n \<sqsubseteq> q \<down> n\<close> for p q :: \<open>'a \<times> 'b\<close> and n
+    by (simp add: below_prod_def restriction_prod_def mono_restriction_below)
+next
+  show \<open>\<not> p \<sqsubseteq> q \<Longrightarrow> \<exists>n. \<not> p \<down> n \<sqsubseteq> q \<down> n\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (simp add: below_prod_def restriction_prod_def)
+      (metis ex_not_restriction_below)
+qed
+
+
+instance prod :: (cpo_restriction_space, cpo_restriction_space) cpo_restriction_space
+  using is_lub_prod by intro_classes blast
+
+instance prod :: (pcpo_restriction_space, pcpo_restriction_space) pcpo_restriction_space
+  by (intro_classes) (simp add: pcpo_class.least)
+
+
+section \<open>Functions\<close>
+
+instance \<open>fun\<close> :: (type, po_restriction_space) po_restriction_space
+proof intro_classes
+  show \<open>f \<down> 0 \<sqsubseteq> g \<down> 0\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (metis below_refl restriction_0_related)
+next
+  show \<open>f \<sqsubseteq> g \<Longrightarrow> f \<down> n \<sqsubseteq> g \<down> n\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close> and n
+    by (simp add: fun_below_iff mono_restriction_below restriction_fun_def)
+next
+  show \<open>\<not> f \<sqsubseteq> g \<Longrightarrow> \<exists>n. \<not> f \<down> n \<sqsubseteq> g \<down> n\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (metis belowRS.all_ge_restriction_related_iff_related fun_below_iff restriction_fun_def)
+qed
+
+instance \<open>fun\<close> :: (type, cpo_restriction_space) cpo_restriction_space
+  by intro_classes (simp add: cpo_class.cpo)
+
+instance \<open>fun\<close> :: (type, pcpo_restriction_space) pcpo_restriction_space
+  by intro_classes (simp add: pcpo_class.least)
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/HOLCF/document/root.tex
--- /dev/null
+++ thys/Restriction_Spaces-Examples/HOLCF/document/root.tex
@@ -0,0 +1,65 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{latexsym}
+\usepackage{amssymb}
+\usepackage{amsmath}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{HOLCF support for restriction Spaces}
+\author{Benoît Ballenghien \and Benjamin Puyobro \and Burkhart Wolff}
+\maketitle
+
+\abstract{
+  In this session, we provide support for restriction spaces that are also equipped with a partial order structure inherited from \verb'HOLCF'.
+}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/ROOT
--- /dev/null
+++ thys/Restriction_Spaces-Examples/ROOT
@@ -0,0 +1,66 @@
+chapter AFP
+
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+
+session "Restriction_Spaces-Examples" = "HOL-Analysis" +
+  description "Miscellaneous examples of restriction spaces."
+  options [timeout = 300]
+  sessions Restriction_Spaces
+  theories
+    RS_Any_Type
+    RS_Bool
+    RS_Nat
+    RS_Int
+    RS_Option
+    RS_List
+    RS_Tree
+    RS_Decimals
+    RS_Trace_Model_CSP
+    RS_Formal_Power_Series
+  document_files "root.tex"
+
+
+
+
+session "Restriction_Spaces-HOLCF" in HOLCF = HOLCF +
+  description "HOLCF support for restriction spaces."
+  options [timeout = 300]
+  sessions Restriction_Spaces
+  theories "Restriction_Spaces-HOLCF" (global)
+  document_files "root.tex"
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Any_Type.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Any_Type.thy
@@ -0,0 +1,89 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Trivial Construction\<close>
+
+(*<*)
+theory RS_Any_Type
+  imports Restriction_Spaces
+begin 
+  (*>*)
+
+
+text \<open>Restriction instance for any type.\<close>
+
+typedef 'a type' = \<open>UNIV :: 'a set\<close> by auto
+
+instantiation type' :: (type) restriction
+begin
+
+lift_definition restriction_type' :: \<open>'a type' \<Rightarrow> nat \<Rightarrow> 'a type'\<close>
+  is \<open>\<lambda>x n. if n = 0 then undefined else x\<close> .
+
+instance by (intro_classes, transfer, simp add: min_def)
+
+end
+
+
+lemma restriction_type'_0_is_undefined [simp] :
+  \<open>x \<down> 0 = undefined\<close> for x :: \<open>'a type'\<close> by transfer simp
+
+
+instance type' :: (type) restriction_space
+  by (intro_classes, simp, transfer, auto)
+
+
+lemma restriction_tendsto_type'_iff :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<exists>n0. \<forall>n\<ge>n0. \<sigma> n = \<Sigma>)\<close> for \<Sigma> :: \<open>'a type'\<close>
+  by (simp add: restriction_tendsto_def, transfer, auto)
+
+
+lemma restriction_chain_type'_iff :
+  \<open>chain\<down> \<sigma> \<longleftrightarrow> \<sigma> 0 = undefined \<and> (\<forall>n\<ge>Suc 0. \<sigma> n = \<sigma> (Suc 0))\<close>
+  for \<sigma> :: \<open>nat \<Rightarrow> 'a type'\<close>
+  by (simp add: restriction_chain_def_ter, transfer, simp)
+    (safe, (simp_all)[3], metis Suc_le_D Suc_le_eq zero_less_Suc)
+
+
+
+instance type' :: (type) complete_restriction_space
+  by intro_classes
+    (auto simp add: restriction_chain_type'_iff restriction_convergent_def
+      restriction_tendsto_type'_iff)
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Bool.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Bool.thy
@@ -0,0 +1,101 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Booleans\<close>
+
+(*<*)
+theory RS_Bool
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+text \<open>Restriction instance for typ\<open>bool\<close>.\<close>
+
+instantiation bool :: restriction
+begin
+
+definition restriction_bool :: \<open>bool \<Rightarrow> nat \<Rightarrow> bool\<close>
+  where \<open>b \<down> n \<equiv> if n = 0 then False else b\<close>
+
+instance by (intro_classes) (auto simp add: restriction_bool_def)
+end
+
+
+lemma restriction_bool_0_is_False [simp] : \<open>b \<down> 0 = False\<close>
+  by (simp add: restriction_bool_def)
+
+
+text \<open>Restriction space instance for typ\<open>bool\<close>.\<close>
+
+instance bool :: restriction_space
+  by intro_classes (simp_all add: restriction_bool_def gt_ex)
+
+
+
+text \<open>Complete Restriction space instance for typ\<open>bool\<close>.\<close>
+
+lemma restriction_tendsto_bool_iff :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<exists>n. \<forall>k\<ge>n. \<sigma> k = \<Sigma>)\<close> for \<Sigma> :: bool
+  unfolding restriction_tendsto_def
+  by (auto simp add: restriction_bool_def)
+
+
+instance bool :: complete_restriction_space
+proof intro_classes
+  fix \<sigma> :: \<open>nat \<Rightarrow> bool\<close> assume \<open>chain\<down> \<sigma>\<close>
+  hence \<open>(\<forall>n>0. \<not> \<sigma> n) \<or> (\<forall>n>0. \<sigma> n)\<close>
+    by (simp add: restriction_chain_def restriction_bool_def split: if_split_asm)
+      (metis One_nat_def Zero_not_Suc gr0_conv_Suc nat_induct_non_zero zero_induct)
+  hence \<open>\<sigma> \<midarrow>\<down>\<rightarrow> False \<or> \<sigma> \<midarrow>\<down>\<rightarrow> True\<close>
+    by (metis (full_types) gt_ex order.strict_trans2 restriction_tendsto_def)
+  thus \<open>convergent\<down> \<sigma>\<close>
+    using restriction_convergentI by blast
+qed
+
+
+
+lemma restriction_cont_imp_restriction_adm :
+  \<open>cont\<down> P \<Longrightarrow> adm\<down> P\<close> for P :: \<open>'a :: restriction_space \<Rightarrow> bool\<close>
+  unfolding restriction_adm_def restriction_cont_on_def restriction_cont_at_def
+  by (auto simp add: restriction_tendsto_bool_iff)
+
+
+lemma restriction_compact_bool : \<open>compact\<down> (UNIV :: bool set)\<close>
+  by (simp add: finite_imp_restriction_compact)
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Decimals.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Decimals.thy
@@ -0,0 +1,256 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Decimals of a Number\<close>
+
+(*<*)
+theory RS_Decimals
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+typedef (overloaded) 'a :: zero decimals = \<open>{\<sigma> :: nat \<Rightarrow> 'a. \<sigma> 0 = 0}\<close>
+  morphisms from_decimals to_decimals by auto
+
+setup_lifting type_definition_decimals
+
+
+declare from_decimals [simp] to_decimals_cases[simp]
+  to_decimals_inject[simp] to_decimals_inverse [simp]
+
+
+declare from_decimals_inject  [simp]
+  from_decimals_inverse [simp]
+
+lemmas to_decimals_inject_simplified [simp] = to_decimals_inject [simplified]
+  and to_decimals_inverse_simplified[simp] = to_decimals_inverse[simplified]
+
+(* useful ? *)
+lemmas to_decimals_induct_simplified = to_decimals_induct[simplified]
+  and to_decimals_cases_simplified  = to_decimals_cases [simplified]
+  and from_decimals_induct_simplified = from_decimals_induct[simplified]
+  and from_decimals_cases_simplified  = from_decimals_cases [simplified]
+
+
+
+instantiation decimals :: (zero) restriction
+begin
+
+lift_definition restriction_decimals :: \<open>'a decimals \<Rightarrow> nat \<Rightarrow> 'a decimals\<close> 
+  is \<open>\<lambda>\<sigma> m n. if n \<le> m then \<sigma> n else 0\<close> by simp
+
+instance by (intro_classes, transfer, rule ext, simp)
+
+end
+
+
+instance decimals :: (zero) restriction_space
+  by (intro_classes; transfer, auto)
+    (metis (no_types, lifting) ext order_refl)
+
+
+lemma restriction_decimals_eq_iff :
+  \<open>x \<down> n = y \<down> n \<longleftrightarrow> (\<forall>i\<le>n. from_decimals x i = from_decimals y i)\<close>
+  by transfer meson
+
+
+lemma restriction_decimals_eqI :
+  \<open>(\<And>i. i \<le> n \<Longrightarrow> from_decimals x i = from_decimals y i) \<Longrightarrow> x \<down> n = y \<down> n\<close>
+  by (simp add: restriction_decimals_eq_iff)
+
+lemma restriction_decimals_eqD :
+  \<open>x \<down> n = y \<down> n \<Longrightarrow> i \<le> n \<Longrightarrow> from_decimals x i = from_decimals y i\<close>
+  by (simp add: restriction_decimals_eq_iff)
+
+
+text \<open>This space is actually complete.\<close>
+
+instance decimals :: (zero) complete_restriction_space
+proof (intro_classes, rule restriction_convergentI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a decimals\<close> assume \<open>chain\<down> \<sigma>\<close>
+  let ?\<Sigma> = \<open>to_decimals (\<lambda>n. from_decimals (\<sigma> n) n)\<close>
+  have \<open>?\<Sigma> \<down> n = \<sigma> n\<close> for n
+  proof (subst restricted_restriction_chain_is[OF \<open>chain\<down> \<sigma>\<close>, symmetric],
+      rule restriction_decimals_eqI)
+    fix i assume \<open>i \<le> n\<close>
+    from restriction_chain_def_ter
+      [THEN iffD1, OF \<open>restriction_chain \<sigma>\<close>, rule_format, OF \<open>i \<le> n\<close>]
+    show \<open>from_decimals ?\<Sigma> i = from_decimals (\<sigma> n) i\<close>
+      by (subst to_decimals_inverse_simplified, use from_decimals in blast)
+        (metis dual_order.refl restriction_decimals.rep_eq)
+  qed
+  thus \<open>restriction_chain \<sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> ?\<Sigma>\<close> 
+  proof -
+    have \<open>(\<down>) (to_decimals (\<lambda>n. from_decimals (\<sigma> n) n)) = \<sigma>\<close>
+      using \<open>\<And>n. to_decimals (\<lambda>n. from_decimals (\<sigma> n) n) \<down> n = \<sigma> n\<close> by force
+    then show ?thesis
+      by (metis restriction_tendsto_restrictions)
+  qed
+qed
+
+
+
+
+typedef nat_0_9 = \<open>{0.. 9::nat}\<close>
+  morphisms from_nat_0_9 to_nat_0_9 by auto
+
+
+setup_lifting type_definition_nat_0_9
+
+
+instantiation nat_0_9 :: zero
+begin
+
+lift_definition zero_nat_0_9 :: nat_0_9 is 0 by simp
+
+instance ..
+
+end
+
+instantiation nat_0_9 :: one
+begin
+
+lift_definition one_nat_0_9 :: nat_0_9 is 1 by simp
+
+instance ..
+
+end
+
+
+lift_definition update_nth_decimal :: \<open>[nat_0_9 decimals, nat, nat] \<Rightarrow> nat_0_9 decimals\<close>
+  is \<open>\<lambda>s index value.   if index = 0 \<or> 9 < value then from_decimals s
+                      else (from_decimals s)(index := to_nat_0_9 value)\<close>
+  using from_decimals by auto
+
+
+lemma no_update_nth_decimal [simp] :
+  \<open>index = 0 \<Longrightarrow> update_nth_decimal s index val = s\<close>
+  \<open>9 < val   \<Longrightarrow> update_nth_decimal s index val = s\<close>
+  by (simp_all add: update_nth_decimal.abs_eq)
+
+
+lemma non_destructive_update_nth_decimal : \<open>non_destructive update_nth_decimal\<close>
+proof (rule non_destructiveI)
+  show \<open>update_nth_decimal x \<down> n = update_nth_decimal y \<down> n\<close> if \<open>x \<down> n = y \<down> n\<close> for x y n
+  proof (unfold restriction_fun_def, intro ext restriction_decimals_eqI)
+    fix index val i assume \<open>i \<le> n\<close>
+    from restriction_decimals_eqD[OF \<open>x \<down> n = y \<down> n\<close> \<open>i \<le> n\<close>]
+    show \<open>from_decimals (update_nth_decimal x index val) i =
+          from_decimals (update_nth_decimal y index val) i\<close>
+      by (simp add: update_nth_decimal.rep_eq)
+  qed
+qed
+
+
+lift_definition shift_decimal_right :: \<open>nat_0_9 decimals \<Rightarrow> nat_0_9 decimals\<close>
+  is \<open>\<lambda>s n. case n of 0 \<Rightarrow> to_nat_0_9 0 | Suc n' \<Rightarrow> from_decimals s n'\<close>
+  by (simp add: zero_nat_0_9_def)
+
+
+lemma constructive_shift_decimal_right : \<open>constructive shift_decimal_right\<close>
+proof (rule constructiveI)
+  show \<open>shift_decimal_right x \<down> Suc n = shift_decimal_right y \<down> Suc n\<close> if \<open>x \<down> n = y \<down> n\<close> for x y n
+  proof (intro restriction_decimals_eqI)
+    fix index val i assume \<open>i \<le> Suc n\<close>
+    hence \<open>i - 1 \<le> n\<close> by simp
+    from restriction_decimals_eqD[OF \<open>x \<down> n = y \<down> n\<close> \<open>i - 1 \<le> n\<close>]
+    show \<open>from_decimals (shift_decimal_right x) i = from_decimals (shift_decimal_right y) i\<close>
+      by (simp add: shift_decimal_right.rep_eq Nitpick.case_nat_unfold)
+  qed
+qed
+
+
+lift_definition shift_decimal_left :: \<open>nat_0_9 decimals \<Rightarrow> nat_0_9 decimals\<close>
+  is \<open>\<lambda>s n. if n = 0 then to_nat_0_9 0 else from_decimals s (Suc n)\<close>
+  by (simp add: zero_nat_0_9_def)
+
+lemma non_too_destructive_shift_decimal_left : \<open>non_too_destructive shift_decimal_left\<close>
+proof (rule non_too_destructiveI)
+  show \<open>shift_decimal_left x \<down> n = shift_decimal_left y \<down> n\<close> if \<open>x \<down> Suc n = y \<down> Suc n\<close> for x y n
+  proof (intro restriction_decimals_eqI)
+    fix index val i assume \<open>i \<le> n\<close>
+    hence \<open>Suc i \<le> Suc n\<close> by simp
+    from restriction_decimals_eqD[OF \<open>x \<down> Suc n = y \<down> Suc n\<close> \<open>Suc i \<le> Suc n\<close>]
+    show \<open>from_decimals (shift_decimal_left x) i = from_decimals (shift_decimal_left y) i\<close>
+      by (simp add: shift_decimal_left.rep_eq)
+  qed
+qed
+
+
+
+lemma restriction_fix_shift_decimal_right : \<open>(\<upsilon> x. shift_decimal_right x) = to_decimals (\<lambda>_. 0)\<close>
+proof (rule restriction_fix_unique)
+  show \<open>constructive shift_decimal_right\<close> 
+    by (fact constructive_shift_decimal_right)
+next
+  show \<open>shift_decimal_right (to_decimals (\<lambda>_. 0)) = to_decimals (\<lambda>_. 0)\<close>
+    by (simp add: shift_decimal_right.abs_eq)
+      (metis nat.case_eq_if to_decimals_inverse_simplified zero_nat_0_9_def)
+qed
+
+
+
+text \<open>Example of a predicate that is not admissible.\<close>
+
+lemma one_in_decimals_not_admissible :
+  defines P_def: \<open>P \<equiv> \<lambda>x. (1 :: nat_0_9) \<in> range (from_decimals x)\<close>
+  shows \<open>\<not> adm\<down> P\<close>
+proof (rule ccontr)
+  assume * : \<open>\<not> \<not> adm\<down> P\<close>
+
+  define x where \<open>x \<equiv> to_decimals (\<lambda>n. if n = 0 then 0 else 1 :: nat_0_9)\<close>
+
+  have \<open>P (\<upsilon> x. shift_decimal_right x)\<close>
+  proof (induct rule: restriction_fix_ind)
+    show \<open>constructive shift_decimal_right\<close> by (fact constructive_shift_decimal_right)
+  next
+    from "*" show \<open>adm\<down> P\<close> by simp
+  next
+    show \<open>P x\<close> by (auto simp add: P_def x_def image_iff)
+  next
+    show \<open>P x \<Longrightarrow> P (shift_decimal_right x)\<close> for x
+      by (simp add: P_def image_def shift_decimal_right.rep_eq)
+        (metis old.nat.simps(5))
+  qed
+  moreover have \<open>\<not> P (\<upsilon> x. shift_decimal_right x)\<close>
+    by (simp add: P_def restriction_fix_shift_decimal_right)
+      (transfer, simp)
+  ultimately show False by blast
+qed
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Formal_Power_Series.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Formal_Power_Series.thy
@@ -0,0 +1,145 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Formal power Series\<close>
+
+(*<*)
+theory RS_Formal_Power_Series
+  imports Restriction_Spaces "HOL-Analysis.FPS_Convergence"
+begin
+  (*>*)
+
+instantiation fps ::("{comm_ring_1}") restriction_space begin
+definition restriction_fps :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a fps"  
+  where \<open>restriction_fps a n \<equiv> \<Sum>i<n. fps_const (fps_nth a i)*fps_X^i\<close>
+
+lemma intersection_equality:\<open>(n::nat) \<le> m \<Longrightarrow> {..<m} \<inter> {i. i < n} = {i. i<n}\<close>
+  by auto
+
+lemma exist_noneq:\<open>x \<noteq> y \<Longrightarrow>
+           \<exists>n. (\<Sum>i\<in>{x. x < n}. fps_const (fps_nth x i) * fps_X ^ i) \<noteq>
+               (\<Sum>i\<in>{x. x < n}. fps_const (fps_nth y i) * fps_X ^ i)\<close> for x y::\<open>'a fps\<close>
+proof -
+  assume \<open>x\<noteq>y\<close>
+  then have \<open>\<exists>n. (n = (LEAST n. fps_nth x n \<noteq> fps_nth y n))\<close> 
+    using fps_nth_inject by blast
+  then obtain n where \<open> (n = (LEAST n. fps_nth x n \<noteq> fps_nth y n))\<close> by blast
+  then have \<open>\<forall>i<n. fps_nth x i = fps_nth y i\<close>
+    using not_less_Least by blast
+  then have f0:\<open>(\<Sum>i<n. fps_const (fps_nth x i) * fps_X ^ i) = (\<Sum>i<n. fps_const (fps_nth y i) * fps_X ^ i)\<close>
+    by(auto)
+  have rule:\<open>a\<noteq>c \<Longrightarrow> a+b \<noteq> c+b\<close> for a b c::"'a fps"
+    unfolding fps_plus_def using fps_ext 
+    by(auto simp:fun_eq_iff fps_ext Abs_fps_inverse fps_nth_inverse Abs_fps_inject) 
+  have \<open>fps_nth x n \<noteq> fps_nth y n\<close> 
+    by (metis (mono_tags, lifting) LeastI_ex \<open>n = (LEAST n. fps_nth x n \<noteq> fps_nth y n)\<close> \<open>x \<noteq> y\<close> fps_ext)
+  then have \<open> (\<Sum>i<n. fps_const (fps_nth x i) * fps_X ^ i) + fps_const (fps_nth x n) * fps_X ^ n \<noteq>
+    (\<Sum>i<n. fps_const (fps_nth y i) * fps_X ^ i) + fps_const (fps_nth y n) * fps_X ^ n\<close>
+    by (metis (no_types, lifting) f0 add_left_imp_eq fps_nth_fps_const fps_shift_times_fps_X_power') 
+  moreover have \<open> (\<Sum>i\<in>{x. x < Suc n}. fps_const (fps_nth z i) * fps_X ^ i) 
+= (\<Sum>i<n. fps_const (fps_nth z i) * fps_X ^ i) + fps_const (fps_nth z n) * fps_X ^ n\<close> for z::"'a fps"
+  proof -
+    have "\<forall>n. {..<n::nat} = {na. na < n}"
+      by (simp add: lessThan_def)
+    then show ?thesis
+      using sum.lessThan_Suc by auto
+  qed
+  ultimately show \<open>\<exists>n. (\<Sum>i\<in>{x. x < n}. fps_const (fps_nth x i) * fps_X ^ i) \<noteq>
+               (\<Sum>i\<in>{x. x < n}. fps_const (fps_nth y i) * fps_X ^ i)\<close>
+    by(auto intro: exI[where x=\<open>Suc n\<close>]) 
+qed
+
+
+instance
+  using intersection_equality exist_noneq 
+  by intro_classes
+    (auto cong: if_cong simp add:
+     Collect_mono Int_absorb2 restriction_fps_def fps_sum_nth
+     if_distrib[where f=fps_const]  if_distrib[where f=\<open>\<lambda>x. a * x\<close> for a]
+     if_distrib[where f=\<open>\<lambda>x. x * a\<close> for a] lessThan_def sum.If_cases min_def)
+
+end                                                           
+
+lemma fps_sum_rep_nthb: "fps_nth (\<Sum>i<m. fps_const(a i)*fps_X^i) n = (if n < m then a n else 0)"
+  by (simp add: fps_sum_nth if_distrib cong del: if_weak_cong)
+
+lemma restriction_eq_iff :\<open>a\<down>n = b\<down>n \<longleftrightarrow> (\<forall>i<n. fps_nth a i = fps_nth b i)\<close>
+  by (auto simp:restriction_fps_def) 
+    (metis (full_types) fps_sum_rep_nthb)
+
+
+lemma restriction_eqI :
+  \<open>(\<And>i. i < n \<Longrightarrow> fps_nth x i = fps_nth y i) \<Longrightarrow> x \<down> n = y \<down> n\<close>
+  by (simp add: restriction_eq_iff)
+
+lemma restriction_eqI' :
+  \<open>(\<And>i. i \<le> n \<Longrightarrow> fps_nth x i = fps_nth y i) \<Longrightarrow> x \<down> n = y \<down> n\<close>
+  by (simp add: restriction_eq_iff)
+
+instantiation fps :: (comm_ring_1) complete_restriction_space 
+begin
+instance
+proof (intro_classes, rule restriction_convergentI)
+  fix \<sigma> :: \<open>nat \<Rightarrow>'a fps\<close> assume h:\<open>restriction_chain \<sigma>\<close>
+  have h':\<open>\<forall>n. (\<Sum>i<n. fps_const (fps_nth (\<sigma> (Suc n)) i) * fps_X ^ i) = \<sigma> n\<close>
+    using h unfolding restriction_chain_def restriction_fps_def by auto
+  let ?\<Sigma> = \<open>Abs_fps (\<lambda>n. fps_nth (\<sigma> (Suc n)) n)\<close>
+  have \<open>?\<Sigma> \<down> (n) = \<sigma> n\<close> for n
+  proof (subst restricted_restriction_chain_is[OF \<open>restriction_chain \<sigma>\<close>, symmetric],
+      rule restriction_eqI)
+    fix i assume \<open>i < n\<close> 
+    then have \<open>i \<le> n\<close> by auto
+    from restriction_chain_def_ter
+      [THEN iffD1, OF \<open>restriction_chain \<sigma>\<close>, rule_format, OF \<open>i \<le> n\<close>]
+    show \<open>fps_nth (Abs_fps (\<lambda>n. fps_nth (\<sigma> (Suc n)) n)) i = fps_nth (\<sigma> n) i\<close>
+      by (subst Abs_fps_inverse, use Abs_fps_inject restriction_fps_def in blast)
+        (smt (verit, ccfv_threshold) Suc_leI \<open>i < n\<close> h le_refl lessI restriction_eq_iff 
+          restriction_chain_def restriction_chain_def_ter)     
+  qed
+  thus \<open>restriction_chain \<sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> ?\<Sigma>\<close> 
+  proof -
+    have "(\<down>) (Abs_fps (\<lambda>n. fps_nth (\<sigma> (Suc n)) n)) = \<sigma>"
+      using \<open>\<And>n. Abs_fps (\<lambda>n. fps_nth (\<sigma> (Suc n)) n) \<down> n = \<sigma> n\<close> by force
+    then show ?thesis
+      by (metis restriction_tendsto_restrictions)
+  qed
+qed
+
+end
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Int.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Int.thy
@@ -0,0 +1,80 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Integers\<close>
+
+(*<*)
+theory RS_Int
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+instantiation int :: restriction
+begin
+
+definition restriction_int :: \<open>int \<Rightarrow> nat \<Rightarrow> int\<close>
+  where \<open>x \<down> n \<equiv> if \<bar>x\<bar> \<le> int n then x else if 0 \<le> x then int n else - int n\<close>
+
+instance by intro_classes (simp add: restriction_int_def min_def)
+
+end
+
+
+instance int :: restriction_space
+  by (intro_classes, simp_all add: restriction_int_def)
+    (metis le_eq_less_or_eq linorder_not_less nat_le_iff)
+
+
+
+lemma restriction_int_0_is_0 [simp] : \<open>x \<down> 0 = (0 :: int)\<close>
+  by (simp add: restriction_int_def)
+
+
+text \<open>Restriction shift plus\<close>
+
+lemma restriction_shift_on_pos_plus : \<open>restriction_shift_on (\<lambda>x. x + k) k {x. 0 \<le> x}\<close>
+  by (intro restriction_shift_onI)
+    (simp add: restriction_int_def split: if_split_asm)
+
+lemma restriction_shift_on_neg_minus : \<open>restriction_shift_on (\<lambda>x. x - k) k {x. x \<le> 0}\<close>
+  by (intro restriction_shift_onI)
+    (simp add: restriction_int_def split: if_split_asm)
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_List.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_List.thy
@@ -0,0 +1,95 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Lists\<close>
+
+(*<*)
+theory RS_List
+  imports Restriction_Spaces "HOL-Library.Prefix_Order"
+begin
+  (*>*)
+
+text \<open>List is a restriction space using const\<open>take\<close> as the restriction function\<close>
+
+instantiation list :: (type) restriction
+begin
+
+definition restriction_list :: \<open>'a list \<Rightarrow> nat \<Rightarrow> 'a list\<close> 
+  where \<open>L \<down> n \<equiv> take n L\<close>
+
+instance by intro_classes (simp add: restriction_list_def min.commute)
+
+end
+
+
+instance list :: (type) order_restriction_space
+proof intro_classes
+  show \<open>L \<down> 0 \<le> M \<down> 0\<close> for L M :: \<open>'a list\<close>
+    by (simp add: restriction_list_def)
+next
+  show \<open>L \<le> M \<Longrightarrow> L \<down> n \<le> M \<down> n\<close> for L M :: \<open>'a list\<close> and n
+    unfolding restriction_list_def
+    by (metis less_eq_list_def prefix_def take_append)
+next
+  show \<open>\<not> L \<le> M \<Longrightarrow> \<exists>n. \<not> L \<down> n \<le> M \<down> n\<close> for M L :: \<open>'a list\<close>
+    unfolding restriction_list_def
+    by (metis linorder_linear take_all_iff)
+qed
+
+
+lemma \<open>OFCLASS('a list, restriction_space_class)\<close> ..
+
+
+
+text \<open>Of course, this space is not complete. We prove this with by exhibiting a counter-example.\<close>
+
+notepad begin
+  define \<sigma> :: \<open>nat \<Rightarrow> 'a list\<close>
+    where \<open>\<sigma> n = replicate n undefined\<close> for n
+
+  have \<open>chain\<down> \<sigma>\<close>
+    by (intro restriction_chainI ext)
+      (simp add: \<sigma>_def restriction_list_def flip: replicate_append_same)
+
+  hence \<open>\<nexists>\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (metis \<sigma>_def convergent_restriction_chain_imp_ex1 length_replicate
+        lessI nat_less_le restriction_convergentI restriction_list_def take_all)
+
+end
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Nat.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Nat.thy
@@ -0,0 +1,104 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Naturals\<close>
+
+(*<*)
+theory RS_Nat
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+text \<open>Restriction instance for typ\<open>nat\<close>.\<close>
+
+instantiation nat :: restriction
+begin
+
+definition restriction_nat :: \<open>nat \<Rightarrow> nat \<Rightarrow> nat\<close>
+  where \<open>x \<down> n \<equiv> if x \<le> n then x else n\<close>
+
+instance by intro_classes (simp add: restriction_nat_def)
+
+end
+
+
+lemma restriction_nat_0_is_0 [simp] : \<open>x \<down> 0 = (0 :: nat)\<close>
+  by (simp add: restriction_nat_def)
+
+
+text \<open>Restriction Space instance for typ\<open>nat\<close>.\<close>
+
+instance nat :: restriction_space
+  by intro_classes (use nat_le_linear in \<open>auto simp add: restriction_nat_def\<close>)
+
+
+
+text \<open>Constructive Suc\<close>
+
+lemma constructive_Suc : \<open>constructive Suc\<close>
+proof (rule constructiveI)
+  show \<open>x \<down> n = y \<down> n \<Longrightarrow> Suc x \<down> Suc n = Suc y \<down> Suc n\<close> for x y n
+    by (simp add: restriction_nat_def split: if_split_asm)
+qed
+
+
+text \<open>Non too destructive pred\<close>
+
+lemma non_too_destructive_pred : \<open>non_too_destructive nat.pred\<close>
+proof (rule non_too_destructiveI)
+  show \<open>x \<down> Suc n = y \<down> Suc n \<Longrightarrow> nat.pred x \<down> n = nat.pred y \<down> n\<close> for x y n
+    by (cases x; cases y) (simp_all add: restriction_nat_def split: if_split_asm)
+qed
+
+
+text \<open>Restriction shift plus\<close>
+
+lemma restriction_shift_plus : \<open>restriction_shift (\<lambda>x. x + k) (int k)\<close>
+proof (intro restriction_shiftI)
+  show \<open>x \<down> n = y \<down> n \<Longrightarrow> x + k \<down> nat (int n + int k) = y + k \<down> nat (int n + int k)\<close> for x y n
+    by (simp add: restriction_nat_def nat_int_add split: if_split_asm)
+qed
+
+lemma \<open>restriction_shift (\<lambda>x. k + x) (int k)\<close>
+  by (simp add: add.commute restriction_shift_plus)
+
+\<comment> \<open>In particular, constructive if term\<open>1 < k\<close>.\<close>
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Option.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Option.thy
@@ -0,0 +1,142 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Option Type\<close>
+
+(*<*)
+theory RS_Option
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+subsection \<open>Restriction option type\<close>
+
+instantiation option :: (restriction) restriction
+begin
+
+definition restriction_option :: \<open>'a option \<Rightarrow> nat \<Rightarrow> 'a option\<close>
+  where \<open>x \<down> n \<equiv> if n = 0 then None else map_option (\<lambda>a. a \<down> n) x\<close>
+
+instance
+  by intro_classes
+    (simp add: restriction_option_def option.map_comp comp_def min_def)
+
+end
+
+lemma restriction_option_0_is_None [simp] : \<open>x \<down> 0 = None\<close>
+  by (simp add: restriction_option_def)
+
+lemma restriction_option_None [simp] : \<open>None \<down> n = None\<close>
+  by (simp add: restriction_option_def)
+
+lemma restriction_option_Some [simp] : \<open>Some x \<down> n = (if n = 0 then None else Some (x \<down> n))\<close>
+  by (simp add: restriction_option_def)
+
+lemma restriction_option_eq_None_iff : \<open>x \<down> n = None \<longleftrightarrow> n = 0 \<or> x = None\<close>
+  by (cases x) simp_all
+
+lemma restriction_option_eq_Some_iff : \<open>x \<down> n = Some y \<longleftrightarrow> n \<noteq> 0 \<and> x \<noteq> None \<and> y = the x \<down> n\<close>
+  by (cases x) auto
+
+
+
+subsection \<open>Restriction space option type\<close>
+
+instance option :: (restriction_space) restriction_space
+proof intro_classes
+  show \<open>x \<down> 0 = y \<down> 0\<close> for x y :: \<open>'a option\<close> by simp
+next
+  show \<open>x \<noteq> y \<Longrightarrow> \<exists>n. x \<down> n \<noteq> y \<down> n\<close> for x y :: \<open>'a option\<close>
+    by (cases x; cases y, simp_all add: gt_ex)
+      (metis bot_nat_0.not_eq_extremum ex_not_restriction_related restriction_0_related)
+qed
+
+
+
+subsection \<open>Complete restriction space option type\<close>
+
+lemma option_restriction_chainE :
+  fixes \<sigma> :: \<open>nat \<Rightarrow> 'a :: restriction_space option\<close> assumes \<open>chain\<down> \<sigma>\<close>
+  obtains \<open>\<sigma> = (\<lambda>n. None)\<close>
+  | \<sigma>' where \<open>chain\<down> \<sigma>'\<close> and \<open>\<sigma> = (\<lambda>n. if n = 0 then None else Some (\<sigma>' n))\<close>
+proof -
+  from \<open>chain\<down> \<sigma>\<close> consider \<open>\<forall>n. \<sigma> n = None\<close> | \<open>\<forall>n>0. \<sigma> n \<noteq> None\<close>
+    by (metis bot_nat_0.not_eq_extremum linorder_neqE_nat
+        restriction_chain_def_bis restriction_option_eq_None_iff)
+  thus thesis
+  proof cases
+    from that(1) show \<open>\<forall>n. \<sigma> n = None \<Longrightarrow> thesis\<close> by fast
+  next
+    define \<sigma>' where \<open>\<sigma>' n \<equiv> if n = 0 then undefined \<down> 0 else the (\<sigma> n)\<close> for n
+    assume \<open>\<forall>n>0. \<sigma> n \<noteq> None\<close>
+    with \<open>chain\<down> \<sigma>\<close> have \<open>chain\<down> \<sigma>'\<close> \<open>\<sigma> = (\<lambda>n. if n = 0 then undefined \<down> 0 else Some (\<sigma>' n))\<close>
+      by (simp_all add: \<sigma>'_def restriction_chain_def)
+        (metis option.sel restriction_option_eq_Some_iff,
+          metis \<sigma>'_def bot_nat_0.not_eq_extremum option.sel restriction_option_0_is_None)
+    with that(2) show thesis by force
+  qed
+qed
+
+
+lemma non_destructive_Some : \<open>non_destructive Some\<close>
+  by (simp add: non_destructiveI)
+
+lemma restriction_cont_Some : \<open>cont\<down> (Some :: 'a :: restriction_space \<Rightarrow> 'a option)\<close>
+  by (rule restriction_shift_imp_restriction_cont[where k = 0])
+    (simp add: restriction_shiftI)
+
+
+instance option :: (complete_restriction_space) complete_restriction_space
+proof intro_classes
+  show \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close> for \<sigma> :: \<open>nat \<Rightarrow> 'a option\<close>
+  proof (elim option_restriction_chainE)
+    show \<open>\<sigma> = (\<lambda>n. None) \<Longrightarrow> convergent\<down> \<sigma>\<close> by simp
+  next
+    fix \<sigma>' assume \<open>chain\<down> \<sigma>'\<close> and \<sigma>_def : \<open>\<sigma> = (\<lambda>n. if n = 0 then None else Some (\<sigma>' n))\<close>
+    from \<open>chain\<down> \<sigma>'\<close> have \<open>convergent\<down> \<sigma>'\<close> by (simp add: restriction_chain_imp_restriction_convergent)
+    hence \<open>convergent\<down> (\<lambda>n. \<sigma>' (n + 1))\<close> by (unfold restriction_convergent_shift_iff)
+    then obtain \<Sigma>' where \<open>(\<lambda>n. \<sigma>' (n + 1)) \<midarrow>\<down>\<rightarrow> \<Sigma>'\<close> by (blast dest: restriction_convergentD')
+    hence \<open>(\<lambda>n. Some (\<sigma>' (n + 1))) \<midarrow>\<down>\<rightarrow> Some \<Sigma>'\<close> by (fact restriction_contD[OF restriction_cont_Some])
+    hence \<open>convergent\<down> (\<lambda>n. Some (\<sigma>' (n + 1)))\<close> by (blast intro: restriction_convergentI)
+    hence \<open>convergent\<down> (\<lambda>n. \<sigma> (n + 1))\<close> by (simp add: \<sigma>_def)
+    thus \<open>convergent\<down> \<sigma>\<close> using restriction_convergent_shift_iff by blast
+  qed
+qed
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Trace_Model_CSP.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Trace_Model_CSP.thy
@@ -0,0 +1,1056 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Trace Model of CSP\<close>
+
+(*<*)
+theory RS_Trace_Model_CSP
+  imports Restriction_Spaces "HOL-Library.Prefix_Order"
+begin 
+  (*>*)
+
+text \<open>
+In the AFP one can already find verbatim\<open>HOL-CSP\<close>, a shallow embedding of the failure-divergence
+model of denotational semantics proposed by Hoare, Roscoe and Brookes in the eighties.
+Here, we simplify the example by restraining ourselves to a trace model.\<close>
+
+subsection \<open>Prerequisites\<close>
+
+datatype 'a event = ev (of_ev : 'a) | tick (\<open>\<checkmark>\<close>)
+
+type_synonym 'a trace = \<open>'a event list\<close>
+
+
+definition tickFree :: \<open>'a trace \<Rightarrow> bool\<close> (\<open>tF\<close>)
+  where \<open>tickFree t \<equiv> \<checkmark> \<notin> set t\<close>
+
+definition front_tickFree :: \<open>'a trace \<Rightarrow> bool\<close> (\<open>ftF\<close>)
+  where \<open>front_tickFree s \<equiv> s = [] \<or> tickFree (tl (rev s))\<close>
+
+
+lemma tickFree_Nil        [simp] : \<open>tF []\<close>
+  and tickFree_Cons_iff   [simp] : \<open>tF (a # t) \<longleftrightarrow> a \<noteq> \<checkmark> \<and> tF t\<close>
+  and tickFree_append_iff [simp] : \<open>tF (s @ t) \<longleftrightarrow> tF s    \<and> tF t\<close>
+  and tickFree_rev_iff    [simp] : \<open>tF (rev t) \<longleftrightarrow> tF t\<close>
+  and non_tickFree_tick   [simp] : \<open>\<not> tF [\<checkmark>]\<close>
+  by (auto simp add: tickFree_def)
+
+lemma tickFree_iff_is_map_ev : \<open>tF t \<longleftrightarrow> (\<exists>u. t = map ev u)\<close>
+  by (metis event.collapse event.distinct(1) ex_map_conv tickFree_def)
+
+lemma front_tickFree_Nil   [simp] : \<open>ftF []\<close>
+  and front_tickFree_single[simp] : \<open>ftF [a]\<close>
+  by (simp_all add: front_tickFree_def)
+
+
+lemma tickFree_tl : \<open>tF s \<Longrightarrow> tF (tl s)\<close>
+  by (cases s) simp_all
+
+lemma non_tickFree_imp_not_Nil: \<open>\<not> tF s \<Longrightarrow> s \<noteq> []\<close>
+  using tickFree_Nil by blast
+
+lemma tickFree_butlast: \<open>tF s \<longleftrightarrow> tF (butlast s) \<and> (s \<noteq> [] \<longrightarrow> last s \<noteq> \<checkmark>)\<close>
+  by (induct s) simp_all
+
+lemma front_tickFree_iff_tickFree_butlast: \<open>ftF s \<longleftrightarrow> tF (butlast s)\<close>
+  by (induct s) (auto simp add: front_tickFree_def)
+
+lemma front_tickFree_Cons_iff: \<open>ftF (a # s) \<longleftrightarrow> s = [] \<or> a \<noteq> \<checkmark> \<and> ftF s\<close>
+  by (simp add: front_tickFree_iff_tickFree_butlast)
+
+lemma front_tickFree_append_iff:
+  \<open>ftF (s @ t) \<longleftrightarrow> (if t = [] then ftF s else tF s \<and> ftF t)\<close>
+  by (simp add: butlast_append front_tickFree_iff_tickFree_butlast)
+
+lemma tickFree_imp_front_tickFree [simp] : \<open>tF s \<Longrightarrow> ftF s\<close>
+  by (simp add: front_tickFree_def tickFree_tl)
+
+lemma front_tickFree_charn: \<open>ftF s \<longleftrightarrow> s = [] \<or> (\<exists>a t. s = t @ [a] \<and> tF t)\<close>
+  by (cases s rule: rev_cases) (simp_all add: front_tickFree_def)
+
+
+lemma nonTickFree_n_frontTickFree: \<open>\<not> tF s \<Longrightarrow> ftF s \<Longrightarrow> \<exists>t r. s = t @ [\<checkmark>]\<close>
+  by (metis front_tickFree_charn tickFree_Cons_iff
+      tickFree_Nil tickFree_append_iff)
+
+lemma front_tickFree_dw_closed : \<open>ftF (s @ t) \<Longrightarrow> ftF s\<close>
+  by (metis front_tickFree_append_iff tickFree_imp_front_tickFree)
+
+lemma front_tickFree_append: \<open>tF s \<Longrightarrow> ftF t \<Longrightarrow> ftF (s @ t)\<close>
+  by (simp add: front_tickFree_append_iff)
+
+lemma tickFree_imp_front_tickFree_snoc: \<open>tF s \<Longrightarrow> ftF (s @ [a])\<close>
+  by (simp add: front_tickFree_append)
+
+lemma front_tickFree_nonempty_append_imp: \<open>ftF (t @ r) \<Longrightarrow> r \<noteq> [] \<Longrightarrow> tF t \<and> ftF r\<close>
+  by (simp add: front_tickFree_append_iff)
+
+lemma tickFree_map_ev [simp] : \<open>tF (map ev t)\<close>
+  by (induct t) simp_all
+
+lemma tickFree_map_ev_comp [simp] : \<open>tF (map (ev \<circ> f) t)\<close>
+  by (metis list.map_comp tickFree_map_ev)
+
+lemma front_tickFree_map_map_event_iff :
+  \<open>ftF (map (map_event f) t) \<longleftrightarrow> ftF t\<close>
+  by (induct t) (simp_all add: front_tickFree_Cons_iff)
+
+
+
+
+definition is_process :: \<open>'a trace set \<Rightarrow> bool\<close>
+  where \<open>is_process T \<equiv> [] \<in> T \<and> (\<forall>t. t \<in> T \<longrightarrow> ftF t) \<and> (\<forall>t u. t @ u \<in> T \<longrightarrow> t \<in> T)\<close>
+
+typedef 'a process = \<open>{T :: 'a trace set. is_process T}\<close>
+  morphisms Traces to_process
+proof (rule exI)
+  show \<open>{[]} \<in> {T. is_process T}\<close>
+    by (simp add: is_process_def front_tickFree_def)
+qed
+
+setup_lifting type_definition_process
+
+
+notation Traces (\<open>\<T>\<close>)
+
+lemma is_process_inv :
+  \<open>[] \<in> \<T> P \<and> (\<forall>t. t \<in> \<T> P \<longrightarrow> ftF t) \<and> (\<forall>t u. t @ u \<in> \<T> P \<longrightarrow> t \<in> \<T> P)\<close>
+  by (metis is_process_def mem_Collect_eq to_process_cases to_process_inverse)
+
+lemma Nil_elem_T       : \<open>[] \<in> \<T> P\<close>
+  and front_tickFree_T : \<open>t \<in> \<T> P \<Longrightarrow> ftF t\<close>
+  and T_dw_closed      : \<open>t @ u \<in> \<T> P \<Longrightarrow> t \<in> \<T> P\<close>
+  by (simp_all add: is_process_inv)
+
+
+lemma process_eq_spec : \<open>P = Q \<longleftrightarrow> \<T> P = \<T> Q\<close>
+  by transfer simp
+
+
+
+subsection \<open>First Processes\<close>
+
+lift_definition BOT :: \<open>'a process\<close> is \<open>{t. ftF t}\<close>
+  by (auto simp add: is_process_def front_tickFree_append_iff)
+
+lemma T_BOT : \<open>\<T> BOT = {t. ftF t}\<close>
+  by (simp add: BOT.rep_eq)
+
+
+lift_definition SKIP :: \<open>'a process\<close> is \<open>{[], [\<checkmark>]}\<close>
+  by (simp add: is_process_def append_eq_Cons_conv)
+
+lemma T_SKIP : \<open>\<T> SKIP = {[], [\<checkmark>]}\<close>
+  by (simp add: SKIP.rep_eq)
+
+
+lift_definition STOP :: \<open>'a process\<close> is \<open>{[]}\<close>
+  by (simp add: is_process_def)
+
+lemma T_STOP : \<open>\<T> STOP = {[]}\<close>
+  by (simp add: STOP.rep_eq)
+
+
+lift_definition Sup_processes ::
+  \<open>(nat \<Rightarrow> 'a process) \<Rightarrow> 'a process\<close> is \<open>\<lambda>\<sigma>. \<Inter>i. \<T> (\<sigma> i)\<close>
+proof -
+  show \<open>is_process (\<Inter>i. \<T> (\<sigma> i))\<close> for \<sigma> :: \<open>nat \<Rightarrow> 'a process\<close>
+  proof (unfold is_process_def, intro conjI allI impI)
+    show \<open>[] \<in> (\<Inter>i. \<T> (\<sigma> i))\<close> by (simp add: Nil_elem_T)
+  next
+    show \<open>t \<in> (\<Inter>i. \<T> (\<sigma> i)) \<Longrightarrow> ftF t\<close> for t
+      by (auto intro: front_tickFree_T)
+  next
+    show \<open>t @ u \<in> (\<Inter>i. \<T> (\<sigma> i)) \<Longrightarrow> t \<in> (\<Inter>i. \<T> (\<sigma> i))\<close> for t u
+      by (auto intro: T_dw_closed)
+  qed
+qed
+
+lemma T_Sup_processes : \<open>\<T> (Sup_processes \<sigma>) = (\<Inter>i. \<T> (\<sigma> i))\<close>
+  by (simp add: Sup_processes.rep_eq)
+
+
+
+subsection \<open>Instantiations\<close>
+
+instantiation process :: (type) order
+begin
+
+definition less_eq_process :: \<open>'a process \<Rightarrow> 'a process \<Rightarrow> bool\<close>
+  where \<open>P \<le> Q \<equiv> \<T> Q \<subseteq> \<T> P\<close>
+
+definition less_process :: \<open>'a process \<Rightarrow> 'a process \<Rightarrow> bool\<close>
+  where \<open>P < Q \<equiv> \<T> Q \<subset> \<T> P\<close>
+
+instance
+  by intro_classes
+    (auto simp add: less_eq_process_def less_process_def process_eq_spec)
+
+end
+
+
+instantiation process :: (type) order_restriction_space
+begin
+
+lift_definition restriction_process :: \<open>'a process \<Rightarrow> nat \<Rightarrow> 'a process\<close>
+  is \<open>\<lambda>P n. \<T> P \<union> {t @ u |t u. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u}\<close>
+proof -
+  show \<open>?thesis P n\<close> (is \<open>is_process ?t\<close>) for P n
+  proof (unfold is_process_def, intro conjI allI impI)
+    show \<open>[] \<in> ?t\<close> by (simp add: Nil_elem_T)
+  next
+    show \<open>t \<in> ?t \<Longrightarrow> ftF t\<close> for t
+      by (auto simp add: front_tickFree_append_iff front_tickFree_T)
+  next
+    fix t u assume \<open>t @ u \<in> ?t\<close>
+    then consider \<open>t @ u \<in> \<T> P\<close>
+      | (@) t' u' where \<open>t @ u = t' @ u'\<close> \<open>t' \<in> \<T> P\<close>
+        \<open>length t' = n\<close> \<open>tF t'\<close> \<open>ftF u'\<close> by blast
+    thus \<open>t \<in> ?t\<close>
+    proof cases
+      from T_dw_closed show \<open>t @ u \<in> \<T> P \<Longrightarrow> t \<in> ?t\<close> by blast
+    next
+      case @
+      show \<open>t \<in> ?t\<close>
+      proof (cases \<open>length t \<le> length t'\<close>)
+        assume \<open>length t \<le> length t'\<close>
+        with "@"(1) obtain t'' where \<open>t' = t @ t''\<close>
+          by (metis append_eq_append_conv_if append_eq_conv_conj)
+        with "@"(2) T_dw_closed show \<open>t \<in> ?t\<close> by blast
+      next
+        assume \<open>\<not> length t \<le> length t'\<close>
+        hence \<open>length t' \<le> length t\<close> by simp
+        with "@"(1,4,5) \<open>\<not> length t \<le> length t'\<close>
+        obtain t'' where \<open>t = t' @ t''\<close> \<open>ftF t''\<close>
+          by (metis append_eq_conv_conj drop_eq_Nil front_tickFree_append
+              front_tickFree_dw_closed front_tickFree_nonempty_append_imp
+              take_all_iff take_append)
+        with "@"(2,3,4) show \<open>t \<in> ?t\<close> by blast
+      qed
+    qed
+  qed
+qed
+
+lemma T_restriction_process :
+  \<open>\<T> (P \<down> n) = \<T> P \<union> {t @ u |t u. t \<in> \<T> P \<and> length t = n \<and> tF t \<and> ftF u}\<close>
+  by (simp add: restriction_process.rep_eq)
+
+
+lemma restriction_process_0 [simp] : \<open>P \<down> 0 = BOT\<close>
+  by transfer (auto simp add: front_tickFree_T Nil_elem_T)
+
+lemma T_restriction_processE :
+  \<open>t \<in> \<T> (P \<down> n) \<Longrightarrow>
+   (t \<in> \<T> P \<Longrightarrow> length t \<le> n \<Longrightarrow> thesis) \<Longrightarrow>
+   (\<And>u v. t = u @ v \<Longrightarrow> u \<in> \<T> P \<Longrightarrow> length u = n \<Longrightarrow> tF u \<Longrightarrow> ftF v \<Longrightarrow> thesis) \<Longrightarrow>
+   thesis\<close>
+  by (simp add: T_restriction_process)
+    (metis (no_types) T_dw_closed append_take_drop_id drop_eq_Nil
+      front_tickFree_T front_tickFree_nonempty_append_imp length_take min_def)
+
+
+instance
+proof intro_classes
+  show \<open>P \<down> 0 \<le> Q \<down> 0\<close> for P Q :: \<open>'a process\<close> by simp
+next
+  show \<open>P \<down> n \<down> m = P \<down> min n m\<close> for P :: \<open>'a process\<close> and n m
+  proof (subst process_eq_spec, intro subset_antisym subsetI)
+    show \<open>t \<in> \<T> (P \<down> n \<down> m) \<Longrightarrow> t \<in> \<T> (P \<down> min n m)\<close> for t
+      by (elim T_restriction_processE)
+        (auto simp add: T_restriction_process intro: front_tickFree_append)
+  next
+    show \<open>t \<in> \<T> (P \<down> min n m) \<Longrightarrow> t \<in> \<T> (P \<down> n \<down> m)\<close> for t
+      by (elim T_restriction_processE)
+        (auto simp add: T_restriction_process min_def split: if_split_asm)
+  qed
+next
+  show \<open>P \<le> Q \<Longrightarrow> P \<down> n \<le> Q \<down> n\<close> for P Q :: \<open>'a process\<close> and n
+    by (auto simp add: less_eq_process_def T_restriction_process)
+next
+  fix P Q :: \<open>'a process\<close> assume \<open>\<not> P \<le> Q\<close>
+  then obtain t where \<open>t \<in> \<T> Q\<close> \<open>t \<notin> \<T> P\<close>
+    unfolding less_eq_process_def by blast
+  hence \<open>t \<in> \<T> (Q \<down> Suc (length t)) \<and> t \<notin> \<T> (P \<down> Suc (length t))\<close>
+    by (auto simp add: T_restriction_process) 
+  hence \<open>\<not> P \<down> Suc (length t) \<le> Q \<down> Suc (length t)\<close>
+    unfolding less_eq_process_def by blast
+  thus \<open>\<exists>n. \<not> P \<down> n \<le> Q \<down> n\<close> ..
+qed
+
+
+text \<open>Of course, we recover the structure of class\<open>restriction_space\<close>.\<close>
+
+lemma \<open>OFCLASS ('a process, restriction_space_class)\<close>
+  by intro_classes
+
+end
+
+
+lemma restricted_Sup_processes_is :
+  \<open>(\<lambda>n. Sup_processes \<sigma> \<down> n) = \<sigma>\<close> if \<open>restriction_chain \<sigma>\<close>
+proof (subst (2) restricted_restriction_chain_is
+    [OF \<open>restriction_chain \<sigma>\<close>, symmetric], rule ext)
+  fix n
+  have \<open>length t \<le> n \<Longrightarrow> t \<in> \<T> (\<sigma> n) \<longleftrightarrow> (\<forall>i. t \<in> \<T> (\<sigma> i))\<close> for t n
+  proof safe
+    show \<open>t \<in> \<T> (\<sigma> i)\<close> if \<open>length t \<le> n\<close> \<open>t \<in> \<T> (\<sigma> n)\<close> for i
+    proof (cases \<open>i \<le> n\<close>)
+      from restriction_chain_def_ter[THEN iffD1, OF \<open>restriction_chain \<sigma>\<close>]
+      show \<open>i \<le> n \<Longrightarrow> t \<in> \<T> (\<sigma> i)\<close>
+        by (metis (lifting) \<open>t \<in> \<T> (\<sigma> n)\<close> T_restriction_process Un_iff)
+    next
+      from \<open>length t \<le> n\<close> \<open>t \<in> \<T> (\<sigma> n)\<close> show \<open>\<not> i \<le> n \<Longrightarrow> t \<in> \<T> (\<sigma> i)\<close>
+        by (induct n, simp_all add: Nil_elem_T)
+          (metis (no_types) T_restriction_processE
+            append_eq_conv_conj linorder_linear take_all_iff
+            restriction_chain_def_ter[THEN iffD1, OF \<open>restriction_chain \<sigma>\<close>])
+    qed
+  next
+    show \<open>\<forall>i. t \<in> \<T> (\<sigma> i) \<Longrightarrow> t \<in> \<T> (\<sigma> n)\<close> by simp
+  qed
+  hence * : \<open>length t \<le> n \<Longrightarrow> t \<in> \<T> (\<sigma> n) \<longleftrightarrow> t \<in> \<T> (Sup_processes \<sigma>)\<close> for t n
+    by (simp add: T_Sup_processes)
+
+  show \<open>Sup_processes \<sigma> \<down> n = \<sigma> n \<down> n\<close> for n
+  proof (subst process_eq_spec, intro subset_antisym subsetI)
+    show \<open>t \<in> \<T> (Sup_processes \<sigma> \<down> n) \<Longrightarrow> t \<in> \<T> (\<sigma> n \<down> n)\<close> for t
+    proof (elim T_restriction_processE)
+      show \<open>t \<in> \<T> (Sup_processes \<sigma>) \<Longrightarrow> length t \<le> n \<Longrightarrow> t \<in> \<T> (\<sigma> n \<down> n)\<close>
+        by (simp add: "*" T_restriction_process)
+    next
+      show \<open>\<lbrakk>t = u @ v; u \<in> \<T> (Sup_processes \<sigma>); length u = n; tF u; ftF v\<rbrakk>
+            \<Longrightarrow> t \<in> \<T> (\<sigma> n \<down> n)\<close> for u v
+        by (auto simp add: "*" T_restriction_process)
+    qed
+  next
+    from "*" show \<open>t \<in> \<T> (\<sigma> n \<down> n) \<Longrightarrow> t \<in> \<T> (Sup_processes \<sigma> \<down> n)\<close> for t
+      by (elim T_restriction_processE)
+        (auto simp add: T_restriction_process)
+  qed
+qed
+
+
+instance process :: (type) complete_restriction_space 
+proof intro_classes
+  show \<open>restriction_convergent \<sigma>\<close> if \<open>restriction_chain \<sigma>\<close>
+    for \<sigma> :: \<open>nat \<Rightarrow> 'a process\<close>
+  proof (rule restriction_convergentI)
+    have \<open>\<sigma> = (\<lambda>n. Sup_processes \<sigma> \<down> n)\<close>
+      by (simp add: restricted_Sup_processes_is \<open>restriction_chain \<sigma>\<close>)
+    moreover have \<open>(\<lambda>n. Sup_processes \<sigma> \<down> n) \<midarrow>\<down>\<rightarrow> Sup_processes \<sigma>\<close>
+      by (fact restriction_tendsto_restrictions)
+    ultimately show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> Sup_processes \<sigma>\<close> by simp
+  qed
+qed
+
+
+
+subsection \<open>Operators\<close>
+
+lift_definition Choice :: \<open>'a process \<Rightarrow> 'a process \<Rightarrow> 'a process\<close> (infixl \<open>\<box>\<close> 82)
+  is \<open>\<lambda>P Q. \<T> P \<union> \<T> Q\<close>
+  by (auto simp add: is_process_def Nil_elem_T front_tickFree_T intro: T_dw_closed)
+
+lemma T_Choice : \<open>\<T> (P \<box> Q) = \<T> P \<union> \<T> Q\<close>
+  by (simp add: Choice.rep_eq)
+
+
+lift_definition GlobalChoice :: \<open>['b set, 'b \<Rightarrow> 'a process] \<Rightarrow> 'a process\<close>
+  is \<open>\<lambda>A P. if A = {} then {[]} else \<Union>a\<in>A. \<T> (P a)\<close>
+  by (auto simp add: is_process_def Nil_elem_T front_tickFree_T intro: T_dw_closed)
+
+syntax "_GlobalChoice" :: \<open>[pttrn, 'b set, 'a process] \<Rightarrow> 'a process\<close>
+  (\<open>(3\<box>((_)/\<in>(_))./ (_))\<close> [78,78,77] 77)
+syntax_consts "_GlobalChoice" \<rightleftharpoons> GlobalChoice
+translations  "\<box> a \<in> A. P" \<rightleftharpoons> "CONST GlobalChoice A (\<lambda>a. P)"
+
+lemma T_GlobalChoice : \<open>\<T> (\<box>a \<in> A. P a) = (if A = {} then {[]} else \<Union>a\<in>A. \<T> (P a))\<close>
+  by (simp add: GlobalChoice.rep_eq)
+
+
+lift_definition Seq :: \<open>'a process \<Rightarrow> 'a process \<Rightarrow> 'a process\<close> (infixl \<open>;\<close> 74)
+  is \<open>\<lambda>P Q. {t \<in> \<T> P. tF t} \<union> {t @ u |t u. t @ [\<checkmark>] \<in> \<T> P \<and> u \<in> \<T> Q}\<close>
+  by (auto simp add: is_process_def Nil_elem_T append_eq_append_conv2 intro: T_dw_closed)
+    (metis front_tickFree_T front_tickFree_append_iff
+      front_tickFree_dw_closed not_Cons_self,
+      meson front_tickFree_append_iff is_process_inv snoc_eq_iff_butlast)
+
+lemma T_Seq : \<open>\<T> (P ; Q) = {t \<in> \<T> P. tF t} \<union> {t @ u |t u. t @ [\<checkmark>] \<in> \<T> P \<and> u \<in> \<T> Q}\<close>
+  by (simp add: Seq.rep_eq)
+
+
+lift_definition Renaming :: \<open>['a process, 'a \<Rightarrow> 'b] \<Rightarrow> 'b process\<close>
+  is \<open>\<lambda>P f. {map (map_event f) u |u. u \<in> \<T> P}\<close>
+  by (auto simp add: is_process_def Nil_elem_T front_tickFree_map_map_event_iff
+      front_tickFree_T append_eq_map_conv intro: T_dw_closed)
+
+lemma T_Renaming : \<open>\<T> (Renaming P f) = {map (map_event f) u |u. u \<in> \<T> P}\<close>
+  by (simp add: Renaming.rep_eq)
+
+
+lift_definition Mprefix :: \<open>['a set, 'a \<Rightarrow> 'a process] \<Rightarrow> 'a process\<close>
+  is \<open>\<lambda>A P. insert [] {ev a # t |a t. a \<in> A \<and> t \<in> \<T> (P a)}\<close>
+  by (auto simp add: is_process_def front_tickFree_Cons_iff
+      front_tickFree_T append_eq_Cons_conv intro: T_dw_closed)
+
+syntax "_Mprefix" :: \<open>[pttrn, 'a set, 'a process] \<Rightarrow> 'a process\<close>	
+  (\<open>(3\<box>((_)/\<in>(_))/ \<rightarrow> (_))\<close> [78,78,77] 77)
+syntax_consts "_Mprefix" \<rightleftharpoons> Mprefix
+translations "\<box>a\<in>A \<rightarrow> P" \<rightleftharpoons> "CONST Mprefix A (\<lambda>a. P)"
+
+lemma T_Mprefix : \<open>\<T> (\<box>a \<in> A \<rightarrow> P a) = insert [] {ev a # t |a t. a \<in> A \<and> t \<in> \<T> (P a)}\<close>
+  by (simp add: Mprefix.rep_eq)
+
+
+
+
+
+
+fun setinterleaving :: \<open>'a trace \<times> 'a set \<times> 'a trace \<Rightarrow> 'a trace set\<close>
+  where Nil_setinterleaving_Nil : \<open>setinterleaving ([], A, []) = {[]}\<close>
+
+|     ev_setinterleaving_Nil :
+  \<open>setinterleaving (ev a # u, A, []) =
+         (if a \<in> A then {} else {ev a # t| t. t \<in> setinterleaving (u, A, [])})\<close>
+|     tick_setinterleaving_Nil : \<open>setinterleaving (\<checkmark> # u, A, []) = {}\<close>
+
+|     Nil_setinterleaving_ev :
+  \<open>setinterleaving ([], A, ev b # v) =
+         (if b \<in> A then {} else {ev b # t| t. t \<in> setinterleaving ([], A, v)})\<close>
+|     Nil_setinterleaving_tick  : \<open>setinterleaving ([], A, \<checkmark> # v) = {}\<close>
+
+|     ev_setinterleaving_ev : 
+  \<open>setinterleaving (ev a # u, A, ev b # v) =
+         (  if a \<in> A
+          then   if b \<in> A 
+                then   if a = b
+                     then {ev a # t |t. t \<in> setinterleaving (u, A, v)}
+                     else {}
+                else {ev b # t |t. t \<in> setinterleaving (ev a # u, A, v)}
+           else   if b \<in> A then {ev a # t |t. t \<in> setinterleaving (u, A, ev b # v)}
+                else {ev a # t |t. t \<in> setinterleaving (u, A, ev b # v)} \<union>
+                     {ev b # t |t. t \<in> setinterleaving (ev a # u, A, v)})\<close>
+|     ev_setinterleaving_tick :
+  \<open>setinterleaving (ev a # u, A, \<checkmark> # v) =
+         (if a \<in> A then {} else {ev a # t |t. t \<in> setinterleaving (u, A, \<checkmark> # v)})\<close>
+|     tick_setinterleaving_ev :
+  \<open>setinterleaving (\<checkmark> # u, A, ev b # v) =
+         (if b \<in> A then {} else {ev b # t |t. t \<in> setinterleaving (\<checkmark> # u, A, v)})\<close>
+|     tick_setinterleaving_tick :
+  \<open>setinterleaving (\<checkmark> # u, A, \<checkmark> # v) = {\<checkmark> # t |t. t \<in> setinterleaving (u, A, v)}\<close>
+
+
+lemmas setinterleaving_induct
+  [case_names Nil_setinterleaving_Nil ev_setinterleaving_Nil tick_setinterleaving_Nil
+    Nil_setinterleaving_ev Nil_setinterleaving_tick ev_setinterleaving_ev
+    ev_setinterleaving_tick tick_setinterleaving_ev tick_setinterleaving_tick] =
+  setinterleaving.induct
+
+
+
+lemma Cons_setinterleaving_Nil :
+  \<open>setinterleaving (e # u, A, []) =
+   (case e of ev a \<Rightarrow> (  if a \<in> A then {}
+                       else {ev a # t |t. t \<in> setinterleaving (u, A, [])})
+            | \<checkmark> \<Rightarrow> {})\<close>
+  by (cases e) simp_all
+
+lemma Nil_setinterleaving_Cons :
+  \<open>setinterleaving ([], A, e # v) =
+   (case e of ev a \<Rightarrow> (  if a \<in> A then {}
+                       else {ev a # t |t. t \<in> setinterleaving ([], A, v)})
+            | \<checkmark> \<Rightarrow> {})\<close>
+  by (cases e) simp_all
+
+lemma Cons_setinterleaving_Cons :
+  \<open>setinterleaving (e # u, A, f # v) =
+   (case e of ev a \<Rightarrow>
+    (case f of ev b \<Rightarrow> 
+       if a \<in> A
+     then   if b \<in> A 
+           then   if a = b
+                then {ev a # t |t. t \<in> setinterleaving (u, A, v)}
+                else {}
+           else {ev b # t |t. t \<in> setinterleaving (ev a # u, A, v)}
+      else   if b \<in> A then {ev a # t |t. t \<in> setinterleaving (u, A, ev b # v)}
+           else {ev a # t |t. t \<in> setinterleaving (u, A, ev b # v)} \<union>
+                {ev b # t |t. t \<in> setinterleaving (ev a # u, A, v)}
+             | \<checkmark> \<Rightarrow>   if a \<in> A then {}
+                    else {ev a # t |t. t \<in> setinterleaving (u, A, \<checkmark> # v)})
+            | \<checkmark> \<Rightarrow>
+    (case f of ev b \<Rightarrow>   if b \<in> A then {}
+                       else {ev b # t| t. t \<in> setinterleaving (\<checkmark> # u, A, v)}
+             | \<checkmark> \<Rightarrow> {\<checkmark> # t |t. t \<in> setinterleaving (u, A, v)}))\<close>
+  by (cases e; cases f) simp_all
+
+
+lemmas setinterleaving_simps =
+  Cons_setinterleaving_Nil Nil_setinterleaving_Cons Cons_setinterleaving_Cons 
+
+
+
+abbreviation setinterleaves ::
+  \<open>['a trace, 'a trace, 'a trace, 'a set] \<Rightarrow> bool\<close>
+  (\<open>(_ /(setinterleaves)/ '(()'(_, _')(), _'))\<close> [63,0,0,0] 64)
+  where \<open>t setinterleaves ((u, v), A) \<equiv> t \<in> setinterleaving (u, A, v)\<close>
+
+
+
+
+lemma tickFree_setinterleaves_iff :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> tF t \<longleftrightarrow> tF u \<and> tF v\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+lemma setinterleaves_tickFree_imp :
+  \<open>tF u \<or> tF v \<Longrightarrow> t setinterleaves ((u, v), A) \<Longrightarrow> tF t \<and> tF u \<and> tF v\<close>
+  by (elim disjE; induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+
+
+lemma setinterleaves_NilL_iff :
+  \<open>t setinterleaves (([], v), A) \<longleftrightarrow>
+   tF v \<and> set v \<inter> ev ` A = {} \<and> t = map ev (map of_ev v)\<close>
+  by (induct \<open>([] :: 'a trace, A, v)\<close> arbitrary: t v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+lemma setinterleaves_NilR_iff :
+  \<open>t setinterleaves ((u, []), A) \<longleftrightarrow>
+   tF u \<and> set u \<inter> ev ` A = {} \<and> t = map ev (map of_ev u)\<close>
+  by (induct \<open>(u, A, [] :: 'a trace)\<close>
+      arbitrary: t u rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+lemma Nil_setinterleaves :
+  \<open>[] setinterleaves ((u, v), A) \<Longrightarrow> u = [] \<and> v = []\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: u v rule: setinterleaving_induct)
+    (simp_all split: if_split_asm)
+
+
+lemma front_tickFree_setinterleaves_iff :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> ftF t \<longleftrightarrow> ftF u \<and> ftF v\<close>
+proof (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+  case (tick_setinterleaving_tick u v) thus ?case
+    by (simp add: split: if_split_asm)
+      (metis Nil_setinterleaves Nil_setinterleaving_Nil front_tickFree_Cons_iff singletonD)
+qed (simp add: setinterleaves_NilL_iff setinterleaves_NilR_iff split: if_split_asm;
+    metis event.simps(3) front_tickFree_Cons_iff front_tickFree_Nil)+
+
+
+lemma setinterleaves_snoc_notinL :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> a \<notin> A \<Longrightarrow>
+   t @ [ev a] setinterleaves ((u @ [ev a], v), A)\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+lemma setinterleaves_snoc_notinR :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> a \<notin> A \<Longrightarrow>
+   t @ [ev a] setinterleaves ((u, v @ [ev a]), A)\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+lemma setinterleaves_snoc_inside :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> a \<in> A \<Longrightarrow>
+   t @ [ev a] setinterleaves ((u @ [ev a], v @ [ev a]), A)\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+
+lemma setinterleaves_snoc_tick :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> t @ [\<checkmark>] setinterleaves ((u @ [\<checkmark>], v @ [\<checkmark>]), A)\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+
+lemma Cons_tick_setinterleavesE :
+  \<open>\<checkmark> # t setinterleaves ((u, v), A) \<Longrightarrow>
+   (\<And>u' v' r s. \<lbrakk>u = \<checkmark> # u'; v = \<checkmark> # v'; t setinterleaves ((u', v'), A)\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (simp_all split: if_split_asm)
+
+lemma Cons_ev_setinterleavesE :
+  \<open>ev a # t setinterleaves ((u, v), A) \<Longrightarrow>
+   (\<And>u'. a \<notin> A \<Longrightarrow> u = ev a # u' \<Longrightarrow> t setinterleaves ((u', v), A) \<Longrightarrow> thesis) \<Longrightarrow>
+   (\<And>v'. a \<notin> A \<Longrightarrow> v = ev a # v' \<Longrightarrow> t setinterleaves ((u, v'), A) \<Longrightarrow> thesis) \<Longrightarrow>
+   (\<And>u' v'. a \<in> A \<Longrightarrow> u = ev a # u' \<Longrightarrow> v = ev a # v' \<Longrightarrow>
+             t setinterleaves ((u', v'), A) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+proof (induct \<open>(u, A, v)\<close> arbitrary: u v t rule: setinterleaving_induct)
+  case Nil_setinterleaving_Nil thus ?case by simp
+next
+  case (ev_setinterleaving_Nil b u)
+  from ev_setinterleaving_Nil.prems(1) show ?case
+    by (simp add: ev_setinterleaving_Nil.prems(2) split: if_split_asm)
+next
+  case (tick_setinterleaving_Nil r u) thus ?case by simp
+next
+  case (Nil_setinterleaving_ev c v)
+  from Nil_setinterleaving_ev.prems(1) show ?case
+    by (simp add: Nil_setinterleaving_ev.prems(3) split: if_split_asm)
+next
+  case (Nil_setinterleaving_tick s v) thus ?case by simp
+next
+  case (ev_setinterleaving_ev b u c v)
+  from ev_setinterleaving_ev.prems(1) show ?case
+    by (simp add: ev_setinterleaving_ev.prems(2, 3, 4) split: if_split_asm)
+      (use ev_setinterleaving_ev.prems(2, 3) in presburger)
+next
+  case (ev_setinterleaving_tick b u s v)
+  from ev_setinterleaving_tick.prems(1) show ?case
+    by (simp add: ev_setinterleaving_tick.prems(2) split: if_split_asm)
+next
+  case (tick_setinterleaving_ev r u c v)
+  from tick_setinterleaving_ev.prems(1) show ?case
+    by (simp add: tick_setinterleaving_ev.prems(3) split: if_split_asm)
+next
+  case (tick_setinterleaving_tick u v) thus ?case by simp
+qed
+
+
+lemma rev_setinterleaves_rev_rev_iff :
+  \<open>rev t setinterleaves ((rev u, rev v), A)
+   \<longleftrightarrow> t setinterleaves ((u, v), A)\<close>
+proof (rule iffI)
+  show \<open>t setinterleaves ((u, v), A) \<Longrightarrow>
+        rev t setinterleaves ((rev u, rev v), A)\<close> for t u v
+    by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+      (auto simp add: setinterleaves_snoc_notinL setinterleaves_snoc_notinR
+        setinterleaves_snoc_inside setinterleaves_snoc_tick split: if_split_asm)
+  from this[of \<open>rev t\<close> \<open>rev u\<close> \<open>rev v\<close>, simplified]
+  show \<open>rev t setinterleaves ((rev u, rev v), A) \<Longrightarrow>
+        t setinterleaves ((u, v), A)\<close> .
+qed
+
+
+lemma setinterleaves_preserves_ev_notin_set :
+  \<open>\<lbrakk>ev a \<notin> set u; ev a \<notin> set v; t setinterleaves ((u, v), A)\<rbrakk> \<Longrightarrow> ev a \<notin> set t\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+    (auto split: if_split_asm)
+
+
+
+lemma setinterleaves_preserves_ev_inside_set :
+  \<open>\<lbrakk>ev a \<in> set u; ev a \<in> set v; t setinterleaves ((u, v), A)\<rbrakk> \<Longrightarrow> ev a \<in> set t\<close>
+proof (induct \<open>(u, A, v)\<close> arbitrary: t u v rule: setinterleaving_induct)
+  case Nil_setinterleaving_Nil
+  then show ?case by simp
+next
+  case (ev_setinterleaving_Nil a u)
+  then show ?case by simp
+next
+  case (tick_setinterleaving_Nil u)
+  then show ?case by simp
+next
+  case (Nil_setinterleaving_ev b v)
+  then show ?case by simp
+next
+  case (Nil_setinterleaving_tick v)
+  then show ?case by simp
+next
+  case (ev_setinterleaving_ev a u b v)
+  from ev_setinterleaving_ev.prems show ?case
+    by (simp_all split: if_split_asm)
+      (insert ev_setinterleaving_ev.hyps; metis list.set_intros(1,2))+
+next
+  case (ev_setinterleaving_tick a u v)
+  then show ?case by (auto split: if_split_asm)
+next
+  case (tick_setinterleaving_ev u b v)
+  then show ?case by (auto split: if_split_asm)
+next
+  case (tick_setinterleaving_tick u v)
+  then show ?case by auto
+qed
+
+
+lemma ev_notin_both_sets_imp_empty_setinterleaving :
+  \<open>\<lbrakk>ev a \<in> set u \<and> ev a \<notin> set v \<or> ev a \<notin> set u \<and> ev a \<in> set v; a \<in> A\<rbrakk> \<Longrightarrow>
+   setinterleaving (u, A, v) = {}\<close>
+  by (induct \<open>(u, A, v)\<close> arbitrary: u v rule: setinterleaving_induct)
+    (simp_all, safe, auto)
+
+
+lemma append_setinterleaves_imp :
+  \<open>t setinterleaves ((u, v), A) \<Longrightarrow> t' \<le> t \<Longrightarrow>
+   \<exists>u' \<le> u. \<exists>v' \<le> v. t' setinterleaves ((u', v'), A)\<close>
+proof (induct \<open>(u, A, v)\<close> arbitrary: t u v t' rule: setinterleaving_induct)
+  case Nil_setinterleaving_Nil thus ?case by auto
+next
+  case (ev_setinterleaving_Nil a u)
+  from ev_setinterleaving_Nil.prems
+  obtain w w' where \<open>a \<notin> A\<close> \<open>t = ev a # w\<close> \<open>t' = [] \<or> t' = ev a # w'\<close>
+    \<open>w' \<le> w\<close> \<open>w setinterleaves ((u, []), A)\<close>
+    by (simp split: if_split_asm) (metis (no_types) Prefix_Order.prefix_Cons Nil_prefix)
+  from \<open>t' = [] \<or> t' = ev a # w'\<close> show ?case
+  proof (elim disjE)
+    show \<open>t' = [] \<Longrightarrow> ?case\<close> by auto
+  next
+    assume \<open>t' = ev a # w'\<close>
+    from ev_setinterleaving_Nil.hyps[OF \<open>a \<notin> A\<close> \<open>w setinterleaves ((u, []), A)\<close> \<open>w' \<le> w\<close>]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> [] \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>ev a # u' \<le> ev a # u \<and> v' \<le> [] \<and> t' setinterleaves ((ev a # u', v'), A)\<close>
+      by (auto simp add: \<open>a \<notin> A\<close> \<open>t' = ev a # w'\<close>)
+    thus ?case by blast
+  qed
+next
+  case (tick_setinterleaving_Nil r u) thus ?case by simp
+next
+  case (Nil_setinterleaving_ev b v)
+  from Nil_setinterleaving_ev.prems
+  obtain w w' where \<open>b \<notin> A\<close> \<open>t = ev b # w\<close> \<open>t' = [] \<or> t' = ev b # w'\<close>
+    \<open>w' \<le> w\<close> \<open>w setinterleaves (([], v), A)\<close>
+    by (simp split: if_split_asm) (metis (no_types) Prefix_Order.prefix_Cons Nil_prefix)
+  from \<open>t' = [] \<or> t' = ev b # w'\<close> show ?case
+  proof (elim disjE)
+    show \<open>t' = [] \<Longrightarrow> ?case\<close> by auto
+  next
+    assume \<open>t' = ev b # w'\<close>
+    from Nil_setinterleaving_ev.hyps[OF \<open>b \<notin> A\<close> \<open>w setinterleaves (([], v), A)\<close> \<open>w' \<le> w\<close>]
+    obtain u' v' where \<open>u' \<le> [] \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>u' \<le> [] \<and> ev b # v' \<le> ev b # v \<and> t' setinterleaves ((u', ev b # v'), A)\<close>
+      by (auto simp add: \<open>b \<notin> A\<close> \<open>t' = ev b # w'\<close>)
+    thus ?case by blast
+  qed
+next
+  case (Nil_setinterleaving_tick s v) thus ?case by simp
+next
+  case (ev_setinterleaving_ev a u b v)
+  from ev_setinterleaving_ev.prems
+  consider \<open>t' = []\<close>
+    |      (both_in)   w w' where \<open>a \<in> A\<close> \<open>b \<in> A\<close> \<open>a = b\<close> \<open>t = ev a # w\<close> \<open>t' = ev a # w'\<close>
+      \<open>w setinterleaves ((u, v), A)\<close> \<open>w' \<le> w\<close>
+    |      (inR_mvL)   w w' where \<open>a \<notin> A\<close> \<open>b \<in> A\<close> \<open>t = ev a # w\<close> \<open>t' = ev a # w'\<close>
+      \<open>w setinterleaves ((u, ev b # v), A)\<close> \<open>w' \<le> w\<close>
+    |      (inL_mvR)   w w' where \<open>a \<in> A\<close> \<open>b \<notin> A\<close> \<open>t = ev b # w\<close> \<open>t' = ev b # w'\<close>
+      \<open>w setinterleaves ((ev a # u, v), A)\<close> \<open>w' \<le> w\<close>
+    |      (notin_mvL) w w' where \<open>a \<notin> A\<close> \<open>b \<notin> A\<close> \<open>t = ev a # w\<close> \<open>t' = ev a # w'\<close>
+      \<open>w setinterleaves ((u, ev b # v), A)\<close> \<open>w' \<le> w\<close>
+    |      (notin_mvR) w w' where \<open>a \<notin> A\<close> \<open>b \<notin> A\<close> \<open>t = ev b # w\<close> \<open>t' = ev b # w'\<close>
+      \<open>w setinterleaves ((ev a # u, v), A)\<close> \<open>w' \<le> w\<close>
+    by (cases t') (auto split: if_split_asm)
+  thus ?case
+  proof cases
+    from Nil_setinterleaving_Nil show \<open>t' = [] \<Longrightarrow> ?thesis\<close> by blast
+  next
+    case both_in
+    from ev_setinterleaving_ev(1)[OF both_in(1-3, 6-7)]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>ev a # u' \<le> ev a # u \<and> ev b # v' \<le> ev b # v \<and> t' setinterleaves ((ev a # u', ev b # v'), A)\<close>
+      by (auto simp add: both_in(2, 3) \<open>t' = ev a # w'\<close>)
+    thus ?thesis by blast
+  next
+    case inR_mvL
+    from ev_setinterleaving_ev(3)[OF inR_mvL(1, 2, 5, 6)]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> ev b # v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>ev a # u' \<le> ev a # u \<and> v' \<le> ev b # v \<and> t' setinterleaves ((ev a # u', v'), A)\<close>
+      by (cases v') (auto simp add: inR_mvL(1, 4))
+    thus ?thesis by blast
+  next
+    case inL_mvR
+    from ev_setinterleaving_ev(2)[OF inL_mvR(1, 2, 5, 6)]
+    obtain u' v' where \<open>u' \<le> ev a # u \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>u' \<le> ev a # u \<and> ev b # v' \<le> ev b # v \<and> t' setinterleaves ((u', ev b # v'), A)\<close>
+      by (cases u') (auto simp add: inL_mvR(2, 4))
+    thus ?thesis by blast
+  next
+    case notin_mvL
+    from ev_setinterleaving_ev(4)[OF notin_mvL(1, 2, 5, 6)]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> ev b # v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>ev a # u' \<le> ev a # u \<and> v' \<le> ev b # v \<and> t' setinterleaves ((ev a # u', v'), A)\<close>
+      by (cases v') (auto simp add: notin_mvL(1, 4))
+    thus ?thesis by blast
+  next
+    case notin_mvR
+    from ev_setinterleaving_ev(5)[OF notin_mvR(1, 2, 5, 6)]
+    obtain u' v' where \<open>u' \<le> ev a # u \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>u' \<le> ev a # u \<and> ev b # v' \<le> ev b # v \<and> t' setinterleaves ((u', ev b # v'), A)\<close>
+      by (cases u') (auto simp add: notin_mvR(2, 4))
+    thus ?thesis by blast
+  qed
+next
+  case (ev_setinterleaving_tick a u v)
+  from ev_setinterleaving_tick.prems
+  obtain w w' where \<open>a \<notin> A\<close> \<open>t = ev a # w\<close> \<open>t' = [] \<or> t' = ev a # w'\<close>
+    \<open>w' \<le> w\<close> \<open>w setinterleaves ((u, \<checkmark> # v), A)\<close>
+    by (simp split: if_split_asm) (metis (no_types) Prefix_Order.prefix_Cons Nil_prefix)
+  from \<open>t' = [] \<or> t' = ev a # w'\<close> show ?case
+  proof (elim disjE)
+    from Nil_setinterleaving_Nil show \<open>t' = [] \<Longrightarrow> ?case\<close> by blast
+  next
+    assume \<open>t' = ev a # w'\<close>
+    from ev_setinterleaving_tick.hyps[OF \<open>a \<notin> A\<close> \<open>w setinterleaves ((u, \<checkmark> # v), A)\<close> \<open>w' \<le> w\<close>]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> \<checkmark> # v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>ev a # u' \<le> ev a # u \<and> v' \<le> \<checkmark> # v \<and> t' setinterleaves ((ev a # u', v'), A)\<close>
+      by (cases v') (simp_all add: \<open>a \<notin> A\<close> \<open>t' = ev a # w'\<close>)
+    thus ?case by blast
+  qed
+next
+  case (tick_setinterleaving_ev u b v)
+  from tick_setinterleaving_ev.prems
+  obtain w w' where \<open>b \<notin> A\<close> \<open>t = ev b # w\<close> \<open>t' = [] \<or> t' = ev b # w'\<close>
+    \<open>w' \<le> w\<close> \<open>w setinterleaves ((\<checkmark> # u, v), A)\<close>
+    by (simp split: if_split_asm) (metis (no_types) Prefix_Order.prefix_Cons Nil_prefix)
+  from \<open>t' = [] \<or> t' = ev b # w'\<close> show ?case
+  proof (elim disjE)
+    from Nil_setinterleaving_Nil show \<open>t' = [] \<Longrightarrow> ?case\<close> by blast
+  next
+    assume \<open>t' = ev b # w'\<close>
+    from tick_setinterleaving_ev.hyps[OF \<open>b \<notin> A\<close> \<open>w setinterleaves ((\<checkmark> # u, v), A)\<close> \<open>w' \<le> w\<close>]
+    obtain u' v' where \<open>u' \<le> \<checkmark> # u \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>u' \<le> \<checkmark> # u \<and> ev b # v' \<le> ev b # v \<and> t' setinterleaves ((u', ev b # v'), A)\<close>
+      by (cases u') (simp_all add: \<open>b \<notin> A\<close> \<open>t' = ev b # w'\<close>)
+    thus ?case by blast
+  qed
+next
+  case (tick_setinterleaving_tick u v)
+  from tick_setinterleaving_tick.prems obtain w w'
+    where \<open>t = \<checkmark> # w\<close> \<open>t' = [] \<or> t' = \<checkmark> # w'\<close>
+      \<open>w' \<le> w\<close> \<open>w setinterleaves ((u, v), A)\<close>
+    by (cases t') (auto split: option.split_asm)
+  from \<open>t' = [] \<or> t' = \<checkmark> # w'\<close> show ?case
+  proof (elim disjE)
+    from Nil_setinterleaving_Nil show \<open>t' = [] \<Longrightarrow> ?case\<close> by blast
+  next
+    assume \<open>t' = \<checkmark> # w'\<close>
+    from tick_setinterleaving_tick.hyps
+      [OF \<open>w setinterleaves ((u, v), A)\<close> \<open>w' \<le> w\<close>]
+    obtain u' v' where \<open>u' \<le> u \<and> v' \<le> v \<and> w' setinterleaves ((u', v'), A)\<close> by blast
+    hence \<open>\<checkmark> # u' \<le> \<checkmark> # u \<and> \<checkmark> # v' \<le> \<checkmark> # v \<and>
+           t' setinterleaves ((\<checkmark> # u', \<checkmark> # v'), A)\<close>
+      by (simp add: \<open>t' = \<checkmark> # w'\<close>)
+    thus ?case by blast
+  qed
+qed
+
+
+
+lift_definition Sync :: \<open>'a process \<Rightarrow> 'a set \<Rightarrow> 'a process \<Rightarrow> 'a process\<close>
+  (\<open>(3(_ \<lbrakk>_\<rbrakk>/ _))\<close> [70, 0, 71] 70)
+  is \<open>\<lambda>P A Q. {t. \<exists>t_P t_Q. t_P \<in> \<T> P \<and> t_Q \<in> \<T> Q \<and> t setinterleaves ((t_P, t_Q), A)}\<close>
+proof -
+  show \<open>?thesis P A Q\<close> (is \<open>is_process ?t\<close>) for P A Q
+  proof (unfold is_process_def, intro conjI allI impI)
+    from Nil_elem_T Nil_setinterleaving_Nil show \<open>[] \<in> ?t\<close> by blast
+  next
+    fix t assume \<open>t \<in> ?t\<close>
+    then obtain t_P t_Q where \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+      \<open>t setinterleaves ((t_P, t_Q), A)\<close> by blast
+    from \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> front_tickFree_T
+    have \<open>ftF t_P\<close> \<open>ftF t_Q\<close> by auto
+    with \<open>t setinterleaves ((t_P, t_Q), A)\<close>
+    show \<open>ftF t\<close> by (simp add: front_tickFree_setinterleaves_iff)
+  next
+    fix t u assume \<open>t @ u \<in> ?t\<close>
+    then obtain t_P t_Q where \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close>
+      \<open>t @ u setinterleaves ((t_P, t_Q), A)\<close> by blast
+    from this(3) obtain t_P' t_P'' t_Q' t_Q''
+      where \<open>t_P = t_P' @ t_P''\<close> \<open>t_Q = t_Q' @ t_Q''\<close>
+        \<open>t setinterleaves ((t_P', t_Q'), A)\<close>
+      by (meson Prefix_Order.prefixE Prefix_Order.prefixI append_setinterleaves_imp)
+    from \<open>t_P \<in> \<T> P\<close> \<open>t_Q \<in> \<T> Q\<close> this(1, 2) have \<open>t_P' \<in> \<T> P\<close> \<open>t_Q' \<in> \<T> Q\<close>
+      by (auto intro: T_dw_closed)
+    with \<open>t setinterleaves ((t_P', t_Q'), A)\<close> show \<open>t \<in> ?t\<close> by blast
+  qed
+qed
+
+lemma T_Sync :
+  \<open>\<T> (P \<lbrakk>A\<rbrakk> Q) = {t. \<exists>t_P t_Q. t_P \<in> \<T> P \<and> t_Q \<in> \<T> Q \<and> t setinterleaves ((t_P, t_Q), A)}\<close>
+  by (simp add: Sync.rep_eq)
+
+
+lift_definition Interrupt :: \<open>'a process \<Rightarrow> 'a process \<Rightarrow> 'a process\<close> (infixl \<open>\<triangle>\<close> 81)
+  is \<open>\<lambda>P Q. \<T> P \<union> {t @ u |t u. t \<in> \<T> P \<and> tF t \<and> u \<in> \<T> Q}\<close>
+proof -
+  show \<open>?thesis P Q\<close> (is \<open>is_process ?t\<close>) for P Q
+  proof (unfold is_process_def, intro conjI allI impI)
+    show \<open>[] \<in> ?t\<close> by (simp add: Nil_elem_T)
+  next
+    show \<open>t \<in> ?t \<Longrightarrow> ftF t\<close> for t
+      by (auto simp add: front_tickFree_append_iff intro: front_tickFree_T)
+  next
+    show \<open>t @ u \<in> ?t \<Longrightarrow> t \<in> ?t\<close> for t u
+      by (auto simp add: append_eq_append_conv2 intro: T_dw_closed)
+  qed
+qed
+
+
+
+
+
+subsection \<open>Constructiveness\<close>
+
+lemma restriction_process_Mprefix :
+  \<open>\<box>a \<in> A \<rightarrow> P a \<down> n = (case n of 0 \<Rightarrow> BOT | Suc m \<Rightarrow> \<box>a\<in>A \<rightarrow> (P a \<down> m))\<close>
+  by (auto simp add: process_eq_spec T_restriction_process T_Mprefix T_BOT
+      Nil_elem_T nat.case_eq_if front_tickFree_Cons_iff front_tickFree_T)
+    (metis Cons_eq_append_conv Suc_length_conv event.distinct(1) length_greater_0_conv
+      list.size(3) nat.exhaust_sel tickFree_Cons_iff)
+
+
+lemma constructive_Mprefix [simp] :
+  \<open>constructive (\<lambda>b. \<box>a\<in>A \<rightarrow> f a b)\<close> if \<open>\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)\<close>
+proof -
+  have \<open>\<box>a\<in>A \<rightarrow> f a b = \<box>a\<in>A \<rightarrow> (if a \<in> A then f a b else STOP)\<close> for b
+    by (auto simp add: process_eq_spec T_Mprefix)
+  moreover have \<open>constructive (\<lambda>b. \<box>a\<in>A \<rightarrow> (if a \<in> A then f a b else STOP))\<close>
+  proof (rule constructive_comp_non_destructive[of \<open>\<lambda>P. \<box>a\<in>A \<rightarrow> P a\<close>])
+    show \<open>constructive (\<lambda>P. \<box>a\<in>A \<rightarrow> P a)\<close>
+      by (rule constructiveI) (simp add: restriction_process_Mprefix restriction_fun_def)
+  next
+    show \<open>non_destructive (\<lambda>b a. if a \<in> A then f a b else STOP)\<close>
+      by (simp add: non_destructive_fun_iff, intro allI non_destructive_if_then_else)
+        (simp_all add: \<open>\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)\<close> non_destructiveI)
+  qed
+  ultimately show \<open>constructive (\<lambda>b. \<box>a\<in>A \<rightarrow> f a b)\<close> by simp
+qed
+
+
+
+
+
+
+subsection \<open>Non Destructiveness\<close>
+
+lemma non_destructive_Choice [simp] :
+  \<open>non_destructive (\<lambda>x. f x \<box> g x)\<close>
+  if \<open>non_destructive f\<close> \<open>non_destructive g\<close>
+  for f g :: \<open>'a :: restriction \<Rightarrow> 'b process\<close>
+proof -
+  have * : \<open>non_destructive (\<lambda>(P, Q). P \<box> Q :: 'b process)\<close>
+  proof (rule order_non_destructiveI, clarify)
+    fix P Q P' Q' :: \<open>'b process\<close> and n
+    assume \<open>(P, Q) \<down> n = (P', Q') \<down> n\<close>
+    hence \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>
+      by (simp_all add: restriction_prod_def)
+    show \<open>P \<box> Q \<down> n \<le> P' \<box> Q' \<down> n\<close>
+    proof (unfold less_eq_process_def, rule subsetI)
+      show \<open>t \<in> \<T> (P' \<box> Q' \<down> n) \<Longrightarrow> t \<in> \<T> (P \<box> Q \<down> n)\<close> for t
+      proof (elim T_restriction_processE)
+        show \<open>t \<in> \<T> (P' \<box> Q') \<Longrightarrow> length t \<le> n \<Longrightarrow> t \<in> \<T> (P \<box> Q \<down> n)\<close>
+          by (simp add: T_restriction_process T_Choice)
+            (metis (lifting) T_restriction_process T_restriction_processE
+              Un_iff \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close>)
+      next
+        show \<open>\<lbrakk>t = u @ v; u \<in> \<T> (P' \<box> Q'); length u = n; tF u; ftF v\<rbrakk>
+              \<Longrightarrow> t \<in> \<T> (P \<box> Q \<down> n)\<close> for u v
+          by (simp add: T_restriction_process T_Choice)
+            (metis (lifting) T_restriction_process T_restriction_processE Un_iff
+              \<open>P \<down> n = P' \<down> n\<close> \<open>Q \<down> n = Q' \<down> n\<close> append.right_neutral append_eq_conv_conj)
+      qed
+    qed
+  qed
+  have ** : \<open>non_destructive (\<lambda>x. (f x, g x))\<close>
+    by (fact non_destructive_prod_codomain[OF that])
+  from non_destructive_comp_non_destructive[OF * **, simplified]
+  show \<open>non_destructive (\<lambda>x. f x \<box> g x)\<close> .
+qed
+
+
+lemma restriction_process_GlobalChoice :
+  \<open>\<box>a \<in> A. P a \<down> n = (if A = {} then case n of 0 \<Rightarrow> BOT | Suc m \<Rightarrow> STOP else \<box>a \<in> A. (P a \<down> n))\<close>
+  by (auto simp add: process_eq_spec T_restriction_process T_GlobalChoice T_BOT T_STOP
+      split: nat.split)
+
+
+lemma non_destructive_GlobalChoice [simp] :
+  \<open>non_destructive (\<lambda>b. \<box>a\<in>A. f a b)\<close> if \<open>\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)\<close>
+proof -
+  have \<open>\<box>a\<in>A. f a b = \<box>a\<in>A. (if a \<in> A then f a b else STOP)\<close> for b
+    by (auto simp add: process_eq_spec T_GlobalChoice)
+  moreover have \<open>non_destructive (\<lambda>b. \<box>a\<in>A. (if a \<in> A then f a b else STOP))\<close>
+  proof (rule non_destructive_comp_non_destructive[of \<open>\<lambda>P. \<box>a\<in>A. P a\<close>])
+    show \<open>non_destructive (\<lambda>P. \<box>a\<in>A. P a)\<close>
+      by (rule non_destructiveI) (simp add: restriction_process_GlobalChoice restriction_fun_def)
+  next
+    show \<open>non_destructive (\<lambda>b a. if a \<in> A then f a b else STOP)\<close>
+      by (simp add: non_destructive_fun_iff, intro allI non_destructive_if_then_else)
+        (simp_all add: \<open>\<And>a. a \<in> A \<Longrightarrow> non_destructive (f a)\<close> non_destructiveI)
+  qed
+  ultimately show \<open>non_destructive (\<lambda>b. \<box>a\<in>A. f a b)\<close> by simp
+qed
+
+
+
+subsection \<open>Examples\<close>
+
+notepad begin
+  fix A B :: \<open>'b \<Rightarrow> 'a set\<close>
+  define P :: \<open>'b \<Rightarrow> 'a process\<close>
+    where \<open>P \<equiv> \<upsilon> X. (\<lambda>s. \<box> a \<in> A s \<rightarrow> X s \<box> (\<box> b \<in> B s \<rightarrow> X s))\<close> (is \<open>P \<equiv> \<upsilon> X. ?f X\<close>)
+  have \<open>P = ?f P\<close>
+    by (unfold P_def, subst restriction_fix_eq) simp_all
+
+end
+
+
+lemma \<open>constructive (\<lambda>X \<sigma>. \<box> e \<in> f \<sigma> \<rightarrow> \<box> \<sigma>' \<in> g \<sigma> e. X \<sigma>')\<close>
+  by simp
+
+
+
+lemma length_le_T_restriction_process_iff_T :
+  \<open>length t \<le> n \<Longrightarrow> t \<in> \<T> (P \<down> n) \<longleftrightarrow> t \<in> \<T> P\<close>
+  by (auto simp add: T_restriction_process)
+
+
+
+lemma restriction_adm_notin_T [simp] : \<open>adm\<down> (\<lambda>a. t \<notin> \<T> a)\<close>
+proof (rule restriction_admI)
+  fix \<sigma> and \<Sigma> assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. t \<notin> \<T> (\<sigma> n)\<close>
+  from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> obtain n0 where \<open>\<forall>k\<ge>n0. \<Sigma> \<down> length t = \<sigma> k \<down> length t\<close>
+    by (blast dest: restriction_tendstoD)
+  hence \<open>\<forall>k\<ge>n0. \<T> (\<Sigma> \<down> length t) = \<T> (\<sigma> k \<down> length t)\<close> by simp
+  hence \<open>\<forall>k\<ge>n0. t \<in> \<T> \<Sigma> \<longleftrightarrow> t \<in> \<T> (\<sigma> k)\<close>
+    by (metis dual_order.refl length_le_T_restriction_process_iff_T)
+  with \<open>\<And>n. t \<notin> \<T> (\<sigma> n)\<close> show \<open>t \<notin> \<T> \<Sigma>\<close> by blast
+qed
+
+
+lemma restriction_adm_in_T [simp] : \<open>adm\<down> (\<lambda>a. t \<in> \<T> a)\<close>
+proof (rule restriction_admI)
+  fix \<sigma> and \<Sigma> assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. t \<in> \<T> (\<sigma> n)\<close>
+  from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> obtain n0 where \<open>\<forall>k\<ge>n0. \<Sigma> \<down> length t = \<sigma> k \<down> length t\<close>
+    by (blast dest: restriction_tendstoD)
+  hence \<open>\<forall>k\<ge>n0. \<T> (\<Sigma> \<down> length t) = \<T> (\<sigma> k \<down> length t)\<close> by simp
+  hence \<open>\<forall>k\<ge>n0. t \<in> \<T> \<Sigma> \<longleftrightarrow> t \<in> \<T> (\<sigma> k)\<close>
+    by (metis dual_order.refl length_le_T_restriction_process_iff_T)
+  with \<open>\<And>n. t \<in> \<T> (\<sigma> n)\<close> show \<open>t \<in> \<T> \<Sigma>\<close> by blast
+qed
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/RS_Tree.thy
--- /dev/null
+++ thys/Restriction_Spaces-Examples/RS_Tree.thy
@@ -0,0 +1,97 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Binary Trees\<close>
+
+(*<*)
+theory RS_Tree
+  imports Restriction_Spaces
+begin
+  (*>*)
+
+
+datatype 'a ex_tree = tip | node \<open>'a ex_tree\<close> 'a \<open>'a ex_tree\<close>
+
+
+instantiation ex_tree :: (type) restriction
+begin
+
+fun restriction_ex_tree :: \<open>'a ex_tree \<Rightarrow> nat \<Rightarrow> 'a ex_tree\<close>
+  where \<open>tip \<down> n = tip\<close>
+  |     \<open>(node l val r) \<down> 0 = tip\<close>
+  |     \<open>(node l val r) \<down> Suc n = node (l \<down> n) val (r \<down> n)\<close>
+
+
+lemma restriction_ex_tree_0_is_tip [simp] : \<open>T \<down> 0 = tip\<close>
+  using restriction_ex_tree.elims by blast
+
+instance
+proof intro_classes
+  show \<open>T \<down> n \<down> m = T \<down> min n m\<close> for T :: \<open>'a ex_tree\<close> and n m
+  proof (induct n arbitrary: T m)
+    show \<open>T \<down> 0 \<down> m = T \<down> min 0 m\<close> for T :: \<open>'a ex_tree\<close> and m by simp
+  next
+    fix T :: \<open>'a ex_tree\<close> and m n assume hyp : \<open>T \<down> n \<down> m = T \<down> min n m\<close> for T :: \<open>'a ex_tree\<close> and m
+    show \<open>T \<down> Suc n \<down> m = T \<down> min (Suc n) m\<close>
+      by (cases T; cases m, simp_all add: hyp)
+  qed
+qed
+
+
+end
+
+
+lemma size_le_imp_restriction_ex_tree_eq_self :
+  \<open>size x \<le> n \<Longrightarrow> x \<down> n = x\<close> for x :: \<open>'a ex_tree\<close>
+  by (induct rule: restriction_ex_tree.induct) simp_all
+
+lemma restriction_ex_tree_eqI :
+  \<open>(\<And>i. x \<down> i =  y \<down> i) \<Longrightarrow> x = y\<close> for x y :: \<open>'a ex_tree\<close>
+  by (metis linorder_linear size_le_imp_restriction_ex_tree_eq_self)
+
+lemma restriction_ex_tree_eqI_optimized :
+  \<open>(\<And>i. i \<le> max (size x) (size y) \<Longrightarrow> x \<down> i =  y \<down> i) \<Longrightarrow> x = y\<close> for x y :: \<open>'a ex_tree\<close>
+  by (metis max.cobounded1 max.cobounded2 order_eq_refl size_le_imp_restriction_ex_tree_eq_self)
+
+
+instance ex_tree :: (type) restriction_space
+  by (intro_classes, simp)
+    (use restriction_ex_tree_eqI_optimized in blast)
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Examples/document/root.tex
--- /dev/null
+++ thys/Restriction_Spaces-Examples/document/root.tex
@@ -0,0 +1,71 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{pifont}
+
+%\usepackage{latexsym}
+\usepackage{amssymb}
+\usepackage{amsmath}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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}
+
+
+%\pagestyle{plain}
+
+\begin{document}
+
+\title{Miscellaneous Examples of restriction Spaces}
+\author{Benoît Ballenghien \and Benjamin Puyobro \and Burkhart Wolff}
+\maketitle
+
+\abstract{
+  In this session, a number of examples are provided to illustrate how the \verb'Restriction_Spaces' library works.
+  The simple cases are, of course, covered: trivial construction, booleans, integers, option type, and so on.
+  More elaborate situations are also covered, such as formal series and a trace model of the CSP process algebra.
+}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Banach_Theorem_Extension.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Banach_Theorem_Extension.thy
@@ -0,0 +1,363 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Definitions on Functions of Metric Space\<close>
+
+(*<*)
+theory  Banach_Theorem_Extension
+  imports "HOL-Analysis.Elementary_Metric_Spaces"
+begin
+  (*>*)
+
+text \<open>In this theory, we define the notion of lipschitz map, non-expanding map
+      and contraction map. We also establish correspondences.\<close>
+
+
+subsection \<open>Definitions\<close>
+
+subsubsection \<open>Lipschitz Map\<close>
+
+text \<open>This notion is a generalization of contraction map and non-expanding map.\<close>
+
+definition lipschitz_with_on :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, real, 'a set] \<Rightarrow> bool\<close>
+  where \<open>lipschitz_with_on f \<alpha> A \<equiv> 0 \<le> \<alpha> \<and> (\<forall>x\<in>A. \<forall>y\<in>A. dist (f x) (f y) \<le> \<alpha> * dist x y)\<close>
+
+abbreviation lipschitz_with :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, real] \<Rightarrow> bool\<close>
+  where \<open>lipschitz_with f \<alpha> \<equiv> lipschitz_with_on f \<alpha> UNIV\<close>
+
+
+lemma lipschitz_with_onI :
+  \<open>\<lbrakk>0 \<le> \<alpha>; \<And>x y. \<lbrakk>x \<in> A; y \<in> A; x \<noteq> y; f x \<noteq> f y\<rbrakk> \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<rbrakk> \<Longrightarrow>
+   lipschitz_with_on f \<alpha> A\<close>
+  unfolding lipschitz_with_on_def by (metis dist_eq_0_iff zero_le_dist zero_le_mult_iff)
+
+lemma lipschitz_withI :
+  \<open>\<lbrakk>0 \<le> \<alpha>; \<And>x y. x \<noteq> y \<Longrightarrow> f x \<noteq> f y \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<rbrakk> \<Longrightarrow> 
+   lipschitz_with f \<alpha>\<close>
+  by (rule lipschitz_with_onI)
+
+lemma lipschitz_with_onD1 : \<open>lipschitz_with_on f \<alpha> A \<Longrightarrow> 0 \<le> \<alpha>\<close>
+  unfolding lipschitz_with_on_def by simp
+
+lemma lipschitz_withD1 : \<open>lipschitz_with f \<alpha> \<Longrightarrow> 0 \<le> \<alpha>\<close>
+  by (rule lipschitz_with_onD1)
+
+lemma lipschitz_with_onD2 :
+  \<open>lipschitz_with_on f \<alpha> A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<close>
+  unfolding lipschitz_with_on_def by simp
+
+lemma lipschitz_withD2 :
+  \<open>lipschitz_with f \<alpha> \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<close>
+  unfolding lipschitz_with_on_def by simp
+
+lemma lipschitz_with_imp_lipschitz_with_on: \<open>lipschitz_with f \<alpha> \<Longrightarrow> lipschitz_with_on f \<alpha> A\<close>
+  by (simp add: lipschitz_with_on_def)
+
+
+lemma lipschitz_with_on_imp_lipschitz_with_on_ge : \<open>lipschitz_with_on f \<beta> A\<close>
+  if \<open>\<alpha> \<le> \<beta>\<close> and \<open>lipschitz_with_on f \<alpha> A\<close>
+proof (rule lipschitz_with_onI)
+  show \<open>0 \<le> \<beta>\<close> by (meson order_trans lipschitz_with_onD1 that)
+next
+  fix x y assume \<open>x \<in> A\<close> and \<open>y \<in> A\<close>
+  from \<open>lipschitz_with_on f \<alpha> A\<close>[THEN lipschitz_with_onD2, OF this]
+  have \<open>dist (f x) (f y) \<le> \<alpha> * dist x y\<close> .
+  from order_trans[OF this] show \<open>dist (f x) (f y) \<le> \<beta> * dist x y\<close>
+    by (simp add: mult_right_mono \<open>\<alpha> \<le> \<beta>\<close>)
+qed
+
+
+theorem lipschitz_with_on_comp_lipschitz_with_on :
+  \<open>lipschitz_with_on (\<lambda>x. g (f x)) (\<beta> * \<alpha>) A\<close>
+  if \<open>f ` A \<subseteq> B\<close> \<open>lipschitz_with_on g \<beta> B\<close> \<open>lipschitz_with_on f \<alpha> A\<close>
+proof (rule lipschitz_with_onI)
+  show \<open>0 \<le> \<beta> * \<alpha>\<close> by (metis lipschitz_with_onD1 mult_nonneg_nonneg that(2, 3))
+next
+  fix x y assume \<open>x \<in> A\<close> and \<open>y \<in> A\<close>
+  with \<open>f ` A \<subseteq> B\<close> have \<open>f x \<in> B\<close> and \<open>f y \<in> B\<close> by auto
+  have \<open>dist (g (f x)) (g (f y)) \<le> \<beta> * dist (f x) (f y)\<close>
+    by (fact that(2)[THEN lipschitz_with_onD2, OF \<open>f x \<in> B\<close> \<open>f y \<in> B\<close>])
+  also have \<open>\<dots> \<le> \<beta> * (\<alpha> * dist x y)\<close>
+    by (fact mult_left_mono[OF that(3)[THEN lipschitz_with_onD2, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>]
+                               that(2)[THEN lipschitz_with_onD1]])
+  also have \<open>\<dots> = \<beta> * \<alpha> * dist x y\<close> by simp
+  finally show \<open>dist (g (f x)) (g (f y)) \<le> \<dots>\<close> .
+qed
+
+corollary lipschitz_with_comp_lipschitz_with : 
+  \<open>\<lbrakk>lipschitz_with g \<beta>; lipschitz_with f \<alpha>\<rbrakk> \<Longrightarrow>
+   lipschitz_with (\<lambda>x. g (f x)) (\<beta> * \<alpha>)\<close>
+  using lipschitz_with_on_comp_lipschitz_with_on by blast
+
+
+
+subsubsection \<open>Non-expanding Map\<close>
+
+definition non_expanding_on :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, 'a set] \<Rightarrow> bool\<close>
+  where \<open>non_expanding_on f A \<equiv> lipschitz_with_on f 1 A\<close>
+
+abbreviation non_expanding :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space] \<Rightarrow> bool\<close>
+  where \<open>non_expanding f \<equiv> non_expanding_on f UNIV\<close>
+
+lemma non_expanding_onI :
+  \<open>\<lbrakk>\<And>x y. \<lbrakk>x \<in> A; y \<in> A; x \<noteq> y; f x \<noteq> f y\<rbrakk> \<Longrightarrow> dist (f x) (f y) \<le> dist x y\<rbrakk> \<Longrightarrow> 
+   non_expanding_on f A\<close>
+  by (simp add: lipschitz_with_onI non_expanding_on_def)
+
+lemma non_expandingI :
+  \<open>\<lbrakk>\<And>x y. x \<noteq> y \<Longrightarrow> f x \<noteq> f y \<Longrightarrow> dist (f x) (f y) \<le> dist x y\<rbrakk> \<Longrightarrow> non_expanding f\<close>
+  by (rule non_expanding_onI)
+
+lemma non_expanding_onD :
+  \<open>non_expanding_on f A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> dist (f x) (f y) \<le> dist x y\<close>
+  by (metis lipschitz_with_onD2 mult_1 non_expanding_on_def)
+
+lemma non_expandingD : \<open>non_expanding f \<Longrightarrow> dist (f x) (f y) \<le> dist x y\<close>
+  by (simp add: non_expanding_onD)
+
+lemma non_expanding_imp_non_expanding_on: \<open>non_expanding f \<Longrightarrow> non_expanding_on f A\<close>
+  by (meson non_expandingD non_expanding_onI)
+
+
+lemma non_expanding_on_comp_non_expanding_on :
+  \<open>\<lbrakk>f ` A \<subseteq> B; non_expanding_on g B; non_expanding_on f A\<rbrakk> \<Longrightarrow>
+   non_expanding_on (\<lambda>x. g (f x)) A\<close>
+  unfolding non_expanding_on_def
+  by (metis (no_types) lipschitz_with_on_comp_lipschitz_with_on mult_1)
+
+corollary non_expanding_comp_non_expanding  :
+  \<open>\<lbrakk>non_expanding g; non_expanding f\<rbrakk> \<Longrightarrow> non_expanding (\<lambda>x. g (f x))\<close>
+  by (blast intro: non_expanding_on_comp_non_expanding_on)
+
+
+
+subsubsection \<open>Contraction Map\<close>
+
+definition contraction_with_on :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, real, 'a set] \<Rightarrow> bool\<close>
+  where \<open>contraction_with_on f \<alpha> A \<equiv> \<alpha> < 1 \<and> lipschitz_with_on f \<alpha> A\<close>
+
+abbreviation contraction_with :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, real] \<Rightarrow> bool\<close>
+  where \<open>contraction_with f \<alpha> \<equiv> contraction_with_on f \<alpha> UNIV\<close>
+
+definition contraction_on :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space, 'a set] \<Rightarrow> bool\<close>
+  where \<open>contraction_on f A \<equiv> \<exists>\<alpha>. contraction_with_on f \<alpha> A\<close>
+
+abbreviation contraction :: \<open>['a :: metric_space \<Rightarrow> 'b :: metric_space] \<Rightarrow> bool\<close>
+  where \<open>contraction f \<equiv> contraction_on f UNIV\<close>
+
+
+lemma contraction_with_onI :
+  \<open>\<lbrakk>0 \<le> \<alpha>; \<alpha> < 1; \<And>x y. \<lbrakk>x \<in> A; y \<in> A; x \<noteq> y; f x \<noteq> f y\<rbrakk> \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<rbrakk>
+   \<Longrightarrow> contraction_with_on f \<alpha> A\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_onI)
+
+lemma contraction_withI :
+  \<open>\<lbrakk>0 \<le> \<alpha>; \<alpha> < 1; \<And>x y. x \<noteq> y \<Longrightarrow> f x \<noteq> f y \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<rbrakk> \<Longrightarrow>
+   contraction_with f \<alpha>\<close>
+  by (rule contraction_with_onI)
+
+lemma contraction_with_onD1 : \<open>contraction_with_on f \<alpha> A \<Longrightarrow> 0 \<le> \<alpha>\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_on_def)
+
+lemma contraction_withD1 : \<open>contraction_with f \<alpha> \<Longrightarrow> 0 \<le> \<alpha>\<close>
+  by (simp add: contraction_with_onD1)
+
+lemma contraction_with_onD2 : \<open>contraction_with_on f \<alpha> A \<Longrightarrow> \<alpha> < 1\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_on_def)
+
+lemma contraction_withD2 : \<open>contraction_with f \<alpha> \<Longrightarrow> \<alpha> < 1\<close>
+  by (simp add: contraction_with_onD2)
+
+lemma contraction_with_onD3 :
+  \<open>contraction_with_on f \<alpha> A \<Longrightarrow> x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_on_def)
+
+lemma contraction_withD3 : \<open>contraction_with f \<alpha> \<Longrightarrow> dist (f x) (f y) \<le> \<alpha> * dist x y\<close>
+  by (simp add: contraction_with_on_def lipschitz_withD2)
+
+lemma contraction_with_imp_contraction_with_on:
+  \<open>contraction_with f \<alpha> \<Longrightarrow> contraction_with_on f \<alpha> A\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_imp_lipschitz_with_on)
+
+lemma contraction_imp_contraction_on: \<open>contraction f \<Longrightarrow> contraction_on f A\<close>
+  using contraction_on_def contraction_with_imp_contraction_with_on by blast
+
+
+lemma contraction_with_on_imp_contraction_on :
+  \<open>contraction_with_on f \<alpha> A \<Longrightarrow> contraction_on f A\<close>
+  unfolding contraction_on_def by blast
+
+lemma contraction_with_imp_contraction: \<open>contraction_with f \<alpha> \<Longrightarrow> contraction f\<close>
+  by (simp add: contraction_with_on_imp_contraction_on)
+
+
+lemma contraction_onE:
+  \<open>\<lbrakk>contraction_on f A; \<And>\<alpha>. contraction_with_on f \<alpha> A \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis\<close>
+  unfolding contraction_on_def by blast
+
+lemma contractionE:
+  \<open>\<lbrakk>contraction f; \<And>\<alpha>. contraction_with f \<alpha> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis\<close>
+  by (elim contraction_onE)
+
+
+lemma contraction_with_on_imp_contraction_with_on_ge :
+  \<open>\<lbrakk>\<alpha> \<le> \<beta>; \<beta> < 1; contraction_with_on f \<alpha> A\<rbrakk> \<Longrightarrow> contraction_with_on f \<beta> A\<close>
+  by (simp add: contraction_with_on_def lipschitz_with_on_imp_lipschitz_with_on_ge)
+
+
+
+subsection \<open>Properties\<close>
+
+lemma contraction_with_on_imp_lipschitz_with_on[simp] :
+  \<open>contraction_with_on f \<alpha> A \<Longrightarrow> lipschitz_with_on f \<alpha> A\<close>
+  by (simp add: contraction_with_on_def)
+
+lemma non_expanding_on_imp_lipschitz_with_one_on[simp] :
+  \<open>non_expanding_on f A \<Longrightarrow> lipschitz_with_on f 1 A\<close>
+  by (simp add: non_expanding_on_def)
+
+lemma contraction_on_imp_non_expanding_on[simp] :
+  \<open>contraction_on f A \<Longrightarrow> non_expanding_on f A\<close> 
+proof (elim contraction_onE, rule non_expanding_onI)
+  fix \<alpha> x y assume \<open>x \<in> A\<close> and \<open>y \<in> A\<close> and contra : \<open>contraction_with_on f \<alpha> A\<close>
+  show \<open>dist (f x) (f y) \<le> dist x y\<close>
+    by (rule order_trans[OF contra[THEN contraction_with_onD3, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>]])
+      (metis contra contraction_with_on_def mult_less_cancel_right1 nle_le order_less_le zero_le_dist)
+qed
+
+
+lemma contraction_with_on_comp_contraction_with_on :
+  \<open>contraction_with_on (\<lambda>x. g (f x)) (\<beta> * \<alpha>) A\<close>
+  if \<open>f ` A \<subseteq> B\<close> \<open>contraction_with_on g \<beta> B\<close> \<open>contraction_with_on f \<alpha> A\<close>
+proof (unfold contraction_with_on_def, intro conjI)
+  from that(2, 3)[THEN contraction_with_onD2] that(2)[THEN contraction_with_onD1]
+  show \<open>\<beta> * \<alpha> < 1\<close> by (metis dual_order.strict_trans2 mult_less_cancel_left1 nle_le order_less_le)
+next
+  show \<open>lipschitz_with_on (\<lambda>x. g (f x)) (\<beta> * \<alpha>) A\<close>
+    by (rule lipschitz_with_on_comp_lipschitz_with_on[OF \<open>f ` A \<subseteq> B\<close>])
+      (simp_all add: that(2, 3))
+qed
+
+corollary contraction_with_comp_contraction_with :
+  \<open>\<lbrakk>contraction_with g \<beta>; contraction_with f \<alpha>\<rbrakk> \<Longrightarrow> contraction_with (\<lambda>x. g (f x)) (\<beta> * \<alpha>)\<close>
+  by (blast intro: contraction_with_on_comp_contraction_with_on)
+
+corollary contraction_on_comp_contraction_on :
+  \<open>\<lbrakk>f ` A \<subseteq> B; contraction_on g B; contraction_on f A\<rbrakk> \<Longrightarrow> contraction_on (\<lambda>x. g (f x)) A\<close>
+proof (elim contraction_onE)
+  fix \<alpha> \<beta> assume \<open>f ` A \<subseteq> B\<close> \<open>contraction_with_on g \<beta> B\<close> \<open>contraction_with_on f \<alpha> A\<close>
+  from contraction_with_on_comp_contraction_with_on[OF this]
+  show \<open>contraction_on (\<lambda>x. g (f x)) A\<close> by (fact contraction_with_on_imp_contraction_on)
+qed
+
+corollary contraction_comp_contraction :
+  \<open>\<lbrakk>contraction g; contraction f\<rbrakk> \<Longrightarrow> contraction (\<lambda>x. g (f x))\<close>
+  by (blast intro: contraction_on_comp_contraction_on)
+
+
+lemma contraction_with_on_comp_non_expanding_on :
+  \<open>contraction_with_on (\<lambda>x. g (f x)) \<beta> A\<close>
+  if \<open>f ` A \<subseteq> B\<close> \<open>contraction_with_on g \<beta> B\<close> \<open>non_expanding_on f A\<close>
+proof (unfold contraction_with_on_def, intro conjI)
+  from that(2)[THEN contraction_with_onD2] show \<open>\<beta> < 1\<close> .
+next
+  show \<open>lipschitz_with_on (\<lambda>x. g (f x)) \<beta> A\<close>
+    by (rule lipschitz_with_on_comp_lipschitz_with_on[OF \<open>f ` A \<subseteq> B\<close>, of g \<beta> 1, simplified])
+      (simp_all add: that(2, 3))
+qed
+
+corollary contraction_with_comp_non_expanding  :
+  \<open>\<lbrakk>contraction_with g \<beta>; non_expanding f\<rbrakk> \<Longrightarrow> contraction_with (\<lambda>x. g (f x)) \<beta>\<close>
+  by (blast intro: contraction_with_on_comp_non_expanding_on)
+
+corollary contraction_on_comp_non_expanding_on :
+  \<open>\<lbrakk>f ` A \<subseteq> B; contraction_on g B; non_expanding_on f A\<rbrakk> \<Longrightarrow> contraction_on (\<lambda>x. g (f x)) A\<close>
+  by (metis contraction_on_def contraction_with_on_comp_non_expanding_on)
+
+corollary contraction_comp_non_expanding  :
+  \<open>\<lbrakk>contraction g; non_expanding f\<rbrakk> \<Longrightarrow> contraction (\<lambda>x. g (f x))\<close>
+  by (blast intro: contraction_on_comp_non_expanding_on)
+
+
+lemma non_expanding_on_comp_contraction_with_on :
+  \<open>contraction_with_on (\<lambda>x. g (f x)) \<alpha> A\<close>
+  if \<open>f ` A \<subseteq> B\<close> \<open>non_expanding_on g B\<close> \<open>contraction_with_on f \<alpha> A\<close>
+proof (unfold contraction_with_on_def, intro conjI)
+  from that(3)[THEN contraction_with_onD2] show \<open>\<alpha> < 1\<close> .
+next
+  show \<open>lipschitz_with_on (\<lambda>x. g (f x)) \<alpha> A\<close>
+    by (rule lipschitz_with_on_comp_lipschitz_with_on[OF \<open>f ` A \<subseteq> B\<close>, of g 1, simplified])
+      (simp_all add: that(2, 3))
+qed
+
+corollary non_expanding_comp_contraction_with :
+  \<open>\<lbrakk>non_expanding g; contraction_with f \<alpha>\<rbrakk> \<Longrightarrow> contraction_with (\<lambda>x. g (f x)) \<alpha>\<close>
+  by (blast intro: non_expanding_on_comp_contraction_with_on)
+
+corollary non_expanding_on_comp_contraction_on :
+  \<open>\<lbrakk>f ` A \<subseteq> B; non_expanding_on g B; contraction_on f A\<rbrakk> \<Longrightarrow> contraction_on (\<lambda>x. g (f x)) A\<close>
+  by (metis contraction_on_def non_expanding_on_comp_contraction_with_on)
+
+corollary non_expanding_comp_contraction :
+  \<open>\<lbrakk>non_expanding g; contraction f\<rbrakk> \<Longrightarrow> contraction (\<lambda>x. g (f x))\<close>
+  by (blast intro: non_expanding_on_comp_contraction_on)
+
+
+
+subsection \<open>Banach's fixed-point Theorems\<close>
+
+text \<open>We rewrite the Banach's fixed-point theorems with our new definition.\<close>
+
+theorem Banach_fix_type : \<open>contraction f \<Longrightarrow> \<exists>!x. f x = x\<close>
+  for f :: \<open>'a :: complete_space \<Rightarrow> 'a\<close>
+  by (elim contractionE)
+    (metis banach_fix_type contraction_withD1 contraction_withD2 contraction_withD3)
+
+theorem Banach_fix:
+  \<open>contraction_on f s \<Longrightarrow> \<exists>!x. x \<in> s \<and> f x = x\<close> if \<open>complete s\<close> \<open>s \<noteq> {}\<close> \<open>f ` s \<subseteq> s\<close>
+proof (elim contraction_onE, intro Banach_fix[OF \<open>complete s\<close> \<open>s \<noteq> {}\<close> _ _ \<open>f ` s \<subseteq> s\<close>] ballI)
+  show \<open>contraction_with_on f \<alpha> s \<Longrightarrow> 0 \<le> \<alpha>\<close> for \<alpha> by (fact contraction_with_onD1)
+next
+  show \<open>contraction_with_on f \<alpha> s \<Longrightarrow> \<alpha> < 1\<close> for \<alpha> by (fact contraction_with_onD2)
+next
+  show \<open>contraction_with_on f \<alpha> s \<Longrightarrow> x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow>
+        dist (f x) (f y) \<le> \<alpha> * dist x y\<close> for \<alpha> x y by (fact contraction_with_onD3)
+qed
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Fun_Ultrametric_Restriction_Spaces.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Fun_Ultrametric_Restriction_Spaces.thy
@@ -0,0 +1,264 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+
+section \<open>Functions\<close>
+
+(*<*)
+theory Fun_Ultrametric_Restriction_Spaces
+  imports Ultrametric_Restriction_Spaces Banach_Theorem_Extension
+begin
+  (*>*)
+
+
+subsection \<open>Restriction Space\<close>
+
+
+instantiation \<open>fun\<close> :: (type, restriction_\<sigma>) restriction_\<sigma>
+begin
+
+definition restriction_\<sigma>_fun :: \<open>('a \<Rightarrow> 'b) itself \<Rightarrow> nat \<Rightarrow> real\<close>
+  where \<open>restriction_\<sigma>_fun _ \<equiv> restriction_\<sigma> TYPE('b)\<close>
+
+instance by intro_classes
+
+end
+
+instantiation \<open>fun\<close> :: (type, non_decseq_restriction_space) non_decseq_restriction_space
+begin
+
+definition dist_fun :: \<open>['a \<Rightarrow> 'b, 'a \<Rightarrow> 'b] \<Rightarrow> real\<close>
+  where \<open>dist_fun f g \<equiv> INF n \<in> restriction_related_set f g. restriction_\<sigma> TYPE('a \<Rightarrow> 'b) n\<close>
+
+definition uniformity_fun :: \<open>(('a \<Rightarrow> 'b) \<times> ('a \<Rightarrow> 'b)) filter\<close>
+  where \<open>uniformity_fun \<equiv> INF e\<in>{0 <..}. principal {(x, y). dist x y < e}\<close>
+
+definition open_fun :: \<open>('a \<Rightarrow> 'b) set \<Rightarrow> bool\<close>
+  where \<open>open_fun U \<equiv> \<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity\<close>
+
+instance by intro_classes
+    (simp_all add: restriction_\<sigma>_fun_def dist_fun_def open_fun_def
+      uniformity_fun_def restriction_\<sigma>_tendsto_zero)
+
+end
+
+
+instance \<open>fun\<close> :: (type, decseq_restriction_space) decseq_restriction_space
+  by intro_classes (simp add: restriction_\<sigma>_fun_def decseq_restriction_\<sigma>)
+
+instance \<open>fun\<close> :: (type, strict_decseq_restriction_space) strict_decseq_restriction_space
+  by intro_classes
+    (simp add: restriction_\<sigma>_fun_def strict_decseq_restriction_\<sigma>)
+
+
+instantiation \<open>fun\<close> :: (type, restriction_\<delta>) restriction_\<delta>
+begin
+
+definition restriction_\<delta>_fun :: \<open>('a \<Rightarrow> 'b) itself \<Rightarrow> real\<close>
+  where \<open>restriction_\<delta>_fun _ \<equiv> restriction_\<delta> TYPE('b)\<close>
+
+instance by intro_classes (simp_all add: restriction_\<delta>_fun_def)
+
+end
+
+
+instance \<open>fun\<close> :: (type, at_least_geometric_restriction_space) at_least_geometric_restriction_space
+proof intro_classes
+  show \<open>0 < restriction_\<sigma> TYPE('a \<Rightarrow> 'b) n\<close> for n
+    by (simp add: restriction_\<sigma>_fun_def)
+next
+  show \<open>restriction_\<sigma> TYPE('a \<Rightarrow> 'b) (Suc n)
+        \<le> restriction_\<delta> TYPE('a \<Rightarrow> 'b) * restriction_\<sigma> TYPE('a \<Rightarrow> 'b) n\<close> for n
+    by (simp add: restriction_\<sigma>_le restriction_\<sigma>_fun_def restriction_\<delta>_fun_def)
+next
+  show \<open>dist f g = Inf (restriction_\<sigma>_related_set f g)\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (simp add: dist_fun_def)
+qed
+
+instance \<open>fun\<close> :: (type, geometric_restriction_space) geometric_restriction_space
+proof intro_classes
+  show \<open>restriction_\<sigma> TYPE('a \<Rightarrow> 'b) n = restriction_\<delta> TYPE('a \<Rightarrow> 'b) ^ n\<close> for n
+    by (simp add: restriction_\<sigma>_fun_def restriction_\<sigma>_is restriction_\<delta>_fun_def)
+next
+  show \<open>dist f g = Inf (restriction_\<sigma>_related_set f g)\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (simp add: dist_fun_def)
+qed
+
+
+lemma dist_image_le_dist_fun : \<open>dist (f x) (g x) \<le> dist f g\<close>
+  for f g :: \<open>'a \<Rightarrow> 'b :: non_decseq_restriction_space\<close>
+proof (unfold dist_restriction_is, rule cInf_superset_mono)
+  show \<open>restriction_\<sigma>_related_set f g \<noteq> {}\<close> by simp
+next
+  show \<open>bdd_below (restriction_\<sigma>_related_set (f x) (g x))\<close>
+    by (simp add: bounded_imp_bdd_below bounded_restriction_\<sigma>_related_set)
+next
+  show \<open>restriction_\<sigma>_related_set f g \<subseteq> restriction_\<sigma>_related_set (f x) (g x)\<close>
+    unfolding restriction_fun_def restriction_\<sigma>_fun_def 
+    by (simp add: image_def subset_iff) metis
+qed
+
+
+lemma Sup_dist_image_le_dist_fun : \<open>(SUP x. dist (f x) (g x)) \<le> dist f g\<close>
+  for f g :: \<open>'a \<Rightarrow> 'b :: non_decseq_restriction_space\<close>
+  by (simp add: dist_image_le_dist_fun cSUP_least)
+    \<comment> \<open>The other inequality will require the additional assumption const\<open>decseq\<close>.\<close>
+
+
+
+context fixes f g :: \<open>'a \<Rightarrow> 'b :: decseq_restriction_space\<close> begin
+
+lemma reached_dist_fun : \<open>\<exists>x. dist f g = dist (f x) (g x)\<close>
+proof (cases \<open>f = g\<close>)
+  show \<open>f = g \<Longrightarrow> \<exists>x. dist f g = dist (f x) (g x)\<close> by simp
+next
+  assume \<open>f \<noteq> g\<close>
+  let ?n = \<open>Sup (restriction_related_set f g)\<close>
+  from not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(2, 3)[OF \<open>f \<noteq> g\<close>]
+  obtain x where \<open>\<forall>m\<le>?n. f x \<down> m = g x \<down> m\<close> \<open>f x \<down> Suc ?n \<noteq> g x \<down> Suc ?n\<close>
+    unfolding restriction_fun_def by (meson lessI)
+  hence \<open>dist (f x) (g x) = restriction_\<sigma> TYPE('b) ?n\<close>
+    by (metis (no_types, lifting) le_neq_implies_less not_less_eq_eq
+        not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified)
+  with not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(1)
+    [OF \<open>f \<noteq> g\<close>, unfolded restriction_\<sigma>_fun_def]
+  have \<open>dist f g = dist (f x) (g x)\<close> by simp
+  thus \<open>\<exists>x. dist f g = dist (f x) (g x)\<close> ..
+qed
+
+
+lemma dist_fun_eq_Sup_dist_image : \<open>dist f g = (SUP x. dist (f x) (g x))\<close>
+proof (rule order_antisym)
+  show \<open>(SUP x. dist (f x) (g x)) \<le> dist f g\<close> by (fact Sup_dist_image_le_dist_fun)
+next
+  from reached_dist_fun obtain x where \<open>dist f g = dist (f x) (g x)\<close> ..
+  thus \<open>dist f g \<le> (SUP x. dist (f x) (g x))\<close>
+  proof (rule ord_eq_le_trans)
+    show \<open>dist (f x) (g x) \<le> (SUP x. dist (f x) (g x))\<close>
+    proof (rule cSup_upper)
+      show \<open>dist (f x) (g x) \<in> range (\<lambda>x. dist (f x) (g x))\<close> by simp
+    next
+      show \<open>bdd_above (range (\<lambda>x. dist (f x) (g x)))\<close>
+        by (rule bdd_aboveI[of _ \<open>dist f g\<close>]) (auto intro: dist_image_le_dist_fun)
+    qed
+  qed
+qed
+
+
+lemma fun_restriction_space_Sup_properties :
+  \<open>dist (f x) (g x) \<le> dist f g\<close>
+  \<open>(\<And>x. dist (f x) (g x) \<le> b) \<Longrightarrow> dist f g \<le> b\<close>
+  by (use reached_dist_fun in \<open>auto simp add: dist_image_le_dist_fun\<close>)
+
+
+end
+
+
+subsection \<open>Completeness\<close>
+
+text \<open>Actually we can obtain even better: when the instance typ\<open>'b\<close> of
+      class\<open>decseq_restriction_space\<close> is also an instance of class\<open>complete_space\<close>,
+      the type typ\<open>'a \<Rightarrow> 'b\<close> is an instance of class\<open>complete_space\<close>.\<close>
+
+text \<open>This is because when typ\<open>'b\<close> is an instance of class\<open>decseq_restriction_space\<close>
+      (and not only class\<open>non_decseq_restriction_space\<close>) the distance between two functions
+      is reached (see @{thm reached_dist_fun}).\<close>
+
+
+
+text \<open>The only remaining thing is to prove that completeness is preserved on higher-order.\<close>
+
+instance \<open>fun\<close> :: (type, complete_decseq_restriction_space) complete_decseq_restriction_space
+  by intro_classes (fact restriction_chain_imp_restriction_convergent)
+
+instance \<open>fun\<close> :: (type, complete_strict_decseq_restriction_space) complete_strict_decseq_restriction_space
+  by intro_classes (fact restriction_chain_imp_restriction_convergent)
+
+instance \<open>fun\<close> :: (type, complete_at_least_geometric_restriction_space) complete_at_least_geometric_restriction_space
+  by intro_classes (fact restriction_chain_imp_restriction_convergent)
+
+instance \<open>fun\<close> :: (type, complete_geometric_restriction_space) complete_geometric_restriction_space
+  by intro_classes (fact restriction_chain_imp_restriction_convergent)
+
+
+
+subsection \<open>Kind of Extensionality\<close>
+
+context fixes f :: \<open>['a :: metric_space, 'b :: type] \<Rightarrow>
+                    'c :: decseq_restriction_space\<close> begin
+
+lemma lipschitz_with_simplification:
+  \<open>lipschitz_with f \<alpha> \<longleftrightarrow> (\<forall>y. lipschitz_with (\<lambda>x. f x y) \<alpha>)\<close>
+proof (intro iffI allI)
+  fix y assume assm : \<open>lipschitz_with f \<alpha>\<close>
+  show \<open>lipschitz_with (\<lambda>x. f x y) \<alpha>\<close>
+  proof (rule lipschitz_withI)
+    from assm[THEN lipschitz_withD1] show \<open>0 \<le> \<alpha>\<close> .
+  next
+    show \<open>dist (f x1 y) (f x2 y) \<le> \<alpha> * dist x1 x2\<close> for x1 x2
+      by (rule order_trans[OF _ assm[THEN lipschitz_withD2]])
+        (simp add: dist_image_le_dist_fun)
+  qed
+next
+  assume assm : \<open>\<forall>y. lipschitz_with (\<lambda>x. f x y) \<alpha>\<close>
+  show \<open>lipschitz_with f \<alpha>\<close>
+  proof (rule lipschitz_withI)
+    from assm[rule_format, THEN lipschitz_withD1] show \<open>0 \<le> \<alpha>\<close> .
+  next
+    fix x1 x2
+    obtain y where \<open>dist (f x1) (f x2) = dist (f x1 y) (f x2 y)\<close>
+      by (meson reached_dist_fun)
+    also have \<open>\<dots> \<le> \<alpha> * dist x1 x2\<close> by (rule assm[rule_format, THEN lipschitz_withD2])
+    finally show \<open>dist (f x1) (f x2) \<le> \<alpha> * dist x1 x2\<close> .
+  qed
+qed
+
+
+lemma non_expanding_simplification :
+  \<open>non_expanding f \<longleftrightarrow> (\<forall>y. non_expanding (\<lambda>x. f x y))\<close>
+  by (metis lipschitz_with_simplification non_expanding_on_def)
+
+
+lemma contraction_with_simplification:
+  \<open>contraction_with f \<alpha> \<longleftrightarrow> (\<forall>y. contraction_with (\<lambda>x. f x y) \<alpha>)\<close>
+  by (metis contraction_with_on_def lipschitz_with_simplification)
+
+
+end
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Prod_Ultrametric_Restriction_Spaces.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Prod_Ultrametric_Restriction_Spaces.thy
@@ -0,0 +1,457 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Product\<close>
+
+(*<*)
+theory Prod_Ultrametric_Restriction_Spaces
+  imports Ultrametric_Restriction_Spaces Restriction_Spaces
+begin
+  (*>*)
+
+text \<open>
+The product type typ\<open>'a \<times> 'b\<close> of to metric spaces is already instantiated as
+a metric space by setting @{thm dist_prod_def[no_vars]}.
+Unfortunately, this definition is not compatible with the distance required
+by the class\<open>non_decseq_restriction_space\<close>..
+We first have to define a new product type with a trivial theory_text\<open>typedef\<close>.
+\<close>
+
+subsection \<open>Isomorphic Product Construction\<close>
+
+subsubsection \<open>Definition and First Properties\<close>
+
+typedef ('a, 'b) prodmax (\<open>(_ \<times>max/ _)\<close> [21, 20] 20) = \<open>UNIV :: ('a \<times> 'b) set\<close>
+  morphisms from_prodmax to_prodmax by simp
+
+\<comment>\<open>Simplifications because the theory_text\<open>typedef\<close> is trivial.\<close>
+
+declare from_prodmax_inject  [simp]
+  from_prodmax_inverse [simp]
+
+lemmas to_prodmax_inject_simplified [simp] = to_prodmax_inject [simplified]
+  and to_prodmax_inverse_simplified[simp] = to_prodmax_inverse[simplified]
+
+(* useful ? *)
+lemmas to_prodmax_induct_simplified = to_prodmax_induct[simplified]
+  and to_prodmax_cases_simplified  = to_prodmax_cases [simplified]
+  and from_prodmax_induct_simplified = from_prodmax_induct[simplified]
+  and from_prodmax_cases_simplified  = from_prodmax_cases [simplified]
+
+
+setup_lifting type_definition_prodmax
+
+lift_definition Pairmax :: \<open>'a \<Rightarrow> 'b \<Rightarrow> 'a \<times>max 'b\<close> is Pair .
+
+
+free_constructors case_prodmax for Pairmax fstmax sndmax
+  by (metis Pairmax.abs_eq from_prodmax_inverse surjective_pairing)
+    (metis Pairmax.rep_eq prod.inject)
+
+
+lemma fstmax_def : \<open>fstmax \<equiv> map_fun from_prodmax id fst\<close>
+  by (intro eq_reflection ext, simp add: fstmax_def,
+      metis Pairmax.rep_eq from_prodmax_inverse fstmax_def prod.collapse prodmax.sel(1))
+lemma fstmax_rep_eq : \<open>fstmax x = fst (from_prodmax x)\<close>
+  by (metis Pairmax.rep_eq fst_conv prodmax.collapse)
+
+lemma fstmax_abs_eq [simp] : \<open>fstmax (to_prodmax y) = fst y\<close>
+  by (metis Pairmax.abs_eq prod.exhaust_sel prodmax.sel(1))
+
+lemma fstmax_transfer [transfer_rule]: \<open>rel_fun (pcr_prodmax (=) (=)) (=) fst fstmax\<close>
+  by (metis (mono_tags) Pairmax.rep_eq cr_prodmax_def fst_conv prodmax.collapse prodmax.pcr_cr_eq rel_funI)
+
+
+lemma sndmax_def : \<open>sndmax \<equiv> map_fun from_prodmax id snd\<close>
+  by (intro eq_reflection ext, simp add: sndmax_def,
+      metis Pairmax.rep_eq from_prodmax_inverse prod.collapse prodmax.case)
+
+lemma sndmax_rep_eq : \<open>sndmax x = snd (from_prodmax x)\<close>
+  by (metis Pairmax.rep_eq prodmax.collapse snd_conv)
+
+lemma sndmax_abs_eq [simp] : \<open>sndmax (to_prodmax y) = snd y\<close>
+  by (metis Pairmax.abs_eq prod.exhaust_sel prodmax.sel(2))
+
+lemma sndmax_transfer [transfer_rule] : \<open>rel_fun (pcr_prodmax (=) (=)) (=) snd sndmax\<close>
+  by (metis (mono_tags, lifting) Pairmax.rep_eq cr_prodmax_def
+      prodmax.collapse prodmax.pcr_cr_eq rel_funI snd_conv)
+
+
+lemma case_prodmax_def : \<open>case_prodmax \<equiv> map_fun id (map_fun from_prodmax id) case_prod\<close>
+  by (intro eq_reflection ext, simp add: prodmax.case_eq_if fstmax_rep_eq sndmax_rep_eq split_beta)
+
+lemma case_prodmax_rep_eq : \<open>case_prodmax f p = (case from_prodmax p of (x, y) \<Rightarrow> f x y)\<close>
+  by (simp add: fstmax_rep_eq prodmax.case_eq_if sndmax_rep_eq split_beta)
+
+lemma case_prodmax_abs_eq [simp] : \<open>case_prodmax f (to_prodmax q) = (case q of (x, y) \<Rightarrow> f x y)\<close>
+  by (simp add: prodmax.case_eq_if split_beta)
+
+lemma case_prodmax_transfer [transfer_rule] : \<open>rel_fun (=) (rel_fun (pcr_prodmax (=) (=)) (=)) case_prod case_prodmax\<close>
+  by (simp add: cr_prodmax_def fstmax_rep_eq prodmax.case_eq_if prodmax.pcr_cr_eq rel_fun_def sndmax_rep_eq split_beta)
+
+
+
+subsection \<open>Syntactic Sugar\<close>
+
+text \<open>The following syntactic sugar is of course recovered from the theory theory\<open>HOL.Product_Type\<close>.\<close>
+
+nonterminal tuple_argsmax and patternsmax
+syntax
+  "_tuplemax"      :: "'a \<Rightarrow> tuple_argsmax \<Rightarrow> 'a \<times>max 'b"        (\<open>(1\<lblot>_,/ _\<rblot>)\<close>)
+  "_tuple_argmax"  :: "'a \<Rightarrow> tuple_argsmax"                   (\<open>_\<close>)
+  "_tuple_argsmax" :: "'a \<Rightarrow> tuple_argsmax \<Rightarrow> tuple_argsmax"   (\<open>_,/ _\<close>)
+  "_patternmax"    :: "pttrn \<Rightarrow> patternsmax \<Rightarrow> pttrn"         (\<open>\<lblot>_,/ _\<rblot>\<close>)
+  ""              :: "pttrn \<Rightarrow> patternsmax"                  (\<open>_\<close>)
+  "_patternsmax"   :: "pttrn \<Rightarrow> patternsmax \<Rightarrow> patternsmax"    (\<open>_,/ _\<close>)
+translations
+  "\<lblot>x, y\<rblot>" \<rightleftharpoons> "CONST Pairmax x y"
+  "_patternmax x y" \<rightleftharpoons> "CONST Pairmax x y"
+  "_patternsmax x y" \<rightleftharpoons> "CONST Pairmax x y"
+  "_tuplemax x (_tuple_argsmax y z)" \<rightleftharpoons> "_tuplemax x (_tuple_argmax (_tuplemax y z))"
+  "\<lambda>\<lblot>x, y, zs\<rblot>. b" \<rightleftharpoons> "CONST case_prodmax (\<lambda>x \<lblot>y, zs\<rblot>. b)"
+  "\<lambda>\<lblot>x, y\<rblot>. b" \<rightleftharpoons> "CONST case_prodmax (\<lambda>x y. b)"
+  "_abs (CONST Pairmax x y) t" \<rightharpoonup> "\<lambda>\<lblot>x, y\<rblot>. t"
+  \<comment> \<open>This rule accommodates tuples in \<open>case C \<dots> \<lblot>x, y\<rblot> \<dots> \<Rightarrow> \<dots>\<close>:
+     The \<open>\<lblot>x, y\<rblot>\<close> is parsed as \<open>Pairmax x y\<close> because it is \<open>logic\<close>,
+     not \<open>pttrn\<close>.\<close>
+
+text \<open>With this syntactic sugar, one can write
+term\<open>case a of \<lblot>b, c, d, e\<rblot> \<Rightarrow> \<lblot>c, d\<rblot>\<close>, term\<open>\<lambda>\<lblot>y, u\<rblot>. a\<close>,
+term\<open>\<lambda>\<lblot>a, b\<rblot>. \<lblot>a, b, c , d ,e\<rblot>\<close>, term\<open>\<lambda>\<lblot>a, b, c\<rblot>. a\<close>, \<open>\<dots>\<close>
+as for the type typ\<open>'a \<times> 'b\<close>.\<close>
+
+
+\<comment> \<open>Conversions between \<open>_tuple\<close> and \<open>_tuplemax\<close>.\<close>
+lemmas to_prodmax_tuple     [simp] = Pairmax.abs_eq[symmetric]
+  and from_prodmax_tuplemax [simp] = Pairmax.rep_eq
+
+
+
+subsection \<open>Product\<close>
+
+text \<open>We first redo the work of theory\<open>Restriction_Spaces.Restriction_Spaces_Prod\<close>.\<close>
+
+instantiation prodmax :: (restriction, restriction) restriction
+begin
+
+lift_definition restriction_prodmax :: \<open>'a \<times>max 'b \<Rightarrow> nat \<Rightarrow> 'a \<times>max 'b\<close> is \<open>(\<down>)\<close> .
+
+lemma restriction_prodmax_def' : \<open>p \<down> n = \<lblot>fstmax p \<down> n, sndmax p \<down> n\<rblot>\<close>
+  by transfer (simp add: restriction_prod_def)
+
+instance by (intro_classes, transfer, simp)
+
+end
+
+
+instance prodmax :: (restriction_space, restriction_space) restriction_space
+  by (intro_classes; transfer) (simp_all add: ex_not_restriction_related)
+
+
+
+instantiation prodmax :: (restriction_\<sigma>, restriction_\<sigma>) restriction_\<sigma>
+begin
+
+definition restriction_\<sigma>_prodmax :: \<open>('a \<times>max 'b) itself \<Rightarrow> nat \<Rightarrow> real\<close>
+  where \<open>restriction_\<sigma>_prodmax _ n \<equiv>
+         max (restriction_\<sigma> TYPE('a) n) (restriction_\<sigma> TYPE('b) n)\<close>
+
+instance by intro_classes
+end
+
+
+instantiation prodmax :: (non_decseq_restriction_space, non_decseq_restriction_space)
+  non_decseq_restriction_space
+begin
+
+definition dist_prodmax :: \<open>['a \<times>max 'b, 'a \<times>max 'b] \<Rightarrow> real\<close>
+  where \<open>dist_prodmax f g \<equiv> INF n \<in> restriction_related_set f g. restriction_\<sigma> TYPE('a \<times>max 'b) n\<close>
+
+definition uniformity_prodmax :: \<open>(('a \<times>max 'b) \<times> 'a \<times>max 'b) filter\<close>
+  where \<open>uniformity_prodmax \<equiv> INF e\<in>{0 <..}. principal {(x, y). dist x y < e}\<close>
+
+definition open_prodmax :: \<open>('a \<times>max 'b) set \<Rightarrow> bool\<close>
+  where \<open>open_prodmax U \<equiv> \<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity\<close>
+
+
+instance
+proof intro_classes
+  show \<open>restriction_\<sigma> TYPE('a \<times>max 'b) \<longlonglongrightarrow> 0\<close>
+    by (rule real_tendsto_sandwich
+        [of \<open>\<lambda>n. 0\<close> _ _ \<open>\<lambda>n. restriction_\<sigma> TYPE('a) n + restriction_\<sigma> TYPE('b) n\<close>])
+      (simp_all add: order_less_imp_le restriction_\<sigma>_prodmax_def max_def
+        restriction_\<sigma>_tendsto_zero tendsto_add_zero)
+qed (simp_all add: uniformity_prodmax_def open_prodmax_def
+    restriction_\<sigma>_prodmax_def max_def dist_prodmax_def)
+
+end
+
+instance prodmax :: (decseq_restriction_space, decseq_restriction_space) decseq_restriction_space
+proof intro_classes
+  show \<open>decseq (restriction_\<sigma> TYPE('a \<times>max 'b))\<close>
+  proof (intro decseq_SucI)
+    show \<open>restriction_\<sigma> TYPE('a \<times>max 'b) (Suc n) \<le> restriction_\<sigma> TYPE('a \<times>max 'b) n\<close> for n
+      using decseq_SucD[of \<open>restriction_\<sigma> TYPE('a)\<close> n]
+        decseq_SucD[of \<open>restriction_\<sigma> TYPE('b)\<close> n]
+      by (auto simp add: restriction_\<sigma>_prodmax_def decseq_restriction_\<sigma>)
+  qed
+qed
+
+instance prodmax :: (strict_decseq_restriction_space, strict_decseq_restriction_space) strict_decseq_restriction_space
+proof intro_classes
+  show \<open>strict_decseq (restriction_\<sigma> TYPE('a \<times>max 'b))\<close>
+  proof (intro strict_decseq_SucI)
+    show \<open>restriction_\<sigma> TYPE('a \<times>max 'b) (Suc n) < restriction_\<sigma> TYPE('a \<times>max 'b) n\<close> for n
+      using strict_decseq_SucD[of \<open>restriction_\<sigma> TYPE('a)\<close> n]
+        strict_decseq_SucD[of \<open>restriction_\<sigma> TYPE('b)\<close> n]
+      by (auto simp add: restriction_\<sigma>_prodmax_def strict_decseq_restriction_\<sigma>)
+  qed
+qed
+
+
+instantiation prodmax :: (restriction_\<delta>, restriction_\<delta>) restriction_\<delta>
+begin
+
+definition restriction_\<delta>_prodmax :: \<open>('a \<times>max 'b) itself \<Rightarrow> real\<close>
+  where \<open>restriction_\<delta>_prodmax _ \<equiv> max (restriction_\<delta> TYPE('a)) (restriction_\<delta> TYPE('b))\<close>
+
+instance by intro_classes (simp_all add: restriction_\<delta>_prodmax_def max_def)
+
+end
+
+instance prodmax :: (at_least_geometric_restriction_space, at_least_geometric_restriction_space)
+  at_least_geometric_restriction_space
+proof intro_classes
+  show \<open>0 < restriction_\<sigma> TYPE('a \<times>max 'b) n\<close> for n by simp
+next
+  show \<open>restriction_\<sigma> TYPE('a \<times>max 'b) (Suc n)
+        \<le> restriction_\<delta> TYPE('a \<times>max 'b) * restriction_\<sigma> TYPE('a \<times>max 'b) n\<close> for n
+    by (auto intro: order_trans[OF restriction_\<sigma>_le]
+        simp add: restriction_\<delta>_prodmax_def mult_mono' restriction_\<sigma>_prodmax_def)
+next
+  show \<open>dist p1 p2 = Inf (restriction_\<sigma>_related_set p1 p2)\<close> for p1 p2 :: \<open>'a \<times>max 'b\<close>
+    by (simp add: dist_prodmax_def)
+qed
+
+
+instance prodmax :: (geometric_restriction_space, geometric_restriction_space) geometric_restriction_space
+proof intro_classes
+  show \<open>restriction_\<sigma> TYPE('a \<times>max 'b) n = restriction_\<delta> TYPE('a \<times>max 'b) ^ n\<close> for n
+    by (simp add: restriction_\<sigma>_prodmax_def restriction_\<sigma>_is restriction_\<delta>_prodmax_def max_def)
+      (meson nle_le power_mono zero_le_restriction_\<delta>)
+next
+  show \<open>dist p1 p2 = Inf (restriction_\<sigma>_related_set p1 p2)\<close> for p1 p2 :: \<open>'a \<times>max 'b\<close>
+    by (simp add: dist_prodmax_def)
+qed
+
+
+
+lemma max_dist_projections_le_dist_prodmax :
+  \<open>max (dist (fstmax p1) (fstmax p2)) (dist (sndmax p1) (sndmax p2)) \<le> dist p1 p2\<close>
+proof (unfold dist_restriction_is max_def, split if_split, intro conjI impI)
+  show \<open>Inf (restriction_\<sigma>_related_set (sndmax p1) (sndmax p2)) \<le> Inf (restriction_\<sigma>_related_set p1 p2)\<close>
+  proof (rule cINF_superset_mono[OF nonempty_restriction_related_set])
+    show \<open>bdd_below (restriction_\<sigma>_related_set (sndmax p1) (sndmax p2))\<close>
+      by (meson bdd_belowI2 zero_le_restriction_\<sigma>)
+  qed (simp_all add: subset_iff add: restriction_prodmax_def' restriction_\<sigma>_prodmax_def)
+next
+  show \<open>Inf (restriction_\<sigma>_related_set (fstmax p1) (fstmax p2)) \<le> Inf (restriction_\<sigma>_related_set p1 p2)\<close>
+  proof (rule cINF_superset_mono[OF nonempty_restriction_related_set])
+    show \<open>bdd_below (restriction_\<sigma>_related_set (fstmax p1) (fstmax p2))\<close>
+      by (meson bdd_belowI2 zero_le_restriction_\<sigma>)
+  qed (simp_all add: subset_iff add: restriction_prodmax_def' restriction_\<sigma>_prodmax_def)
+qed
+
+
+subsection \<open>Completeness\<close>
+
+subsubsection \<open>Preliminaries\<close>
+
+default_sort non_decseq_restriction_space \<comment>\<open>Otherwise we should always specify.\<close>
+
+lemma restriction_\<sigma>_prodmax_commute :
+  \<open>restriction_\<sigma> TYPE('b \<times>max 'a) = restriction_\<sigma> TYPE('a \<times>max 'b)\<close>
+  unfolding restriction_\<sigma>_prodmax_def by (rule ext) simp
+
+definition dist_left_prodmax :: \<open>('a \<times>max 'b) itself \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> real\<close>
+  where \<open>dist_left_prodmax _ x y \<equiv> INF n \<in> restriction_related_set x y. restriction_\<sigma> TYPE('a \<times>max 'b) n\<close>
+
+definition dist_right_prodmax :: \<open>('a \<times>max 'b) itself \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> real\<close>
+  where \<open>dist_right_prodmax _ x y \<equiv> INF n \<in> restriction_related_set x y. restriction_\<sigma> TYPE('a \<times>max 'b) n\<close>
+
+lemma dist_right_prodmax_is_dist_left_prodmax :
+  \<open>dist_right_prodmax TYPE('b \<times>max 'a) = dist_left_prodmax TYPE('a \<times>max 'b)\<close>
+  unfolding dist_left_prodmax_def dist_right_prodmax_def
+  by (subst restriction_\<sigma>_prodmax_commute) simp
+
+
+lemma dist_le_dist_left_prodmax : \<open>dist x y \<le> dist_left_prodmax TYPE('a \<times>max 'b) x y\<close>
+proof (unfold dist_left_prodmax_def dist_restriction_is,
+    rule cINF_mono[OF nonempty_restriction_related_set[of x y]])
+  show \<open>bdd_below (restriction_\<sigma>_related_set x y)\<close>
+  by (meson bdd_belowI2 zero_le_restriction_\<sigma>)
+next
+  show \<open>m \<in> restriction_related_set x y \<Longrightarrow>
+        \<exists>n\<in>restriction_related_set x y. \<sigma>\<down> TYPE('a) n \<le> \<sigma>\<down> TYPE('a \<times>max 'b) m\<close> for m
+    by (metis max.cobounded1 restriction_\<sigma>_prodmax_def)
+qed
+
+lemma dist_le_dist_right_prodmax : \<open>dist x y \<le> dist_right_prodmax TYPE('b \<times>max 'a) x y\<close>
+  by (simp add: dist_le_dist_left_prodmax dist_right_prodmax_is_dist_left_prodmax)
+
+
+
+lemma 
+  fixes p1 p2 :: \<open>'a :: decseq_restriction_space \<times>max 'b :: decseq_restriction_space\<close>
+  shows dist_prodmax_le_max_dist_left_prodmax_dist_right_prodmax : 
+    \<open>dist p1 p2 \<le> max (dist_left_prodmax  TYPE('a \<times>max 'b) (fstmax p1) (fstmax p2))
+                           (dist_right_prodmax TYPE('a \<times>max 'b) (sndmax p1) (sndmax p2))\<close>
+proof -
+  interpret left  : DecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(=)\<close> \<open>UNIV\<close>
+    \<open>restriction_\<sigma> TYPE('a \<times>max 'b)\<close> \<open>dist_left_prodmax TYPE('a \<times>max 'b)\<close>
+    by unfold_locales
+      (simp_all add: restriction_\<sigma>_tendsto_zero dist_left_prodmax_def decseq_restriction_\<sigma>)
+
+  interpret right : DecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(=)\<close> \<open>UNIV :: 'b set\<close>
+    \<open>restriction_\<sigma> TYPE('a \<times>max 'b)\<close> \<open>dist_right_prodmax TYPE('a \<times>max 'b)\<close>
+    by unfold_locales
+      (simp_all add: restriction_\<sigma>_tendsto_zero dist_right_prodmax_def decseq_restriction_\<sigma>)
+
+  show \<open>dist p1 p2 \<le> max (dist_left_prodmax  TYPE('a \<times>max 'b) (fstmax p1) (fstmax p2))
+                          (dist_right_prodmax TYPE('a \<times>max 'b) (sndmax p1) (sndmax p2))\<close>
+    by (auto simp add: dist_restriction_is_bis left.dist_restriction_is_bis
+        right.dist_restriction_is_bis prodmax.expand restriction_prodmax_def')
+      (smt (verit, best) Collect_cong nle_le restriction_related_le)
+qed
+
+
+default_sort type \<comment>\<open>Back to normal.\<close>
+
+
+
+subsubsection \<open>Complete Restriction Space\<close>
+
+text \<open>When the instances typ\<open>'a\<close> and typ\<open>'b\<close> of class\<open>decseq_restriction_space\<close>
+      are also instances of class\<open>complete_space\<close>,
+      the type typ\<open>'a \<times>max 'b\<close> is an instance of class\<open>complete_space\<close>.\<close>
+
+lemma restriction_chain_prodmax_iff :
+  \<open>restriction_chain \<sigma> \<longleftrightarrow> restriction_chain (\<lambda>n. fstmax (\<sigma> n)) \<and>
+                              restriction_chain (\<lambda>n. sndmax (\<sigma> n))\<close>
+  by (simp add: restriction_chain_def, transfer)
+    (metis fst_conv prod.collapse restriction_prod_def snd_conv)
+
+lemma restriction_tendsto_prodmax_iff :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<lambda>n. fstmax (\<sigma> n)) \<midarrow>\<down>\<rightarrow> fstmax \<Sigma> \<and> (\<lambda>n. sndmax (\<sigma> n)) \<midarrow>\<down>\<rightarrow> sndmax \<Sigma>\<close>
+  by (simp add: restriction_tendsto_def, transfer, simp add: restriction_prod_def)
+    (meson nle_le order.trans)
+
+lemma restriction_convergent_prodmax_iff :
+  \<open>restriction_convergent \<sigma> \<longleftrightarrow> restriction_convergent (\<lambda>n. fstmax (\<sigma> n)) \<and>
+                                restriction_convergent (\<lambda>n. sndmax (\<sigma> n))\<close>
+  by (simp add: restriction_convergent_def restriction_tendsto_prodmax_iff)
+    (metis prodmax.sel)
+
+
+instance prodmax :: (complete_decseq_restriction_space, complete_decseq_restriction_space)
+  complete_decseq_restriction_space
+proof (intro_classes, transfer)
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a \<times>max 'b\<close> assume \<open>chain\<down> \<sigma>\<close>
+  hence \<open>chain\<down> (\<lambda>n. fstmax (\<sigma> n))\<close> \<open>chain\<down> (\<lambda>n. sndmax (\<sigma> n))\<close>
+    by (simp_all add: restriction_chain_prodmax_iff)
+  hence \<open>convergent\<down> (\<lambda>n. fstmax (\<sigma> n))\<close> \<open>convergent\<down> (\<lambda>n. sndmax (\<sigma> n))\<close>
+    by (simp_all add: restriction_chain_imp_restriction_convergent)
+  thus \<open>convergent\<down> \<sigma>\<close>
+    by (simp add: restriction_convergent_prodmax_iff)
+qed    
+
+
+
+instance prodmax :: (complete_strict_decseq_restriction_space, complete_strict_decseq_restriction_space)
+  complete_strict_decseq_restriction_space
+  by intro_classes (simp add: restriction_chain_imp_restriction_convergent)
+
+instance prodmax :: (complete_at_least_geometric_restriction_space, complete_at_least_geometric_restriction_space)
+  complete_at_least_geometric_restriction_space
+  by intro_classes (simp add: restriction_chain_imp_restriction_convergent)
+
+instance prodmax :: (complete_geometric_restriction_space, complete_geometric_restriction_space)
+  complete_geometric_restriction_space
+  by intro_classes (simp add: restriction_chain_imp_restriction_convergent)
+
+
+text \<open>When the types typ\<open>'a\<close> and typ\<open>'b\<close> share the same restriction sequence,
+      we have the following equality.\<close>
+
+
+lemma same_restriction_\<sigma>_imp_restriction_\<sigma>_prodmax_is [simp] :
+  \<open>restriction_\<sigma> TYPE('b :: non_decseq_restriction_space) =
+   restriction_\<sigma> TYPE('a :: non_decseq_restriction_space) \<Longrightarrow>
+   restriction_\<sigma> TYPE('a \<times>max 'b) = restriction_\<sigma> TYPE('a)\<close>
+  unfolding restriction_\<sigma>_prodmax_def by simp
+
+
+lemma same_restriction_\<sigma>_imp_dist_prodmax_eq_max_dist_projections :
+  \<open>dist p1 p2 = max (dist (fstmax p1) (fstmax p2)) (dist (sndmax p1) (sndmax p2))\<close>
+  if same_restriction_\<sigma> [simp] : \<open>restriction_\<sigma> TYPE('b) = restriction_\<sigma> TYPE('a)\<close>
+  for p1 p2 :: \<open>'a :: decseq_restriction_space \<times>max 'b :: decseq_restriction_space\<close>
+proof (rule order_antisym)
+  have \<open>dist_left_prodmax TYPE('a \<times>max 'b) (fstmax p1) (fstmax p2) = dist (fstmax p1) (fstmax p2)\<close>
+    by (simp add: dist_left_prodmax_def dist_restriction_is)
+  moreover have \<open>dist_right_prodmax TYPE('a \<times>max 'b) (sndmax p1) (sndmax p2) = dist (sndmax p1) (sndmax p2)\<close>
+    by (simp add: dist_right_prodmax_def dist_restriction_is)
+  ultimately show \<open>dist p1 p2 \<le> max (dist (fstmax p1) (fstmax p2)) (dist (sndmax p1) (sndmax p2))\<close>
+    by (metis dist_prodmax_le_max_dist_left_prodmax_dist_right_prodmax)
+next
+  show \<open>max (dist (fstmax p1) (fstmax p2)) (dist (sndmax p1) (sndmax p2)) \<le> dist p1 p2\<close>
+    by (fact max_dist_projections_le_dist_prodmax)
+qed
+
+
+
+text \<open>Now, one can write things like term\<open>\<upsilon> \<lblot>x, y\<rblot> . f \<lblot>x, y\<rblot>\<close>.\<close>
+
+
+text \<open>We could of course imagine more support for typ\<open>'a \<times>max 'b\<close> type,
+      but as restriction spaces are intended to be used without recourse
+      to metric spaces, we have not undertaken this task for the time being.\<close>
+
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/ROOT
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/ROOT
@@ -0,0 +1,50 @@
+chapter AFP
+
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+session "Restriction_Spaces-Ultrametric" = Elementary_Ultrametric_Spaces +
+  description "Providing restriction spaces with an ultrametric structure."
+  options [timeout = 300]
+  sessions Restriction_Spaces
+  theories
+    Banach_Theorem_Extension
+    Ultrametric_Restriction_Spaces_Locale
+    Ultrametric_Restriction_Spaces
+    Fun_Ultrametric_Restriction_Spaces
+    Prod_Ultrametric_Restriction_Spaces
+    "Restriction_Spaces-Ultrametric" (global)
+  document_files "root.tex"
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Restriction_Spaces-Ultrametric.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Restriction_Spaces-Ultrametric.thy
@@ -0,0 +1,52 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Main entry Point\<close>
+
+(*<*)
+theory "Restriction_Spaces-Ultrametric"
+  imports Fun_Ultrametric_Restriction_Spaces Prod_Ultrametric_Restriction_Spaces
+begin
+  (*<*)
+
+
+text \<open>This is the entry point session\<open>Restriction_Spaces-Ultrametric\<close> should be imported from.\<close>
+
+(*<*)
+end
+  (*>*)
+
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces.thy
@@ -0,0 +1,947 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Ultrametric Structure of restriction Spaces\<close>
+
+(*<*)
+theory Ultrametric_Restriction_Spaces
+  imports Ultrametric_Restriction_Spaces_Locale
+    Restriction_Spaces Banach_Theorem_Extension
+begin
+  (*>*)
+
+
+
+text \<open>This has only be proven with the sort constraint,
+      not inside the context of the class class\<open>metric_space\<close> ...\<close>
+
+context metric_space begin
+
+lemma LIMSEQ_def : \<open>X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (X n) L < r)\<close>
+  unfolding tendsto_iff eventually_sequentially ..
+
+lemma LIMSEQ_iff_nz: \<open>X \<longlonglongrightarrow> L \<longleftrightarrow> (\<forall>r>0. \<exists>no>0. \<forall>n\<ge>no. dist (X n) L < r)\<close>
+  by (meson Suc_leD LIMSEQ_def zero_less_Suc)
+
+lemma metric_LIMSEQ_I: \<open>(\<And>r. 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r) \<Longrightarrow> X \<longlonglongrightarrow> L\<close>
+  by (simp add: LIMSEQ_def)
+
+lemma metric_LIMSEQ_D: \<open>X \<longlonglongrightarrow> L \<Longrightarrow> 0 < r \<Longrightarrow> \<exists>no. \<forall>n\<ge>no. dist (X n) L < r\<close>
+  by (simp add: LIMSEQ_def)
+
+
+
+lemma LIMSEQ_dist_iff:
+  \<open>f \<longlonglongrightarrow> l \<longleftrightarrow> (\<lambda>x. dist (f x) l) \<longlonglongrightarrow> 0\<close>
+proof (unfold LIMSEQ_def, rule iffI)
+  show \<open>(\<lambda>n. dist (f n) l) \<longlonglongrightarrow> 0 \<Longrightarrow> (\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (f n) l < r)\<close>
+    by (metis (mono_tags, lifting) eventually_at_top_linorder order_tendstoD(2))
+next
+  show \<open>\<forall>r>0. \<exists>no. \<forall>n\<ge>no. dist (f n) l < r \<Longrightarrow> (\<lambda>n. dist (f n) l) \<longlonglongrightarrow> 0\<close>
+    by (simp add: metric_space_class.LIMSEQ_def)
+qed
+
+
+lemma Cauchy_converges_subseq:
+  fixes u::"nat \<Rightarrow> 'a"
+  assumes "Cauchy u"
+    "strict_mono r"
+    "(u \<circ> r) \<longlonglongrightarrow> l"
+  shows "u \<longlonglongrightarrow> l"
+proof -
+  have *: "eventually (\<lambda>n. dist (u n) l < e) sequentially" if "e > 0" for e
+  proof -
+    have "e/2 > 0" using that by auto
+    then obtain N1 where N1: "\<And>m n. m \<ge> N1 \<Longrightarrow> n \<ge> N1 \<Longrightarrow> dist (u m) (u n) < e/2"
+      using \<open>Cauchy u\<close> unfolding Cauchy_def by blast
+    obtain N2 where N2: "\<And>n. n \<ge> N2 \<Longrightarrow> dist ((u \<circ> r) n) l < e / 2"
+      using order_tendstoD(2)[OF iffD1[OF LIMSEQ_dist_iff \<open>(u \<circ> r) \<longlonglongrightarrow> l\<close>] \<open>e/2 > 0\<close>]
+      unfolding eventually_sequentially by auto
+    have "dist (u n) l < e" if "n \<ge> max N1 N2" for n
+    proof -
+      have "dist (u n) l \<le> dist (u n) ((u \<circ> r) n) + dist ((u \<circ> r) n) l"
+        by (rule dist_triangle)
+      also have "\<dots> < e/2 + e/2"
+      proof (intro add_strict_mono)
+        show "dist (u n) ((u \<circ> r) n) < e / 2"
+          using N1[of n "r n"] N2[of n] that unfolding comp_def
+          by (meson assms(2) le_trans max.bounded_iff strict_mono_imp_increasing)
+        show "dist ((u \<circ> r) n) l < e / 2"
+          using N2 that by auto
+      qed
+      finally show ?thesis by simp
+    qed 
+    then show ?thesis unfolding eventually_sequentially by blast
+  qed
+  have "(\<lambda>n. dist (u n) l) \<longlonglongrightarrow> 0"
+    by (simp add: less_le_trans * order_tendstoI)
+  then show ?thesis using LIMSEQ_dist_iff by auto
+qed
+
+
+end
+
+
+
+
+subsection \<open>The Construction with Classes\<close>
+
+class restriction_\<sigma> = restriction_space +
+  fixes restriction_\<sigma> :: \<open>'a itself \<Rightarrow> nat \<Rightarrow> real\<close> (\<open>\<sigma>\<down>\<close>)
+
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, NONE)\<close>
+  \<comment> \<open>To be able to use const\<open>dist\<close> out of the class\<open>metric_space\<close> class.\<close>
+
+
+class non_decseq_restriction_space =
+  uniformity_dist + open_uniformity + restriction_\<sigma> +
+  \<comment>\<open>We do not assume the restriction sequence to be const\<open>decseq\<close> yet.\<close>
+  assumes restriction_\<sigma>_tendsto_zero' : \<open>\<sigma>\<down> TYPE('a) \<longlonglongrightarrow> 0\<close> 
+    and zero_less_restriction_\<sigma>' [simp] : \<open>0 < \<sigma>\<down> TYPE('a) n\<close>
+    (* and decseq_restriction_\<sigma> : \<open>decseq (\<sigma>\<down> TYPE('a))\<close> *)
+    and dist_restriction_is' : \<open>dist x y = (INF n \<in> {n. x \<down> n = y \<down> n}. \<sigma>\<down> TYPE('a) n)\<close>
+begin 
+
+sublocale NonDecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(=)\<close> \<open>UNIV\<close> \<open>\<sigma>\<down> TYPE ('a)\<close> dist
+  by unfold_locales (simp_all add: restriction_\<sigma>_tendsto_zero' dist_restriction_is')
+
+end
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, SOME typ\<open>'a :: metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
+  \<comment> \<open>Only allow const\<open>dist\<close> in class class\<open>metric_space\<close> (back to normal).\<close>
+
+text \<open>We hide duplicated facts
+@{thm restriction_\<sigma>_tendsto_zero zero_less_restriction_\<sigma> dist_restriction_is}.\<close>
+hide_fact restriction_\<sigma>_tendsto_zero' zero_less_restriction_\<sigma>' dist_restriction_is'
+
+
+context non_decseq_restriction_space begin
+
+subclass ultrametric_space
+proof unfold_locales
+  show \<open>dist x y = 0 \<longleftrightarrow> x = y\<close> for x y
+    by (simp add: "restriction_dist_eq_0_iff_related")
+
+  have dist_commute : \<open>dist x y = dist y x\<close> for x y
+    by (simp add: dist_restriction_is) metis
+
+  show \<open>dist x y \<le> max (dist x z) (dist y z)\<close> for x y z
+  proof -
+    consider \<open>x \<noteq> y\<close> and \<open>y \<noteq> z\<close> and \<open>x \<noteq> z\<close> | \<open>x = y \<or> y = z \<or> x = z\<close> by blast
+
+    thus \<open>dist x y \<le> max (dist x z) (dist y z)\<close>
+    proof cases
+      assume \<open>x \<noteq> y\<close> and \<open>y \<noteq> z\<close> and \<open>x \<noteq> z\<close>
+      from this(1)[THEN not_related_imp_dist_restriction_le_some_restriction_\<sigma>]
+        this(2, 3)[THEN not_related_imp_dist_restriction_is_some_restriction_\<sigma>]
+      obtain l m n
+        where   * : \<open>dist x y \<le> \<sigma>\<down> TYPE('a) l\<close> \<open>x \<down> Suc l \<noteq> y \<down> Suc l\<close>
+          and  ** : \<open>dist y z = \<sigma>\<down> TYPE('a) m\<close> \<open>\<forall>k\<le>m. y \<down> k = z \<down> k\<close>
+          and *** : \<open>dist x z = \<sigma>\<down> TYPE('a) n\<close> \<open>\<forall>k\<le>n. x \<down> k = z \<down> k\<close> by blast
+      have \<open>min n m \<le> l\<close>
+      proof (rule ccontr)
+        assume \<open>\<not> min n m \<le> l\<close>
+        hence \<open>Suc l \<le> n\<close> and \<open>Suc l \<le> m\<close> by simp_all
+        with "**"(2) "***"(2) Suc_le_lessD
+        have \<open>x \<down> Suc l = z \<down> Suc l\<close> \<open>y \<down> Suc l = z \<down> Suc l\<close> by simp_all
+        hence \<open>x \<down> Suc l = y \<down> Suc l\<close> by simp
+        with \<open>x \<down> Suc l \<noteq> y \<down> Suc l\<close> show False ..
+      qed
+
+      have \<open>dist x y \<le> \<sigma>\<down> TYPE('a) (min n m)\<close>
+        by (rule restriction_space_Inf_properties(1))
+          (simp add: "**"(2) "***"(2))
+      also have \<open>\<dots> \<le> max (dist x z) (dist y z)\<close>
+        unfolding "**"(1) "***"(1) by linarith
+      finally show \<open>dist x y \<le> max (dist x z) (dist y z)\<close> .
+    next
+      show \<open>x = y \<or> y = z \<or> x = z \<Longrightarrow> dist x y \<le> max (dist x z) (dist y z)\<close>
+        by (elim disjE, simp_all add: dist_commute)
+          (metis not_related_imp_dist_restriction_is_some_restriction_\<sigma>
+            restriction_dist_eq_0_iff_related zero_le_restriction_\<sigma> dual_order.refl)
+    qed
+  qed
+qed
+
+end
+
+context non_decseq_restriction_space begin
+
+
+lemma restriction_tendsto_self: \<open>(\<lambda>n. x \<down> n) \<longlonglongrightarrow> x\<close>
+proof -
+  have \<open>(\<lambda>n. dist (x \<down> n) x) \<longlonglongrightarrow> 0\<close>
+  proof (rule real_tendsto_sandwich[OF _ ])
+    show \<open>\<forall>F n in sequentially. 0 \<le> dist (x \<down> n) x\<close> by simp
+  next
+    show \<open>\<forall>F n in sequentially. dist (x \<down> n) x \<le> \<sigma>\<down> TYPE('a) n\<close>
+      by (auto intro: eventually_sequentiallyI[OF restriction_space_Inf_properties(1)])
+  next
+    show \<open>(\<lambda>n. 0) \<longlonglongrightarrow> 0\<close> by simp
+  next
+    show \<open>\<sigma>\<down> TYPE('a) \<longlonglongrightarrow> 0\<close> by (simp add: restriction_\<sigma>_tendsto_zero)
+  qed
+  thus \<open>(\<lambda>n. x \<down> n) \<longlonglongrightarrow> x\<close>
+    by (subst tendsto_iff[of \<open>(\<lambda>n. x \<down> n)\<close>], subst eventually_sequentially)
+      (simp add: LIMSEQ_iff)
+qed
+
+end
+
+
+
+
+
+class decseq_restriction_space = non_decseq_restriction_space +
+  assumes decseq_restriction_\<sigma>' : \<open>decseq (\<sigma>\<down> TYPE('a))\<close>
+begin
+
+sublocale DecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(=)\<close> \<open>UNIV :: 'a set\<close>
+  \<open>restriction_\<sigma> TYPE('a)\<close> dist
+  by unfold_locales (simp add: decseq_restriction_\<sigma>')
+
+
+\<comment> \<open>Removing term\<open>x \<in> M\<close> and term\<open>y \<in> M\<close>.\<close>
+lemmas dist_restriction_is_bis_simplified = dist_restriction_is_bis[simplified]
+  and not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified =
+  not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq[simplified]
+
+end
+
+text \<open>We hide duplicated fact @{thm decseq_restriction_\<sigma>}.\<close>
+hide_fact decseq_restriction_\<sigma>'
+
+
+
+class strict_decseq_restriction_space = non_decseq_restriction_space +
+  assumes strict_decseq_restriction_\<sigma> : \<open>strict_decseq (\<sigma>\<down> TYPE('a))\<close>
+begin
+
+subclass decseq_restriction_space
+  by unfold_locales
+    (fact strict_decseq_imp_decseq[OF strict_decseq_restriction_\<sigma>])
+
+end
+
+
+
+
+text \<open>Generic Properties\<close>
+
+
+lemma (in metric_space) dist_sequences_tendsto_zero_imp_tendsto_iff :
+  \<open>(\<lambda>n. dist (\<sigma> n) (\<psi> n)) \<longlonglongrightarrow> 0 \<Longrightarrow> \<sigma> \<longlonglongrightarrow> \<Sigma> \<longleftrightarrow> \<psi> \<longlonglongrightarrow> \<Sigma>\<close>
+proof (rule iffI)
+  show \<open>\<psi> \<longlonglongrightarrow> \<Sigma>\<close> if \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close> and \<open>(\<lambda>n. dist (\<sigma> n) (\<psi> n)) \<longlonglongrightarrow> 0\<close> for \<sigma> \<psi>
+  proof -
+    from that have * : \<open>(\<lambda>n. dist (\<sigma> n) (\<psi> n) + dist (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+      unfolding LIMSEQ_dist_iff using tendsto_add_zero by blast
+    have \<open>(\<lambda>n. dist (\<psi> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+      by (rule real_tendsto_sandwich
+          [of \<open>\<lambda>n. 0\<close> \<open>\<lambda>n. dist (\<psi> n) \<Sigma>\<close> _ \<open>\<lambda>n. dist (\<sigma> n) (\<psi> n) + dist (\<sigma> n) \<Sigma>\<close>])
+        (simp_all add: "*" dist_triangle3)
+    with LIMSEQ_dist_iff show \<open>\<psi> \<longlonglongrightarrow> \<Sigma>\<close> by blast
+  qed
+  thus \<open>(\<lambda>n. dist (\<sigma> n) (\<psi> n)) \<longlonglongrightarrow> 0 \<Longrightarrow> \<psi> \<longlonglongrightarrow> \<Sigma> \<Longrightarrow> \<sigma> \<longlonglongrightarrow> \<Sigma>\<close>
+    by (simp add: dist_commute)
+qed
+
+
+lemma (in non_decseq_restriction_space) restricted_sequence_tendsto_iff :
+  \<open>(\<lambda>n. \<sigma> n \<down> n) \<longlonglongrightarrow> \<Sigma> \<longleftrightarrow> \<sigma> \<longlonglongrightarrow> \<Sigma>\<close>
+proof -
+  have \<open>(\<lambda>n. dist (\<sigma> n \<down> n) (\<sigma> n)) \<longlonglongrightarrow> 0\<close>
+  proof (unfold metric_space_class.LIMSEQ_def, intro allI impI)
+    fix \<epsilon> :: real assume \<open>0 < \<epsilon>\<close>
+    from restriction_\<sigma>_tendsto_zero[unfolded metric_space_class.LIMSEQ_def, rule_format, OF \<open>0 < \<epsilon>\<close>]
+    obtain no where \<open>\<forall>n\<ge>no. \<sigma>\<down> TYPE('a) n < \<epsilon>\<close> 
+      by (auto simp add: dist_real_def)
+    have \<open>no \<le> n \<Longrightarrow> dist (dist (\<sigma> n \<down> n) (\<sigma> n)) 0 < \<epsilon>\<close> for n 
+      by (simp, rule le_less_trans[OF restriction_space_Inf_properties(1)[of \<open>\<sigma>\<down> TYPE('a) n\<close>]])
+        (simp, simp add: \<open>\<forall>n\<ge>no. \<sigma>\<down> TYPE('a) n < \<epsilon>\<close>)
+    thus \<open>\<exists>no. \<forall>n\<ge>no. dist (dist (\<sigma> n \<down> n) (\<sigma> n)) 0 < \<epsilon>\<close> by blast
+  qed
+
+  thus \<open>(\<lambda>n. \<sigma> n \<down> n) \<longlonglongrightarrow> \<Sigma> \<longleftrightarrow> \<sigma> \<longlonglongrightarrow> \<Sigma>\<close> 
+    by (simp add: dist_sequences_tendsto_zero_imp_tendsto_iff)
+qed
+
+
+
+lemma (in non_decseq_restriction_space) Cauchy_restriction_chain :
+  \<open>Cauchy \<sigma>\<close> if \<open>chain\<down> \<sigma>\<close>
+proof (rule metric_CauchyI)
+  fix \<epsilon> :: real assume \<open>0 < \<epsilon>\<close>
+  from LIMSEQ_D[OF restriction_\<sigma>_tendsto_zero \<open>0 < \<epsilon>\<close>, simplified]
+  obtain M where \<open>\<sigma>\<down> TYPE('a) M < \<epsilon>\<close> by blast
+  moreover have \<open>M \<le> m \<Longrightarrow> M \<le> n \<Longrightarrow> dist (\<sigma> m) (\<sigma> n) \<le> \<sigma>\<down> TYPE('a) M\<close> for m n
+    by (rule restriction_space_Inf_properties(1), simp add: image_iff)
+      (metis restriction_chain_def_ter \<open>chain\<down> \<sigma>\<close>)
+  ultimately have \<open>\<forall>m\<ge>M. \<forall>n\<ge>M. dist (\<sigma> m) (\<sigma> n) < \<epsilon>\<close>
+    by (meson dual_order.strict_trans2)
+  thus \<open>\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (\<sigma> m) (\<sigma> n) < \<epsilon>\<close> ..
+qed
+
+
+
+
+lemma (in non_decseq_restriction_space) restriction_tendsto_imp_tendsto :
+  \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close> if \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof (rule metric_LIMSEQ_I)
+  fix \<epsilon> :: real assume \<open>0 < \<epsilon>\<close>
+  with restriction_\<sigma>_tendsto_zero[THEN LIMSEQ_D]
+  obtain n0 where \<open>\<forall>k\<ge>n0. \<sigma>\<down> TYPE('a) k < \<epsilon>\<close> by auto
+  hence \<open>\<sigma>\<down> TYPE('a) n0 < \<epsilon>\<close> by simp
+  from restriction_tendstoD[OF \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>]
+  obtain n1 where \<open>\<forall>k\<ge>n1. \<Sigma> \<down> n0 = \<sigma> k \<down> n0\<close> ..
+  hence \<open>\<forall>k\<ge>n1. dist (\<sigma> k) \<Sigma> \<le> \<sigma>\<down> TYPE('a) n0\<close>
+    by (simp add: restriction_space_Inf_properties(1))
+  with \<open>\<sigma>\<down> TYPE('a) n0 < \<epsilon>\<close>
+  have \<open>\<forall>k\<ge>n1. dist (\<sigma> k) \<Sigma> < \<epsilon>\<close> by (meson order_le_less_trans)
+  thus \<open>\<exists>n1. \<forall>k\<ge>n1. dist (\<sigma> k) \<Sigma> < \<epsilon>\<close> ..
+qed
+
+
+
+
+
+
+text \<open>In Decseq Restriction Space\<close>
+
+context decseq_restriction_space begin
+
+lemma le_dist_to_restriction_eqE : 
+  obtains k where \<open>n \<le> k\<close> \<open>\<And>x y :: 'a. dist x y \<le> \<sigma>\<down> TYPE('a) k \<Longrightarrow> x \<down> n = y \<down> n\<close>
+proof -
+  have \<open>\<exists>k\<ge>n. \<forall>x y :: 'a. dist x y \<le> \<sigma>\<down> TYPE('a) k \<longrightarrow> x \<down> n = y \<down> n\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> (\<exists>k\<ge>n. \<forall>x y :: 'a. dist x y \<le> \<sigma>\<down> TYPE('a) k \<longrightarrow> x \<down> n = y \<down> n)\<close>
+    hence \<open>\<forall>k\<ge>n. \<exists>x y :: 'a. dist x y \<le> \<sigma>\<down> TYPE('a) k \<and> x \<down> n \<noteq> y \<down> n\<close> by simp
+    then obtain X Y :: \<open>nat \<Rightarrow> 'a\<close>
+      where * : \<open>\<forall>k\<ge>n. dist (X k) (Y k) \<le> \<sigma>\<down> TYPE('a) k\<close>
+        \<open>\<forall>k\<ge>n. X k \<down> n \<noteq> Y k \<down> n\<close> by metis
+    moreover obtain n0 where \<open>\<forall>k\<ge>n0. \<sigma>\<down> TYPE('a) k < \<sigma>\<down> TYPE('a) n\<close>
+      by (metis LIMSEQ_D zero_less_restriction_\<sigma> abs_of_nonneg diff_zero
+          restriction_\<sigma>_tendsto_zero real_norm_def zero_le_restriction_\<sigma>)
+    ultimately have \<open>\<forall>k\<ge>n+n0. dist (X k) (Y k) < \<sigma>\<down> TYPE('a) n\<close>
+      by (metis dual_order.strict_trans2 add_leE)
+    moreover from "*"(2) have \<open>\<forall>k\<ge>n. \<sigma>\<down> TYPE('a) n \<le> dist (X k) (Y k)\<close>
+      by (simp add: dist_restriction_is_bis)
+        (metis (full_types) decseq_restriction_\<sigma>[THEN decseqD] linorder_linear
+          not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(2))
+    ultimately show False by (metis le_add1 linorder_not_le order_refl)
+  qed
+  thus \<open>(\<And>k. \<lbrakk>n \<le> k; \<And>x y :: 'a. dist x y \<le> \<sigma>\<down> TYPE('a) k \<Longrightarrow> x \<down> n = y \<down> n\<rbrakk>
+        \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close> by blast
+qed
+
+
+
+theorem tendsto_iff_restriction_tendsto : \<open>\<sigma> \<longlonglongrightarrow> \<Sigma> \<longleftrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof (rule iffI)
+  show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<sigma> \<longlonglongrightarrow> \<Sigma>\<close> by (fact restriction_tendsto_imp_tendsto)
+next
+  show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> if \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    fix n
+    obtain n0 where \<open>dist x y \<le> \<sigma>\<down> TYPE('a) n0 \<Longrightarrow> x \<down> n = y \<down> n\<close> for x y :: 'a
+      by (metis le_dist_to_restriction_eqE)
+    moreover from metric_LIMSEQ_D[OF \<open>\<sigma> \<longlonglongrightarrow> \<Sigma>\<close> zero_less_restriction_\<sigma>]
+    obtain n1 where \<open>\<forall>k\<ge>n1. dist (\<sigma> k) \<Sigma> < \<sigma>\<down> TYPE('a) n0\<close> ..
+    ultimately have \<open>\<forall>k\<ge>n1. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> by (metis dual_order.order_iff_strict)
+    thus \<open>\<exists>n1. \<forall>k\<ge>n1. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> ..
+  qed
+qed
+
+corollary convergent_iff_restriction_convergent : \<open>convergent \<sigma> \<longleftrightarrow> convergent\<down> \<sigma>\<close>
+  by (simp add: convergent_def restriction_convergent_def tendsto_iff_restriction_tendsto)
+
+
+
+theorem complete_iff_restriction_complete :
+  \<open>(\<forall>\<sigma>. Cauchy \<sigma> \<longrightarrow> convergent \<sigma>) \<longleftrightarrow> (\<forall>\<sigma>. chain\<down> \<sigma> \<longrightarrow> convergent\<down> \<sigma>)\<close>
+  \<comment> \<open>The following result shows that we have not lost anything with our
+      definitions of convergence, completeness, etc. specific to restriction spaces.\<close>
+proof (intro iffI impI allI)
+  fix \<sigma> assume hyp : \<open>\<forall>\<sigma>. Cauchy \<sigma> \<longrightarrow> convergent \<sigma>\<close> and \<open>chain\<down> \<sigma>\<close>
+  from Cauchy_restriction_chain \<open>chain\<down> \<sigma>\<close> have \<open>Cauchy \<sigma>\<close> by blast
+  hence \<open>convergent \<sigma>\<close> by (simp add: hyp)
+  thus \<open>convergent\<down> \<sigma>\<close> by (simp add: convergent_iff_restriction_convergent)
+next
+  fix \<sigma> assume hyp : \<open>\<forall>\<sigma>. chain\<down> \<sigma> \<longrightarrow> convergent\<down> \<sigma>\<close> and \<open>Cauchy \<sigma>\<close>
+  from \<open>Cauchy \<sigma>\<close> have * : \<open>\<forall>n. \<exists>k. \<forall>l\<ge>k. \<sigma> l \<down> n = \<sigma> k \<down> n\<close>
+    by (metis (mono_tags, opaque_lifting) Cauchy_altdef2 dual_order.order_iff_strict
+        le_dist_to_restriction_eqE zero_less_restriction_\<sigma>) 
+  define f where \<open>f \<equiv> rec_nat
+                      (LEAST k. \<forall>l\<ge>k. \<sigma> l \<down> 0 = \<sigma> k \<down> 0)
+                      (\<lambda>n k. LEAST l. k < l \<and> (\<forall>m\<ge>l. \<sigma> m \<down> Suc n = \<sigma> l \<down> Suc n))\<close>
+  have f_Suc_def : \<open>f (Suc n) = (LEAST l. f n < l \<and> (\<forall>m\<ge>l. \<sigma> m \<down> Suc n = \<sigma> l \<down> Suc n))\<close>
+    (is \<open>f (Suc n) = Least (?f_Suc n)\<close>) for n by (simp add: f_def)
+  from "*" have ** : \<open>\<exists>k>f n. \<forall>m\<ge>k. \<sigma> m \<down> Suc n = \<sigma> k \<down> Suc n\<close> for n
+    by (metis dual_order.trans lessI linorder_not_le order_le_less)
+  have \<open>strict_mono f\<close>
+  proof (unfold strict_mono_Suc_iff, rule allI)
+    show \<open>f n < f (Suc n)\<close> for n
+      by (fact LeastI_ex[of \<open>?f_Suc n\<close>, folded f_Suc_def, OF "**", THEN conjunct1])
+  qed
+
+  have \<open>chain\<down> (\<lambda>n. (\<sigma> \<circ> f) n \<down> n)\<close> \<comment> \<open>key point\<close>
+  proof (rule restriction_chainI)
+    fix n
+    have \<open>\<sigma> (f (Suc n)) \<down> n = \<sigma> (f n) \<down> n\<close>
+    proof (cases n)
+      show \<open>n = 0 \<Longrightarrow> \<sigma> (f (Suc n)) \<down> n = \<sigma> (f n) \<down> n\<close> by simp
+    next
+      show \<open>n = Suc nat \<Longrightarrow> \<sigma> (f (Suc n)) \<down> n = \<sigma> (f n) \<down> n\<close> for nat
+      proof (clarify, intro LeastI_ex[of \<open>?f_Suc nat\<close>,
+            folded f_Suc_def, THEN conjunct2, rule_format, OF "**"])
+        show \<open>n = Suc nat \<Longrightarrow> f (Suc nat) \<le> f (Suc (Suc nat))\<close>
+          by (simp add: \<open>strict_mono f\<close> strict_mono_less_eq)
+      qed
+    qed
+    thus \<open>(\<sigma> \<circ> f) (Suc n) \<down> Suc n \<down> n = (\<sigma> \<circ> f) n \<down> n\<close> by simp
+  qed
+  with hyp have \<open>convergent\<down> (\<lambda>n. (\<sigma> \<circ> f) n \<down> n)\<close> by simp
+  hence \<open>convergent\<down> (\<lambda>n. (\<sigma> \<circ> f) n)\<close>
+    by (simp add: restriction_convergent_restricted_iff_restriction_convergent)
+  hence \<open>convergent (\<lambda>n. (\<sigma> \<circ> f) n)\<close>
+    by (simp add: convergent_iff_restriction_convergent)
+  with Cauchy_converges_subseq[OF \<open>Cauchy \<sigma>\<close> \<open>strict_mono f\<close>]
+  show \<open>convergent \<sigma>\<close> unfolding convergent_def by blast
+qed
+
+end
+
+
+text \<open>The following classes will be useful later.\<close>
+
+
+class complete_decseq_restriction_space = decseq_restriction_space +
+  assumes restriction_chain_imp_restriction_convergent' : \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close>
+begin
+
+subclass complete_restriction_space
+  by unfold_locales (fact restriction_chain_imp_restriction_convergent')
+
+subclass complete_ultrametric_space
+proof (unfold_locales)
+  from complete_iff_restriction_complete restriction_chain_imp_restriction_convergent
+  show \<open>Cauchy \<sigma> \<Longrightarrow> convergent \<sigma>\<close> for \<sigma> by blast
+qed
+
+end
+
+text \<open>We hide duplicated fact @{thm restriction_chain_imp_restriction_convergent}.\<close>
+hide_fact restriction_chain_imp_restriction_convergent'
+
+
+class complete_strict_decseq_restriction_space = strict_decseq_restriction_space + 
+  assumes restriction_chain_imp_restriction_convergent' : \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close>
+begin
+
+subclass complete_decseq_restriction_space
+  by (unfold_locales) (fact restriction_chain_imp_restriction_convergent')
+
+end
+
+text \<open>We hide duplicated fact @{thm restriction_chain_imp_restriction_convergent'}.\<close>
+hide_fact restriction_chain_imp_restriction_convergent'
+
+
+
+class restriction_\<delta> = restriction_\<sigma> +
+  fixes restriction_\<delta> :: \<open>'a itself \<Rightarrow> real\<close> (\<open>\<delta>\<down>\<close>)
+  assumes zero_less_restriction_\<delta> [simp] : \<open>0 < \<delta>\<down> TYPE('a)\<close>
+    and restriction_\<delta>_less_one  [simp] : \<open>\<delta>\<down> TYPE('a) < 1\<close>
+begin
+
+lemma zero_le_restriction_\<delta> [simp] : \<open>0 \<le> \<delta>\<down> TYPE('a)\<close>
+  and restriction_\<delta>_le_one  [simp] : \<open>\<delta>\<down> TYPE('a) \<le> 1\<close>
+  and zero_le_pow_restriction_\<delta> [simp] : \<open>0 \<le> \<delta>\<down> TYPE('a) ^ n\<close>
+  by (simp_all add: order_less_imp_le)
+
+lemma pow_restriction_\<delta>_le_one [simp] : \<open>\<delta>\<down> TYPE('a) ^ n \<le> 1\<close>
+  by (simp add: power_le_one)
+
+lemma pow_restriction_\<delta>_less_one [simp] : \<open>n \<noteq> 0 \<Longrightarrow> \<delta>\<down> TYPE('a) ^ n < 1\<close>
+  by (metis restriction_\<delta>_less_one zero_less_restriction_\<delta> not_gr_zero power_0 power_strict_decreasing)
+
+end
+
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, NONE)\<close>
+  \<comment> \<open>To be able to use const\<open>dist\<close> out of the class\<open>metric_space\<close> class.\<close>
+
+class at_least_geometric_restriction_space =
+  uniformity_dist + open_uniformity + restriction_\<delta> +
+  assumes zero_less_restriction_\<sigma>' : \<open>0 < \<sigma>\<down> TYPE('a) n\<close>
+    and restriction_\<sigma>_le :
+    \<open>\<sigma>\<down> TYPE('a) (Suc n) \<le> \<delta>\<down> TYPE('a) * \<sigma>\<down> TYPE('a) n\<close>
+    and dist_restriction_is' :
+    \<open>dist x y = (INF n \<in> {n. x \<down> n = y \<down> n}. \<sigma>\<down> TYPE('a) n)\<close>
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, SOME typ\<open>'a :: metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
+  \<comment> \<open>Only allow const\<open>dist\<close> in class class\<open>metric_space\<close> (back to normal).\<close>
+
+context at_least_geometric_restriction_space begin
+
+lemma restriction_\<sigma>_le_restriction_\<sigma>_times_pow_restriction_\<delta> :
+  \<open>\<sigma>\<down> TYPE('a) (n + k) \<le> \<sigma>\<down> TYPE('a) n * \<delta>\<down> TYPE('a) ^ k\<close>
+  by (induct k, simp_all)
+    (metis dual_order.trans restriction_\<sigma>_le zero_less_restriction_\<delta>
+      mult.left_commute mult_le_cancel_left_pos)
+
+
+lemma restriction_\<sigma>_le_pow_restriction_\<delta> :
+  \<open>\<sigma>\<down> TYPE('a) n \<le> \<sigma>\<down> TYPE('a) 0 * \<delta>\<down> TYPE('a) ^ n\<close>
+  by (metis add_0 restriction_\<sigma>_le_restriction_\<sigma>_times_pow_restriction_\<delta>)
+
+
+
+
+subclass strict_decseq_restriction_space
+proof unfold_locales
+  have \<open>\<forall>F n in sequentially. 0 \<le> \<sigma>\<down> TYPE('a) n\<close>
+    by (simp add: zero_less_restriction_\<sigma>' order_less_imp_le)
+  moreover have \<open>\<forall>F n in sequentially. \<sigma>\<down> TYPE('a) n
+                 \<le> \<sigma>\<down> TYPE('a) 0 * \<delta>\<down> TYPE('a) ^ n\<close>
+    by (simp add: restriction_\<sigma>_le_pow_restriction_\<delta>)
+  moreover have \<open>(\<lambda>n. 0) \<longlonglongrightarrow> 0\<close> by simp
+  moreover have \<open>(\<lambda>n. \<sigma>\<down> TYPE('a) 0 * \<delta>\<down> TYPE('a) ^ n) \<longlonglongrightarrow> 0\<close>
+    by (simp add: LIMSEQ_abs_realpow_zero2 abs_of_pos tendsto_mult_right_zero)
+  ultimately show \<open>\<sigma>\<down> TYPE('a) \<longlonglongrightarrow> 0\<close> by (fact real_tendsto_sandwich)
+next
+  show \<open>0 < \<sigma>\<down> TYPE('a) n\<close> for n
+    by (fact zero_less_restriction_\<sigma>')
+next
+  show \<open>dist x y = Inf (\<sigma>\<down> TYPE('a) ` {n. x \<down> n = y \<down> n})\<close> for x y
+    by (fact dist_restriction_is')
+next
+  show \<open>strict_decseq (\<sigma>\<down> TYPE('a))\<close>
+    by (rule strict_decseqI, rule le_less_trans[OF restriction_\<sigma>_le])
+      (simp add: zero_less_restriction_\<sigma>')
+qed
+
+lemma \<open>0 < \<delta>\<down> TYPE('a) ^ n\<close> by simp
+
+
+
+end
+
+text \<open>We hide duplicated facts
+@{thm zero_less_restriction_\<sigma> dist_restriction_is}.\<close>
+hide_fact zero_less_restriction_\<sigma>' dist_restriction_is'
+
+
+
+
+class complete_at_least_geometric_restriction_space = at_least_geometric_restriction_space + 
+  assumes restriction_chain_imp_restriction_convergent' : \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close>
+begin
+
+subclass complete_strict_decseq_restriction_space
+  by unfold_locales (fact restriction_chain_imp_restriction_convergent')
+
+end
+
+text \<open>We hide duplicated fact @{thm restriction_chain_imp_restriction_convergent}.\<close>
+hide_fact restriction_chain_imp_restriction_convergent'
+
+
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, NONE)\<close>
+  \<comment> \<open>To be able to use const\<open>dist\<close> out of the class\<open>metric_space\<close> class.\<close>
+
+class geometric_restriction_space = uniformity_dist + open_uniformity + restriction_\<delta> +
+  assumes restriction_\<sigma>_is : \<open>\<sigma>\<down> TYPE('a) n = \<delta>\<down> TYPE('a) ^ n\<close>
+    and dist_restriction_is' : \<open>dist x y = (INF n \<in> {n. x \<down> n = y \<down> n}. \<sigma>\<down> TYPE('a) n)\<close>
+begin
+
+text \<open>This is what ``restriction space'' usually mean in the literature.
+      The previous classes are generalizations of this concept
+      (even this one is a generalization, since we usually
+      have term\<open>\<delta>\<down> TYPE('a) = 1 / 2\<close>).\<close>
+
+subclass at_least_geometric_restriction_space
+proof unfold_locales
+  show \<open>0 < \<sigma>\<down> TYPE('a) n\<close> for n by (simp add: restriction_\<sigma>_is)
+next
+  show \<open>\<sigma>\<down> TYPE('a) (Suc n) \<le> \<delta>\<down> TYPE('a) * \<sigma>\<down> TYPE('a) n\<close> for n
+    by (simp add: restriction_\<sigma>_is)
+next
+  show \<open>dist x y = Inf (\<sigma>\<down> TYPE('a) ` {n. x \<down> n = y \<down> n})\<close> for x y
+    by (fact dist_restriction_is')
+qed
+
+
+lemma \<open>0 < \<delta>\<down> TYPE('a) ^ n\<close> by simp
+
+end
+
+setup \<open>Sign.add_const_constraint (const_name\<open>dist\<close>, SOME typ\<open>'a :: metric_space \<Rightarrow> 'a \<Rightarrow> real\<close>)\<close>
+  \<comment> \<open>Only allow const\<open>dist\<close> in class class\<open>metric_space\<close> (back to normal).\<close>
+
+
+text \<open>We hide duplicated fact @{thm dist_restriction_is}.\<close>
+hide_fact dist_restriction_is'
+
+
+class complete_geometric_restriction_space = geometric_restriction_space + 
+  assumes restriction_chain_imp_restriction_convergent' : \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close>
+begin 
+
+subclass complete_at_least_geometric_restriction_space
+  by (unfold_locales) (fact restriction_chain_imp_restriction_convergent')
+
+end 
+
+text \<open>We hide duplicated fact @{thm restriction_chain_imp_restriction_convergent}.\<close>
+hide_fact restriction_chain_imp_restriction_convergent'
+
+(* remove that, useless now *)
+theorem geometric_restriction_space_completeI : \<open>convergent \<sigma>\<close>
+  if \<open>\<And>\<sigma> :: nat \<Rightarrow> 'a. restriction_chain \<sigma> \<Longrightarrow> \<exists>\<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n\<close>
+    and \<open>Cauchy \<sigma>\<close> for \<sigma> :: \<open>nat \<Rightarrow> 'a :: geometric_restriction_space\<close>
+  by (metis complete_iff_restriction_complete convergent_def
+      convergent_iff_restriction_convergent ext restriction_tendsto_self that)
+
+
+
+\<comment> \<open>Because const\<open>cball\<close> is not defined in class\<open>metric_space\<close>.\<close>
+lemma (in non_decseq_restriction_space) restriction_cball_subset_cball :
+  \<open>\<sigma>\<down> TYPE('a) n \<le> r \<Longrightarrow> \<B>\<down>(\<Sigma>, n) \<subseteq> {x. dist \<Sigma> x \<le> r}\<close>
+  by (simp add: subset_iff restriction_cball_mem_iff)
+    (simp add: dual_order.trans restriction_space_Inf_properties(1))
+
+corollary restriction_cball_subset_cball_bis :
+  \<open>\<sigma>\<down> TYPE('a) n \<le> r \<Longrightarrow> \<B>\<down>(\<Sigma>, n) \<subseteq> cball \<Sigma> r\<close>
+  for \<Sigma> :: \<open>'a :: non_decseq_restriction_space\<close>
+  unfolding cball_def by (fact restriction_cball_subset_cball)
+
+
+lemma (in non_decseq_restriction_space) restriction_ball_subset_ball :
+  \<open>\<sigma>\<down> TYPE('a) n < r \<Longrightarrow> restriction_ball \<Sigma> n \<subseteq> {x. dist \<Sigma> x < r}\<close>
+  by (simp add: subset_iff restriction_cball_mem_iff)
+    (metis (mono_tags, lifting) image_eqI mem_Collect_eq order_le_less_trans
+      restriction_related_pred restriction_space_Inf_properties(1))
+
+corollary restriction_ball_subset_ball_bis :
+  \<open>\<sigma>\<down> TYPE('a) n < r \<Longrightarrow> restriction_ball \<Sigma> n \<subseteq> ball \<Sigma> r\<close>
+  for \<Sigma> :: \<open>'a :: non_decseq_restriction_space\<close>
+  unfolding ball_def by (fact restriction_ball_subset_ball)
+
+
+
+lemma (in strict_decseq_restriction_space)
+  restriction_cball_is_cball : \<open>\<B>\<down>(\<Sigma>, n) = {x. dist \<Sigma> x \<le> \<sigma>\<down> TYPE('a) n}\<close>
+proof (intro subset_antisym subsetI)
+  from restriction_cball_subset_cball
+  show \<open>x \<in> \<B>\<down>(\<Sigma>, n) \<Longrightarrow> x \<in> {x. dist \<Sigma> x \<le> \<sigma>\<down> TYPE('a) n}\<close> for x by blast
+next
+  show \<open>x \<in> \<B>\<down>(\<Sigma>, n)\<close> if \<open>x \<in> {x. dist \<Sigma> x \<le> \<sigma>\<down> TYPE('a) n}\<close> for x
+  proof (cases \<open>\<Sigma> = x\<close>)
+    show \<open>\<Sigma> = x \<Longrightarrow> x \<in> \<B>\<down>(\<Sigma>, n)\<close> by simp
+  next
+    assume \<open>\<Sigma> \<noteq> x\<close>
+    with not_related_imp_dist_restriction_is_some_restriction_\<sigma> obtain m
+      where \<open>dist \<Sigma> x = \<sigma>\<down> TYPE('a) m\<close> \<open>\<forall>k\<le>m. \<Sigma> \<down> k = x \<down> k\<close> by blast
+    from \<open>x \<in> {x. dist \<Sigma> x \<le> \<sigma>\<down> TYPE('a) n}\<close>
+    have \<open>dist \<Sigma> x \<le> \<sigma>\<down> TYPE('a) n\<close> by simp
+    with \<open>dist \<Sigma> x = \<sigma>\<down> TYPE('a) m\<close> have \<open>n \<le> m\<close>
+      by (metis linorder_not_less strict_decseq_restriction_\<sigma> strict_decseq_def_bis)
+    with \<open>\<forall>k\<le>m. \<Sigma> \<down> k = x \<down> k\<close> have \<open>\<Sigma> \<down> n = x \<down> n\<close> by simp
+    thus \<open>x \<in> \<B>\<down>(\<Sigma>, n)\<close> by (simp add: restriction_cball_mem_iff)
+  qed
+qed
+
+lemma restriction_cball_is_cball_bis : \<open>\<B>\<down>(\<Sigma>, n) = cball \<Sigma> (\<sigma>\<down> TYPE('a) n)\<close>
+  for \<Sigma> :: \<open>'a :: strict_decseq_restriction_space\<close>
+  by (simp add: cball_def restriction_cball_is_cball)
+
+
+lemma (in strict_decseq_restriction_space)
+  restriction_ball_is_ball : \<open>restriction_ball \<Sigma> n = {x. dist \<Sigma> x < \<sigma>\<down> TYPE('a) n}\<close>
+proof (intro subset_antisym subsetI)
+  show \<open>x \<in> restriction_ball \<Sigma> n \<Longrightarrow> x \<in> {x. dist \<Sigma> x < \<sigma>\<down> TYPE('a) n}\<close> for x
+    by (simp add: subset_iff)
+      (metis lessI local.restriction_cball_is_cball strict_decseq_restriction_\<sigma>
+        mem_Collect_eq order_le_less_trans strict_decseq_def_bis)
+next
+  show \<open>x \<in> restriction_ball \<Sigma> n\<close> if \<open>x \<in> {x. dist \<Sigma> x < \<sigma>\<down> TYPE('a) n}\<close> for x
+  proof (cases \<open>\<Sigma> = x\<close>)
+    show \<open>\<Sigma> = x \<Longrightarrow> x \<in> restriction_ball \<Sigma> n\<close> by simp
+  next
+    assume \<open>\<Sigma> \<noteq> x\<close>
+    with not_related_imp_dist_restriction_is_some_restriction_\<sigma> obtain m
+      where \<open>dist \<Sigma> x = \<sigma>\<down> TYPE('a) m\<close> \<open>\<forall>k\<le>m. \<Sigma> \<down> k = x \<down> k\<close> by blast
+    from \<open>x \<in> {x. dist \<Sigma> x < \<sigma>\<down> TYPE('a) n}\<close>
+    have \<open>dist \<Sigma> x < \<sigma>\<down> TYPE('a) n\<close> by simp
+    with \<open>dist \<Sigma> x = \<sigma>\<down> TYPE('a) m\<close> have \<open>n < m\<close>
+      using strict_decseq_restriction_\<sigma> strict_decseq_def_bis by auto
+    with \<open>\<forall>k\<le>m. \<Sigma> \<down> k = x \<down> k\<close> have \<open>\<Sigma> \<down> Suc n = x \<down> Suc n\<close> by simp
+    thus \<open>x \<in> restriction_ball \<Sigma> n\<close> by (simp add: restriction_cball_mem_iff)
+  qed
+qed
+
+lemma restriction_ball_is_ball_bis : \<open>restriction_ball \<Sigma> n = ball \<Sigma> (\<sigma>\<down> TYPE('a) n)\<close>
+  for \<Sigma> :: \<open>'a :: strict_decseq_restriction_space\<close>
+  by (simp add: ball_def restriction_ball_is_ball)
+
+
+
+
+
+lemma isCont_iff_restriction_cont_at : \<open>isCont f \<Sigma> \<longleftrightarrow> restriction_cont_at f \<Sigma>\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: decseq_restriction_space\<close>
+  by (unfold restriction_cont_at_def continuous_at_sequentially comp_def,
+      fold tendsto_iff_restriction_tendsto) simp
+
+
+lemma (in strict_decseq_restriction_space)
+  open_iff_restriction_open : \<open>open U \<longleftrightarrow> open\<down> U\<close>
+proof (unfold open_dist restriction_open_iff_restriction_cball_characterization, intro iffI ballI)
+  fix \<Sigma> assume \<open>\<forall>x\<in>U. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> U\<close> and \<open>\<Sigma> \<in> U\<close>
+  then obtain e where \<open>0 < e\<close> \<open>\<forall>y. dist y \<Sigma> < e \<longrightarrow> y \<in> U\<close> by blast
+  from \<open>0 < e\<close> obtain n where \<open>\<sigma>\<down> TYPE('a) n < e\<close>
+    by (metis eventually_at_top_linorder le_refl restriction_\<sigma>_tendsto_zero order_tendstoD(2))
+  with \<open>\<forall>y. dist y \<Sigma> < e \<longrightarrow> y \<in> U\<close> dist_commute restriction_cball_is_cball
+  have \<open>\<B>\<down>(\<Sigma>, n) \<subseteq> U\<close> by auto
+  thus \<open>\<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close> ..
+next
+  fix x assume \<open>\<forall>\<Sigma>\<in>U. \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close> \<open>x \<in> U\<close>
+  then obtain n where \<open>\<B>\<down>(x, n) \<subseteq> U\<close> by blast
+  hence \<open>\<forall>y. dist y x < \<sigma>\<down> TYPE('a) n \<longrightarrow> y \<in> U\<close>
+    by (simp add: dist_commute restriction_cball_is_cball subset_iff)
+  thus \<open>\<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> U\<close>
+    using zero_less_restriction_\<sigma>' by blast
+qed
+
+
+lemma (in strict_decseq_restriction_space)
+  closed_iff_restriction_closed : \<open>closed U \<longleftrightarrow> closed\<down> U\<close>
+  by (simp add: closed_open open_iff_restriction_open restriction_open_Compl_iff)
+
+
+
+
+lemma continuous_on_iff_restriction_cont_on :
+  \<open>open U \<Longrightarrow> continuous_on U f \<longleftrightarrow> restriction_cont_on f U\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: decseq_restriction_space\<close>
+  by (simp add: restriction_cont_on_def continuous_on_eq_continuous_at
+      flip: isCont_iff_restriction_cont_at)
+
+
+
+
+
+subsection \<open>Equivalence between Lipschitz Map and Restriction shift Map\<close>
+
+text \<open>For a function @{term [show_types] \<open>f :: 'a \<Rightarrow> 'b\<close>}, it is equivalent to have
+      term\<open>restriction_shift_on f k A\<close> and
+      term\<open>lipschitz_with_on f (restriction_\<delta> TYPE('b :: geometric_restriction_space) ^ k) A\<close>
+      when typ\<open>'a\<close> is of sort class\<open>geometric_restriction_space\<close> and 
+      term\<open>restriction_\<sigma> TYPE('b :: geometric_restriction_space) =
+           restriction_\<sigma> TYPE('a :: geometric_restriction_space)\<close>
+      (typ\<open>'b\<close> is then necessarily also of sort class\<open>geometric_restriction_space\<close>).\<close>
+
+text \<open>Weaker versions of this result wan be established with weaker assumptions on the sort,
+      this is what we do below.\<close>
+
+lemma restriction_shift_nonneg_on_imp_lipschitz_with_on :
+  \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) ^ k) A\<close> if \<open>restriction_shift_on f (int k) A\<close>
+  and le_restriction_\<sigma> : \<open>\<And>n. restriction_\<sigma> TYPE('b) n \<le> restriction_\<sigma> TYPE('a) n\<close>
+for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: at_least_geometric_restriction_space\<close>
+proof (rule lipschitz_with_onI)
+  show \<open>0 \<le> restriction_\<delta> TYPE('b) ^ k\<close> by simp
+next
+  fix x y assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<noteq> y\<close> \<open>f x \<noteq> f y\<close>
+  from \<open>restriction_shift_on f k A\<close>[THEN restriction_shift_onD, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>]
+  have \<open>i \<in> restriction_related_set x y \<Longrightarrow> i + k \<in> restriction_related_set (f x) (f y)\<close> for i
+    by (simp add: nat_int_add)
+  hence \<open>Sup (restriction_related_set x y) + k \<in> restriction_related_set (f x) (f y)\<close>
+    using \<open>x \<noteq> y\<close> Sup_in_restriction_related_set by blast
+  hence \<open>Sup (restriction_related_set x y) + k \<le> Sup (restriction_related_set (f x) (f y))\<close>
+    by (simp add: \<open>f x \<noteq> f y\<close> bdd_above_restriction_related_set_iff cSup_upper)
+  moreover have \<open>dist (f x) (f y) = restriction_\<sigma> TYPE('b) (Sup (restriction_related_set (f x) (f y)))\<close>
+    by (simp add: \<open>f x \<noteq> f y\<close> dist_restriction_is_bis_simplified)
+  ultimately have \<open>dist (f x) (f y) \<le> restriction_\<sigma> TYPE('b) (Sup (restriction_related_set x y) + k)\<close>
+    by (simp add: decseqD decseq_restriction_space_class.decseq_restriction_\<sigma>)
+  hence \<open>dist (f x) (f y) \<le> restriction_\<sigma> TYPE('b) (Sup (restriction_related_set x y)) * restriction_\<delta> TYPE('b) ^ k\<close>
+    by (meson order.trans restriction_\<sigma>_le_restriction_\<sigma>_times_pow_restriction_\<delta>)
+  also have \<open>\<dots> \<le> restriction_\<sigma> TYPE('a) (Sup (restriction_related_set x y)) * restriction_\<delta> TYPE('b) ^ k\<close>
+    by (simp add: le_restriction_\<sigma>)
+  finally show \<open>dist (f x) (f y) \<le> restriction_\<delta> TYPE('b) ^ k * dist x y\<close>
+    by (simp add: dist_restriction_is_bis[of x y] \<open>x \<noteq> y\<close> mult.commute)
+qed
+
+corollary restriction_shift_nonneg_imp_lipschitz_with :
+  \<open>\<lbrakk>restriction_shift f (int k); \<And>n. restriction_\<sigma> TYPE('b) n \<le> restriction_\<sigma> TYPE('a) n\<rbrakk>
+   \<Longrightarrow> lipschitz_with f (restriction_\<delta> TYPE('b) ^ k)\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: at_least_geometric_restriction_space\<close>
+  using restriction_shift_def restriction_shift_nonneg_on_imp_lipschitz_with_on by blast
+
+
+lemma lipschitz_with_on_imp_restriction_shift_neg_on :
+  \<open>restriction_shift_on f (- int k) A\<close> if \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi - int k) A\<close>
+  and le_restriction_\<sigma> : \<open>\<And>n. restriction_\<sigma> TYPE('a) n \<le> restriction_\<sigma> TYPE('b) n\<close>
+for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: at_least_geometric_restriction_space\<close>
+proof (rule restriction_shift_onI, goal_cases)
+  fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>f x \<noteq> f y\<close> \<open>x \<down> n = y \<down> n\<close>
+  from \<open>f x \<noteq> f y\<close> have \<open>x \<noteq> y\<close> by blast
+  from \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi - int k) A\<close>
+    [THEN lipschitz_with_onD2, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>]
+  have \<open>dist (f x) (f y) \<le> restriction_\<delta> TYPE('b) powi - int k * dist x y\<close> .
+  hence \<open>dist (f x) (f y) * restriction_\<delta> TYPE('b) ^ k \<le> dist x y\<close>
+    by (subst (asm) mult.commute)
+      (drule mult_imp_div_pos_le[rotated]; simp add: power_int_minus_divide)
+  hence \<open>restriction_\<sigma> TYPE('b) (Sup (restriction_related_set (f x) (f y))) * restriction_\<delta> TYPE('b) ^ k
+         \<le> restriction_\<sigma> TYPE('b) (Sup (restriction_related_set x y))\<close>
+    by (simp add: dist_restriction_is_bis \<open>x \<noteq> y\<close> \<open>f x \<noteq> f y\<close> )
+      (drule order_trans[OF _ le_restriction_\<sigma>], simp)
+  from order_trans[OF restriction_\<sigma>_le_restriction_\<sigma>_times_pow_restriction_\<delta>
+      [of \<open>Sup (restriction_related_set (f x) (f y))\<close> k] this]
+  have \<open>restriction_\<sigma> TYPE('b) (Sup (restriction_related_set (f x) (f y)) + k)
+        \<le> restriction_\<sigma> TYPE('b) (Sup (restriction_related_set x y))\<close> .
+  hence \<open>Sup (restriction_related_set x y) \<le> Sup (restriction_related_set (f x) (f y)) + k\<close>
+    using strict_decseq_def_ter strict_decseq_restriction_\<sigma> by blast
+  from order_trans[OF cSup_upper this] have \<open>n \<le> Sup (restriction_related_set (f x) (f y)) + k\<close>
+    by (simp add: \<open>x \<down> n = y \<down> n\<close> \<open>x \<noteq> y\<close> bdd_above_restriction_related_set_iff)
+  hence \<open>nat (int n + - int k) \<le> Sup (restriction_related_set (f x) (f y))\<close> by linarith
+  thus \<open>f x \<down> nat (int n + - int k) = f y \<down> nat (int n + - int k)\<close>
+    by (metis not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(2))
+qed
+
+corollary lipschitz_with_imp_restriction_shift_neg :
+  \<open>\<lbrakk>lipschitz_with f (restriction_\<delta> TYPE('b) powi - int k);
+    \<And>n. restriction_\<sigma> TYPE('a) n \<le> restriction_\<sigma> TYPE('b) n\<rbrakk>
+   \<Longrightarrow> restriction_shift f (- int k)\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: at_least_geometric_restriction_space\<close>
+  using lipschitz_with_on_imp_restriction_shift_neg_on restriction_shift_def by blast
+
+
+
+text \<open>We obtained that const\<open>restriction_shift\<close> implies const\<open>lipschitz_with\<close>
+      when term\<open>0 \<le> k\<close> and that const\<open>lipschitz_with\<close> implies
+      const\<open>restriction_shift\<close> when term\<open>k \<le> 0\<close>. To cover the remaining cases,
+      we give move from class\<open>at_least_geometric_restriction_space\<close>
+      to class\<open>geometric_restriction_space\<close>.\<close>
+
+
+theorem lipschitz_with_on_iff_restriction_shift_on :
+  \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A \<longleftrightarrow> restriction_shift_on f k A\<close>
+  if same_restriction_\<sigma> : \<open>restriction_\<sigma> TYPE('b) = restriction_\<sigma> TYPE('a)\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: geometric_restriction_space\<close>
+proof (rule iffI)
+  \<comment> \<open>We could do a case on term\<open>k\<close>, but both cases are actually handled by the proof
+      required after applying @{thm lipschitz_with_on_imp_restriction_shift_neg_on}.\<close>
+  show \<open>restriction_shift_on f k A\<close> if \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A\<close>
+  proof (rule restriction_shift_onI)
+    fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>f x \<noteq> f y\<close> \<open>x \<down> n = y \<down> n\<close>
+    from \<open>f x \<noteq> f y\<close> have \<open>x \<noteq> y\<close> by blast
+    from \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A\<close>
+      [THEN lipschitz_with_onD2, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>]
+    have \<open>dist (f x) (f y) \<le> restriction_\<delta> TYPE('b) powi k * dist x y\<close> .
+    also have \<open>dist (f x) (f y) = restriction_\<delta> TYPE('b) ^ Sup (restriction_related_set (f x) (f y))\<close>
+      by (simp add: dist_restriction_is_bis \<open>f x \<noteq> f y\<close> restriction_\<sigma>_is)
+    also have \<open>dist x y = restriction_\<delta> TYPE('b) ^ Sup (restriction_related_set x y)\<close>
+      by (simp add: dist_restriction_is_bis \<open>x \<noteq> y\<close> restriction_\<sigma>_is same_restriction_\<sigma>[symmetric])
+    finally have \<open>restriction_\<delta> TYPE('b) ^ Sup (restriction_related_set (f x) (f y))
+                  \<le> restriction_\<delta> TYPE('b) powi k * restriction_\<delta> TYPE('b) ^ Sup (restriction_related_set x y)\<close> .
+    also have \<open>\<dots> = restriction_\<delta> TYPE('b) powi (k + Sup (restriction_related_set x y))\<close>
+      by (subst power_int_add, metis order_less_irrefl zero_less_restriction_\<delta>) simp
+    finally have \<open>restriction_\<delta> TYPE('b) ^ Sup (restriction_related_set (f x) (f y)) \<le> \<dots>\<close> .
+    moreover have \<open>restriction_\<delta> TYPE('b) powi m \<le> restriction_\<delta> TYPE('b) powi m' \<Longrightarrow> m' \<le> m\<close> for m m'
+      by (rule ccontr, unfold not_le,
+          drule power_int_strict_decreasing[OF _ zero_less_restriction_\<delta> restriction_\<delta>_less_one])
+        (fold not_le, blast)
+    ultimately have \<open>k + Sup (restriction_related_set x y) \<le> Sup (restriction_related_set (f x) (f y))\<close> by simp
+    hence \<open>Sup (restriction_related_set x y) \<le> Sup (restriction_related_set (f x) (f y)) - k\<close> by simp
+    with \<open>x \<down> n = y \<down> n\<close> \<open>x \<noteq> y\<close> linorder_not_le
+      not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(3)
+    have \<open>n \<le> Sup (restriction_related_set (f x) (f y)) - k\<close> by fastforce
+    hence \<open>nat (int n + k) \<le> Sup (restriction_related_set (f x) (f y))\<close> by simp
+    thus \<open>f x \<down> nat (int n + k) = f y \<down> nat (int n + k)\<close> 
+      by (metis not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified(2))
+  qed
+next
+  show \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A\<close> if \<open>restriction_shift_on f k A\<close>
+  proof (cases k)
+    show \<open>k = int k' \<Longrightarrow> lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A\<close> for k'
+      by (simp, rule restriction_shift_nonneg_on_imp_lipschitz_with_on)
+        (use \<open>restriction_shift_on f k A\<close> same_restriction_\<sigma> in simp_all)
+  next
+    fix k' assume \<open>k = - int (Suc k')\<close>
+    show \<open>lipschitz_with_on f (restriction_\<delta> TYPE('b) powi k) A\<close>
+    proof (rule lipschitz_with_onI)
+      show \<open>0 \<le> restriction_\<delta> TYPE('b) powi k\<close> by simp
+    next
+      fix x y assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<noteq> y\<close> \<open>f x \<noteq> f y\<close>
+      have \<open>i \<in> restriction_related_set x y \<Longrightarrow> i - Suc k' \<in> restriction_related_set (f x) (f y)\<close> for i
+        using \<open>restriction_shift_on f k A\<close>[THEN restriction_shift_onD, OF \<open>x \<in> A\<close> \<open>y \<in> A\<close>, of i]
+        by (unfold \<open>k = - int (Suc k')\<close>, clarify) (metis diff_conv_add_uminus nat_minus_as_int)
+      hence \<open>Sup (restriction_related_set x y) - Suc k' \<in> restriction_related_set (f x) (f y)\<close>
+        using \<open>x \<noteq> y\<close> not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq_simplified by blast
+      hence \<open>Sup (restriction_related_set x y) - Suc k' \<le> Sup (restriction_related_set (f x) (f y))\<close>
+        by (simp add: \<open>f x \<noteq> f y\<close> bdd_above_restriction_related_set_iff cSup_upper)
+      hence * : \<open>Sup (restriction_related_set x y) + k \<le> Sup (restriction_related_set (f x) (f y))\<close>
+        by (simp add: \<open>k = - int (Suc k')\<close>)
+      have \<open>dist (f x) (f y) = restriction_\<delta> TYPE('b) powi Sup (restriction_related_set (f x) (f y))\<close>
+        by (simp add: \<open>f x \<noteq> f y\<close> dist_restriction_is_bis_simplified restriction_\<sigma>_is)
+      also have \<open>\<dots> \<le> restriction_\<delta> TYPE('b) powi (Sup (restriction_related_set x y) + k)\<close>
+        by (rule power_int_decreasing[OF "*"]; simp)
+          (metis order_less_irrefl zero_less_restriction_\<delta>)
+      also have \<open>\<dots> = restriction_\<delta> TYPE('b) powi k * restriction_\<delta> TYPE('b) powi (Sup (restriction_related_set x y))\<close>
+        by (subst power_int_add, metis order_less_irrefl zero_less_restriction_\<delta>) simp
+      also have \<open>restriction_\<delta> TYPE('b) powi (Sup (restriction_related_set x y)) = dist x y\<close>
+        by (simp add: \<open>x \<noteq> y\<close> dist_restriction_is_bis_simplified
+            restriction_\<sigma>_is same_restriction_\<sigma>[symmetric])
+      finally show \<open>dist (f x) (f y) \<le> restriction_\<delta> TYPE('b) powi k * dist x y\<close> .
+    qed
+  qed
+qed
+
+corollary lipschitz_with_iff_restriction_shift :
+  \<open>restriction_\<sigma> TYPE('b) = restriction_\<sigma> TYPE('a) \<Longrightarrow>
+   lipschitz_with f (restriction_\<delta> TYPE('b) powi k) \<longleftrightarrow> restriction_shift f k\<close>
+  for f :: \<open>'a :: decseq_restriction_space \<Rightarrow> 'b :: geometric_restriction_space\<close>
+  by (simp add: lipschitz_with_on_iff_restriction_shift_on restriction_shift_def)
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces_Locale.thy
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces_Locale.thy
@@ -0,0 +1,351 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+
+section \<open>Locales factorizing the proof Work\<close>
+
+(*<*)
+theory Ultrametric_Restriction_Spaces_Locale
+  imports Restriction_Spaces.Restriction_Spaces_Locales Elementary_Ultrametric_Spaces
+begin
+  (*>*)
+
+subsection \<open>Preliminaries on strictly decreasing Sequences\<close>
+
+abbreviation strict_decseq :: \<open>(nat \<Rightarrow> 'a :: order) \<Rightarrow> bool\<close>
+  where \<open>strict_decseq \<equiv> monotone (<) (\<lambda>x y. y < x)\<close>
+
+lemma strict_decseq_def : \<open>strict_decseq X \<longleftrightarrow> (\<forall>m n. m < n \<longrightarrow> X n < X m)\<close>
+  by (fact monotone_def)
+
+lemma strict_decseqI : \<open>strict_decseq X\<close> if \<open>\<And>n. X (Suc n) < X n\<close>
+  by (metis Suc_le_eq decseqD decseq_Suc_iff le_less_trans nless_le strict_decseq_def that)
+
+lemma strict_decseqD : \<open>strict_decseq X \<Longrightarrow> m < n \<Longrightarrow> X n < X m\<close>
+  using strict_decseq_def by blast
+
+lemma strict_decseq_def_bis : \<open>strict_decseq X \<longleftrightarrow> (\<forall>m n. X n < X m \<longleftrightarrow> m < n)\<close>
+  by (metis linorder_less_linear order_less_imp_not_less strict_decseq_def)
+
+lemma strict_decseq_def_ter : \<open>strict_decseq X \<longleftrightarrow> (\<forall>m n. X n \<le> X m \<longleftrightarrow> m \<le> n)\<close>
+  unfolding strict_decseq_def_bis
+  by (rule iffI, metis le_simps(2) order_le_less, simp add: less_le_not_le)
+
+lemma strict_decseq_imp_decseq : \<open>strict_decseq \<sigma> \<Longrightarrow> decseq \<sigma>\<close>
+  by (simp add: monotone_on_def order_le_less)
+
+
+lemma strict_decseq_SucI : \<open>(\<And>n. X (Suc n) < X n) \<Longrightarrow> strict_decseq X\<close>
+  by (metis Suc_le_eq decseqD decseq_Suc_iff less_le_not_le order_less_le strict_decseq_def)
+
+lemma strict_decseq_SucD : \<open>strict_decseq A \<Longrightarrow> A (Suc i) < A i\<close>
+  by (simp add: strict_decseq_def)
+
+
+
+text \<open>
+Classically, a restriction space is given the structure of a metric space by defining
+\<open>dist x y \<equiv> Inf {(1 / 2) ^ n| n. x \<down> n = y \<down> n}\<close>.
+This obviously also works if we replace term\<open>1 / 2 :: real\<close> by any real term\<open>\<delta>\<close>
+such that term\<open>0 < (\<delta> :: real)\<close> and term\<open>(\<delta> :: real) < 1\<close>.
+But more generally, this still works if we set
+\<open>dist x y \<equiv> Inf {\<sigma> n| n. x \<down> n = y \<down> n}\<close> where term\<open>\<sigma>\<close> is a sequence of typ\<open>real\<close>
+verifying term\<open>\<forall>n. 0 < (\<sigma> :: nat \<Rightarrow> real) n\<close> and term\<open>\<sigma> \<longlonglongrightarrow> 0\<close>.
+As you would expect, the more structure you have, the more powerful theorems you get.
+We explore all these variants in the theory below.
+\<close>
+
+subsection \<open>The Construction with Locales\<close>
+
+text \<open>Our formalization will extend the class class\<open>metric_space\<close>.
+      But some proofs are redundant, especially when it comes to the product type.
+      So first we will be working with locales, and interpret them with the classes.\<close>
+
+locale NonDecseqRestrictionSpace = PreorderRestrictionSpace +
+  \<comment>\<open>Factorization of the proof work.\<close>
+  fixes M :: \<open>'a set\<close> and restriction_\<sigma> :: \<open>nat \<Rightarrow> real\<close> (\<open>\<sigma>\<down>\<close>)
+    and restriction_dist :: \<open>'a \<Rightarrow> 'a \<Rightarrow> real\<close> (\<open>dist\<down>\<close>)
+  assumes restriction_\<sigma>_tendsto_zero : \<open>restriction_\<sigma> \<longlonglongrightarrow> 0\<close>
+    and zero_less_restriction_\<sigma> [simp] : \<open>0 < restriction_\<sigma> n\<close>
+    and dist_restriction_is :
+    \<open>dist\<down> x y = (INF n \<in> restriction_related_set x y. restriction_\<sigma> n)\<close>
+begin
+
+lemma zero_le_restriction_\<sigma> [simp] : \<open>0 \<le> \<sigma>\<down> n\<close> 
+  by (simp add: order_less_imp_le)
+
+lemma restriction_\<sigma>_neq_zero [simp] : \<open>\<sigma>\<down> n \<noteq> 0\<close> 
+  by (metis zero_less_restriction_\<sigma> order_less_irrefl)
+
+
+(* lemma eventually_inf_to:
+  \<open>r \<longlonglongrightarrow> k \<Longrightarrow> \<forall>e>0. k \<in> ball k (k + e) \<longrightarrow> (\<forall>F n in sequentially. r n \<in> ball k (k + e))\<close>
+  unfolding eventually_sequentially inj_on_def
+  using lim_explicit by blast *)
+
+
+lemma bounded_range_restriction_\<sigma>: \<open>bounded (range \<sigma>\<down>)\<close>
+  by (fact convergent_imp_bounded[OF restriction_\<sigma>_tendsto_zero])
+
+abbreviation restriction_\<sigma>_related_set :: \<open>'a \<Rightarrow> 'a \<Rightarrow> real set\<close>
+  where \<open>restriction_\<sigma>_related_set x y \<equiv> \<sigma>\<down> ` restriction_related_set x y\<close>
+
+abbreviation restriction_\<sigma>_not_related_set :: \<open>'a \<Rightarrow> 'a \<Rightarrow> real set\<close>
+  where \<open>restriction_\<sigma>_not_related_set x y \<equiv> \<sigma>\<down> ` restriction_not_related_set x y\<close>
+
+lemma nonempty_restriction_\<sigma>_related_set :
+  \<open>restriction_\<sigma>_related_set x y \<noteq> {}\<close> by simp
+
+lemma restriction_\<sigma>_related_set_Un_restriction_\<sigma>_not_related_set :
+  \<open>restriction_\<sigma>_related_set x y \<union> restriction_\<sigma>_not_related_set x y = range \<sigma>\<down>\<close>
+  by blast
+
+lemma \<open>bdd_above (restriction_\<sigma>_related_set x y)\<close>
+  by (meson bdd_above.I2 bdd_above.unfold bounded_imp_bdd_above
+      bounded_range_restriction_\<sigma> rangeI)
+
+lemma \<open>bdd_above (restriction_\<sigma>_not_related_set x y)\<close>
+  by (meson bdd_above.E bdd_above.I2 bounded_imp_bdd_above
+      bounded_range_restriction_\<sigma> range_eqI)
+
+lemma bounded_restriction_\<sigma>_related_set: \<open>bounded (restriction_\<sigma>_related_set x y)\<close>
+  by (meson bounded_range_restriction_\<sigma> bounded_subset image_mono top_greatest)
+
+lemma bounded_restriction_\<sigma>_not_related_set: \<open>bounded (restriction_\<sigma>_not_related_set x y)\<close>
+  by (meson bounded_range_restriction_\<sigma> bounded_subset image_mono subset_UNIV)
+
+corollary restriction_space_Inf_properties:
+  \<open>a \<in> restriction_\<sigma>_related_set x y \<Longrightarrow> dist\<down> x y \<le> a\<close>
+  \<open>\<lbrakk>\<And>a. a \<in> restriction_\<sigma>_related_set x y \<Longrightarrow> b \<le> a\<rbrakk> \<Longrightarrow> b \<le> dist\<down> x y\<close>
+  unfolding dist_restriction_is
+  by (simp_all add: bounded_has_Inf(1) bounded_restriction_\<sigma>_related_set cInf_greatest)
+
+lemma restriction_\<sigma>_related_set_alt :
+  \<open>restriction_\<sigma>_related_set x y = {\<sigma>\<down> n| n. n \<in> restriction_related_set x y}\<close>
+  by blast
+
+
+lemma exists_less_restriction_\<sigma> : \<open>\<exists>n. m < n \<and> \<sigma>\<down> n < \<sigma>\<down> m\<close>
+proof (rule ccontr)
+  assume \<open>\<not> (\<exists>n>m. \<sigma>\<down> n < \<sigma>\<down> m)\<close>
+  hence \<open>\<forall>n\<ge>m. \<sigma>\<down> m \<le> \<sigma>\<down> n\<close>
+    by (metis linorder_not_le nle_le)
+  hence \<open>\<not> \<sigma>\<down> \<longlonglongrightarrow> 0\<close>
+    by (meson Lim_bounded2 linorder_not_le zero_less_restriction_\<sigma>)
+  with restriction_\<sigma>_tendsto_zero show False by simp
+qed
+
+
+lemma \<open>dist\<down> x y = Inf (restriction_\<sigma>_related_set x y)\<close>
+  by (fact dist_restriction_is)
+
+
+lemma not_related_imp_dist_restriction_is_some_restriction_\<sigma> :
+  \<open>\<exists>n. dist\<down> x y = \<sigma>\<down> n \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m) \<close> if \<open>\<not> x \<lessapprox> y\<close>
+proof -
+  have \<open>finite (restriction_related_set x y)\<close>
+    by (simp add: finite_restriction_related_set_iff \<open>\<not> x \<lessapprox> y\<close>)
+  have \<open>Inf (restriction_\<sigma>_related_set x y) \<in> restriction_\<sigma>_related_set x y\<close>
+    by (rule closed_contains_Inf[OF nonempty_restriction_\<sigma>_related_set])
+      (simp_all add: \<open>finite (restriction_related_set x y)\<close> finite_imp_closed)
+  hence \<open>dist\<down> x y \<in> restriction_\<sigma>_related_set x y\<close>
+    by (fold dist_restriction_is)
+  with restriction_related_le obtain n
+    where \<open>n \<in> restriction_related_set x y\<close> \<open>dist\<down> x y = \<sigma>\<down> n\<close>
+      \<open>\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m\<close>  by blast
+  with \<open>dist\<down> x y \<in> restriction_\<sigma>_related_set x y\<close> show ?thesis by blast
+qed
+
+
+lemma not_related_imp_dist_restriction_le_some_restriction_\<sigma> :
+  \<open>\<not> x \<lessapprox> y \<Longrightarrow> \<exists>n. dist\<down> x y \<le> \<sigma>\<down> n \<and> (\<not> x \<down> Suc n \<lessapprox> y \<down> Suc n) \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m)\<close>
+  by (blast intro: restriction_space_Inf_properties
+      dest: ex_not_restriction_related_optimized)
+
+
+lemma restriction_dist_eq_0_iff_related : \<open>dist\<down> x y = 0 \<longleftrightarrow> x \<lessapprox> y\<close>
+proof (rule iffI)
+  show \<open>dist\<down> x y = 0 \<Longrightarrow> x \<lessapprox> y\<close>
+    by (erule contrapos_pp)
+      (auto dest: not_related_imp_dist_restriction_is_some_restriction_\<sigma>)
+next
+  show \<open>dist\<down> x y = 0\<close> if \<open>x \<lessapprox> y\<close>
+  proof (rule order_antisym)
+    show \<open>0 \<le> dist\<down> x y\<close> by (simp add: cINF_greatest dist_restriction_is)
+  next
+    define \<Delta> where \<open>\<Delta> n \<equiv> - \<sigma>\<down> n\<close> for n
+    have * : \<open>\<Delta> n \<le> - dist\<down> x y\<close> for n
+      unfolding \<Delta>_def using \<open>x \<lessapprox> y\<close> restriction_space_Inf_properties(1)
+      by simp (metis UNIV_restriction_related_set_iff rangeI)
+    from restriction_\<sigma>_tendsto_zero tendsto_minus_cancel_left
+    have \<open>\<Delta> \<longlonglongrightarrow> 0\<close> unfolding \<Delta>_def by force
+    from lim_le[OF convergentI[OF \<open>\<Delta> \<longlonglongrightarrow> 0\<close>] "*"] limI[OF \<open>\<Delta> \<longlonglongrightarrow> 0\<close>]
+    show \<open>dist\<down> x y \<le> 0\<close> by simp
+  qed
+qed
+
+
+end
+
+
+locale DecseqRestrictionSpace = NonDecseqRestrictionSpace +
+  assumes decseq_restriction_\<sigma> : \<open>decseq \<sigma>\<down>\<close>
+begin
+
+lemma dist_restriction_is_bis :
+  \<open>dist\<down> x y = (if x \<lessapprox> y then 0 else \<sigma>\<down> (Sup (restriction_related_set x y)))\<close>
+  if \<open>x \<in> M\<close> and \<open>y \<in> M\<close>
+proof (split if_split, intro conjI impI)
+  from \<open>x \<in> M\<close> and \<open>y \<in> M\<close> show \<open>x \<lessapprox> y \<Longrightarrow> restriction_dist x y = 0\<close>
+    by (simp add: restriction_dist_eq_0_iff_related)
+next
+  show \<open>dist\<down> x y = \<sigma>\<down> (Sup (restriction_related_set x y))\<close> if \<open>\<not> x \<lessapprox> y\<close>
+  proof (rule order_antisym)
+    show \<open>dist\<down> x y \<le> \<sigma>\<down> (Sup (restriction_related_set x y))\<close>
+      unfolding dist_restriction_is
+      by (rule cINF_lower[OF _ Sup_in_restriction_related_set[OF \<open>\<not> x \<lessapprox> y\<close>]])
+        (meson bdd_below.I2 zero_le_restriction_\<sigma>)
+  next
+    show \<open>\<sigma>\<down> (Sup (restriction_related_set x y)) \<le> dist\<down> x y\<close>
+    proof (unfold dist_restriction_is, rule cINF_greatest)
+      show \<open>restriction_related_set x y \<noteq> {}\<close> by (fact nonempty_restriction_related_set)
+    next
+      fix n assume \<open>n \<in> restriction_related_set x y\<close>
+      hence \<open>n \<le> Sup (restriction_related_set x y)\<close>
+        by (metis \<open>n \<in> restriction_related_set x y\<close> le_cSup_finite
+            finite_restriction_related_set_iff \<open>\<not> x \<lessapprox> y\<close>)
+      from decseqD[OF decseq_restriction_\<sigma> this]
+      show \<open>restriction_\<sigma> (Sup (restriction_related_set x y)) \<le> restriction_\<sigma> n\<close> .
+    qed
+  qed
+qed
+
+
+lemma not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq :
+  \<open>dist\<down> x y = \<sigma>\<down> (Sup (restriction_related_set x y))\<close>
+  \<open>\<forall>m\<le>Sup (restriction_related_set x y). x \<down> m \<lessapprox> y \<down> m\<close>
+  \<open>\<forall>m>Sup (restriction_related_set x y). \<not> x \<down> m \<lessapprox> y \<down> m\<close>
+  if \<open>\<not> x \<lessapprox> y\<close> and \<open>x \<in> M\<close> and \<open>y \<in> M\<close>
+  subgoal by (subst dist_restriction_is_bis; simp add: \<open>\<not> x \<lessapprox> y\<close> \<open>x \<in> M\<close> \<open>y \<in> M\<close>)
+  subgoal using Sup_in_restriction_related_set[OF \<open>\<not> x \<lessapprox> y\<close>] restriction_related_le by blast
+  using cSup_upper[OF _ bdd_above_restriction_related_set_iff[THEN iffD2, OF \<open>\<not> x \<lessapprox> y\<close>],
+      of \<open>Suc (Sup (restriction_related_set x y))\<close>]
+  by (metis (mono_tags, lifting) Suc_leI dual_order.refl mem_Collect_eq not_less_eq_eq restriction_related_le)
+
+
+
+theorem restriction_dist_tendsto_zero_independent_of_restriction_\<sigma> :
+  \<comment>\<open>Very powerful theorem: the convergence of the distance to term\<open>0 :: real\<close>
+     is actually independent from the restriction sequence chosen.\<close>
+  \<comment>\<open>This is the theorem for which we had to work with locales first.\<close>
+
+assumes \<open>DecseqRestrictionSpace (\<down>) (\<lessapprox>) restriction_\<sigma>' restriction_dist'\<close>
+  and \<open>\<Sigma> \<in> M\<close> and \<open>range \<sigma> \<subseteq> M\<close>
+shows \<open>(\<lambda>n. dist\<down> (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<lambda>n. restriction_dist' (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+proof -
+  { fix restriction_\<sigma> restriction_dist restriction_\<sigma>' restriction_dist'
+    assume a1 : \<open>DecseqRestrictionSpace (\<down>) (\<lessapprox>) restriction_\<sigma>  restriction_dist \<close>
+      and a2 : \<open>DecseqRestrictionSpace (\<down>) (\<lessapprox>) restriction_\<sigma>' restriction_dist'\<close>
+      and  * : \<open>(\<lambda>n. restriction_dist (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+
+    interpret i1 : DecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(\<lessapprox>)\<close> M restriction_\<sigma>  restriction_dist  by (fact a1)
+    interpret i2 : DecseqRestrictionSpace \<open>(\<down>)\<close> \<open>(\<lessapprox>)\<close> M restriction_\<sigma>' restriction_dist' by (fact a2)
+
+
+
+    have \<open>(\<lambda>n. restriction_dist' (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+    proof (rule metric_LIMSEQ_I)
+      fix \<epsilon> :: real assume \<open>0 < \<epsilon>\<close>
+
+      from metric_LIMSEQ_D[OF i2.restriction_\<sigma>_tendsto_zero \<open>0 < \<epsilon>\<close>]
+      obtain N where ** : \<open>N \<le> n \<Longrightarrow> restriction_\<sigma>' n < \<epsilon>\<close> for n by auto
+
+      fix N' :: nat
+
+      have \<open>\<exists>N'. \<forall>n\<ge>N'. N \<in> restriction_related_set (\<sigma> n) \<Sigma>\<close>
+      proof (rule ccontr)
+        assume \<open>\<nexists>N'. \<forall>n\<ge>N'. N \<in> restriction_related_set (\<sigma> n) \<Sigma>\<close>
+        hence *** : \<open>\<forall>N'. \<exists>n\<ge>N'. \<not> \<sigma> n \<down> N \<lessapprox> \<Sigma> \<down> N\<close> by simp
+        have **** : \<open>\<forall>N'. \<exists>n\<ge>N'. restriction_\<sigma> N \<le> restriction_dist (\<sigma> n) \<Sigma>\<close>
+        proof (rule allI)
+          fix N' :: nat
+          from "***" obtain n where \<open>N' \<le> n\<close> \<open>\<not> \<sigma> n \<down> N \<lessapprox> \<Sigma> \<down> N\<close> by blast
+          hence \<open>\<not> \<sigma> n \<lessapprox> \<Sigma>\<close> using mono_restriction_related by blast
+          have \<open>restriction_\<sigma> N \<le> restriction_dist (\<sigma> n) \<Sigma>\<close>
+          proof (subst i1.dist_restriction_is_bis)
+            show \<open>\<sigma> n \<in> M\<close> \<open>\<Sigma> \<in> M\<close>
+              by (simp_all add: assms(2, 3) range_subsetD)
+          next
+            show \<open>restriction_\<sigma> N \<le> (if \<sigma> n \<lessapprox> \<Sigma> then 0
+                  else restriction_\<sigma> (Sup (restriction_related_set (\<sigma> n) \<Sigma>)))\<close>
+              using \<open>\<not> \<sigma> n \<lessapprox> \<Sigma>\<close> \<open>\<not> \<sigma> n \<down> N \<lessapprox> \<Sigma> \<down> N\<close> assms(2, 3) nle_le
+                not_eq_imp_dist_restriction_is_restriction_\<sigma>_Sup_restriction_eq(2)
+              by (fastforce intro: decseqD[OF i1.decseq_restriction_\<sigma>])
+          qed
+          with \<open>N' \<le> n\<close> show \<open>\<exists>n\<ge>N'. restriction_\<sigma> N \<le> restriction_dist (\<sigma> n) \<Sigma>\<close> by blast
+        qed
+        from metric_LIMSEQ_D[OF "*", of \<open>restriction_\<sigma> N\<close>]
+        have \<open>\<exists>N''. \<forall>n\<ge>N''. restriction_dist (\<sigma> n) \<Sigma> < restriction_\<sigma> N\<close>
+          by (metis abs_of_nonneg i1.zero_le_restriction_\<sigma> i1.zero_less_restriction_\<sigma> norm_conv_dist order_trans real_norm_def
+              verit_comp_simplify1(3))
+        with "****" show False by fastforce
+      qed
+
+      then obtain N' where *** : \<open>N' \<le> n \<Longrightarrow> N \<in> restriction_related_set (\<sigma> n) \<Sigma>\<close> for n by blast
+      have \<open>restriction_dist' (\<sigma> n) \<Sigma> < \<epsilon>\<close> if \<open>N' \<le> n\<close> for n
+      proof (rule le_less_trans[OF i2.restriction_space_Inf_properties(1), of \<open>restriction_\<sigma>' N\<close>])
+        from \<open>N' \<le> n\<close> "***"
+        show \<open>restriction_\<sigma>' N \<in> i2.restriction_\<sigma>_related_set (\<sigma> n) \<Sigma>\<close> by blast
+      next
+        show \<open>restriction_\<sigma>' N < \<epsilon>\<close> by (simp add: "**")
+      qed
+      thus \<open>\<exists>N. \<forall>n\<ge>N. dist (restriction_dist' (\<sigma> n) \<Sigma>) 0 < \<epsilon>\<close>
+        by (metis abs_of_nonneg dist_0_norm dist_commute i2.not_related_imp_dist_restriction_is_some_restriction_\<sigma>
+            i2.restriction_dist_eq_0_iff_related i2.zero_less_restriction_\<sigma> order_less_imp_not_less real_norm_def
+            verit_comp_simplify1(3))
+    qed }
+  note * = this
+
+  show \<open>(\<lambda>n. dist\<down> (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0 = (\<lambda>n. restriction_dist' (\<sigma> n) \<Sigma>) \<longlonglongrightarrow> 0\<close>
+    using "*" DecseqRestrictionSpace_axioms assms(1) by blast
+qed
+
+end
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces-Ultrametric/document/root.tex
--- /dev/null
+++ thys/Restriction_Spaces-Ultrametric/document/root.tex
@@ -0,0 +1,74 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{latexsym}
+\usepackage{amssymb}
+\usepackage{amsmath}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{Providing restriction Spaces with an ultrametric Structure}
+\author{Benoît Ballenghien \and Benjamin Puyobro \and Burkhart Wolff}
+\maketitle
+
+\abstract{
+  We investigate the relationship between restriction spaces and classical metric
+  structures by instantiating the former as ultrametric spaces.
+  This is classically captured by defining the distance as
+  $$\mathrm{dist} \ x \ y = \inf_{x \downarrow n \ = \ y \downarrow n} \left(\frac{1}{2}\right)^ n$$
+  but we actually generalize this perspective by introducing a hierarchy of increasingly
+  refined type classes to systematically relate ultrametric and restriction-based notions.
+  This layered approach enables a precise comparison of structural and topological properties.
+  In the end, our main result establishes that completeness in the sense of restriction spaces
+  coincides with standard metric completeness, thus bridging the gap between
+  \verb'Restriction_Spaces' and Banach's fixed-point theorem established in \verb'HOL-Analysis'.  
+}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/ROOT
--- /dev/null
+++ thys/Restriction_Spaces/ROOT
@@ -0,0 +1,50 @@
+chapter AFP
+
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+session Restriction_Spaces = "HOL-Library" +
+  description "Restriction Spaces: a Fixed-Point Theory"
+  options [timeout = 300]
+  theories
+    Restriction_Spaces_Locales
+    Restriction_Spaces_Classes
+    Restriction_Spaces_Prod
+    Restriction_Spaces_Fun
+    Restriction_Spaces_Topology
+    Restriction_Spaces_Induction
+    Restriction_Spaces (global)
+  document_files "root.tex"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces.thy
@@ -0,0 +1,90 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+(*<*)
+theory Restriction_Spaces
+  imports Restriction_Spaces_Induction
+begin
+  (*>*)
+
+section \<open>Entry Point\<close>
+
+text \<open>This is the file session\<open>Restriction_Spaces\<close> should be imported from.\<close>
+
+declare
+  restriction_shift_introset [intro!]
+  restriction_shift_simpset  [simp  ]
+  restriction_cont_simpset   [simp  ]
+  restriction_adm_simpset    [simp  ]
+
+
+
+
+
+text \<open>We already have @{thm non_destructive_id(2)}, and can easily notice
+      term\<open>non_destructive (\<lambda>f. f x)\<close>, but also term\<open>non_destructive (\<lambda>f. f x y)\<close>, etc.
+      We add a theory_text\<open>simproc_setup\<close> to enable the simplifier to automatically handle goals
+      of this form, regardless of the number of arguments on which the function is applied.\<close>
+
+
+simproc_setup apply_non_destructiveness (\<open>non_destructive (\<lambda>f. E f)\<close>) = \<open>
+  K (fn ctxt => fn ct =>
+     let val t = Thm.term_of ct
+         val foo = case t of _ $ foo => foo
+     in  case foo of Abs (_, _, expr) =>
+              if   fst (strip_comb expr) = Bound 0
+              (* since \<open>\<lambda>f. E f\<close> is too permissive, we ensure that the term is of the
+                 form \<open>\<lambda>f. f ...\<close> (for example \<open>\<lambda>f. f x\<close>, or \<open>\<lambda>f. f x y\<close>, etc.) *)
+              then let val tac = Metis_Tactic.metis_tac ["no_types"] "combs" ctxt
+                                 @{thms non_destructive_fun_iff non_destructive_id(2)}
+                       val rwrt_ct  = HOLogic.mk_judgment (Thm.apply cterm\<open>\<lambda>lhs. lhs = True\<close> ct)
+                       val rwrt_thm = Goal.prove_internal ctxt [] rwrt_ct (fn _ => tac 1)
+                   in  SOME (mk_meta_eq rwrt_thm)
+                   end
+              else NONE
+            | _ => NONE
+     end
+    )
+\<close>
+
+
+
+lemma \<open>non_destructive (\<lambda>f. f a b c d e f' g h i j k l m n o' p q r s t u v w x y z)\<close>
+  using [[simp_trace]] by simp \<comment> \<open>test\<close>
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Classes.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Classes.thy
@@ -0,0 +1,755 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Class Implementation\<close>
+
+(*<*)
+theory Restriction_Spaces_Classes
+  imports Restriction_Spaces_Locales
+begin
+  (*>*)
+
+
+subsection \<open>Preliminaries\<close>
+
+text \<open>Small lemma from verbatim\<open>HOL-Library.Infinite_Set\<close> to avoid dependency.\<close>
+
+lemma INFM_nat_le: \<open>(\<exists>\<infinity>n :: nat. P n) \<longleftrightarrow> (\<forall>m. \<exists>n\<ge>m. P n)\<close>
+  unfolding cofinite_eq_sequentially frequently_sequentially by simp
+
+text \<open>We need to be able to extract a subsequence verifying a predicate.\<close>
+
+
+fun extraction_subseq :: \<open>[nat \<Rightarrow> 'a, 'a \<Rightarrow> bool] \<Rightarrow> nat \<Rightarrow> nat\<close>
+  where \<open>extraction_subseq \<sigma> P 0 = (LEAST k. P (\<sigma> k))\<close>
+  | \<open>extraction_subseq \<sigma> P (Suc n) = (LEAST k. extraction_subseq \<sigma> P n < k \<and> P (\<sigma> k))\<close>
+
+
+lemma exists_extraction_subseq :
+  assumes \<open>\<exists>\<infinity>k. P (\<sigma> k)\<close>
+  defines f_def : \<open>f \<equiv> extraction_subseq \<sigma> P\<close>
+  shows \<open>strict_mono f\<close> and \<open>P (\<sigma> (f k))\<close>
+proof -
+  have \<open>f n < f (Suc n) \<and> P (\<sigma> (f n))\<close> for n
+    by (cases n, simp_all add: f_def)
+      (metis (mono_tags, lifting) \<open>\<exists>\<infinity>k. P (\<sigma> k)\<close>[unfolded INFM_nat_le] LeastI_ex Suc_le_eq)+
+  thus \<open>strict_mono f\<close> and \<open>P (\<sigma> (f k))\<close>
+    by (simp_all add: strict_mono_Suc_iff)
+qed
+
+
+lemma extraction_subseqD :
+  \<open>\<exists>f :: nat \<Rightarrow> nat. strict_mono f \<and> (\<forall>k. P (\<sigma> (f k)))\<close> if \<open>\<exists>\<infinity>k. P (\<sigma> k)\<close>
+proof (rule exI)
+  show \<open>strict_mono (extraction_subseq \<sigma> P) \<and> (\<forall>k. P (\<sigma> ((extraction_subseq \<sigma> P) k)))\<close>
+    by (simp add: \<open>\<exists>\<infinity>k. P (\<sigma> k)\<close> exists_extraction_subseq)
+qed
+
+lemma extraction_subseqE :
+  \<comment> \<open>The idea is to abstract the concrete construction of this extraction function,
+      we only need the fact that there is one.\<close>
+  \<open>\<exists>\<infinity>k. P (\<sigma> k) \<Longrightarrow> (\<And>f :: nat \<Rightarrow> nat. strict_mono f \<Longrightarrow> (\<And>k. P (\<sigma> (f k))) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  by (blast dest: extraction_subseqD) 
+
+
+
+subsection \<open>Basic Notions for Restriction\<close>
+
+class restriction = 
+  fixes restriction :: \<open>['a, nat] \<Rightarrow> 'a\<close>  (infixl \<open>\<down>\<close> 60)
+  assumes [simp] : \<open>x \<down> n \<down> m = x \<down> min n m\<close>
+begin
+
+sublocale Restriction \<open>(\<down>)\<close> \<open>(=)\<close> by unfold_locales simp
+    \<comment> \<open>Just to recover const\<open>restriction_related_set\<close> and const\<open>restriction_not_related_set\<close>.\<close>
+end
+
+class restriction_space = restriction +
+  assumes [simp] : \<open>x \<down> 0 = y \<down> 0\<close>
+    and ex_not_restriction_eq : \<open>x \<noteq> y \<Longrightarrow> \<exists>n. x \<down> n \<noteq> y \<down> n\<close>
+begin
+
+sublocale PreorderRestrictionSpace \<open>(\<down>)\<close> \<open>(=)\<close>
+  by unfold_locales (simp_all add: ex_not_restriction_eq)
+
+
+
+lemma restriction_related_set_commute :
+  \<open>restriction_related_set x y = restriction_related_set y x\<close> by auto
+
+lemma restriction_not_related_set_commute :
+  \<open>restriction_not_related_set x y = restriction_not_related_set y x\<close> by auto
+
+end
+
+context restriction_space begin
+
+sublocale Restriction_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'b :: restriction \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(=)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(=)\<close> ..
+
+text \<open>With this we recover constants like const\<open>restriction_shift_on\<close>.\<close>
+
+sublocale PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'b :: restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(=)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(=)\<close> ..
+
+text \<open>With that we recover theorems like @{thm constructive_restriction_restriction}.\<close>
+
+sublocale Restriction_2_PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'c :: restriction \<Rightarrow> nat \<Rightarrow> 'c\<close> \<open>(=)\<close>
+  \<open>(\<down>) :: 'b :: restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(=)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(=)\<close> ..
+
+text \<open>And with that we recover theorems like @{thm constructive_on_comp_non_destructive_on}.\<close>
+
+
+
+lemma restriction_shift_const [restriction_shift_simpset] :
+  \<open>restriction_shift (\<lambda>x. c) k\<close> by (simp add: restriction_shiftI)
+
+lemma constructive_const [restriction_shift_simpset] :
+  \<open>constructive (\<lambda>x. c)\<close> by (simp add: constructiveI)
+
+
+
+
+end
+
+lemma restriction_shift_on_restricted [restriction_shift_simpset] :
+  \<open>restriction_shift_on (\<lambda>x. f x \<down> n) k A\<close> if \<open>restriction_shift_on f k A\<close>
+proof (rule restriction_shift_onI)
+  fix x y m assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down> m = y \<down> m\<close>
+  from restriction_shift_onD[OF \<open>restriction_shift_on f k A\<close> this]
+  have \<open>f x \<down> nat (int m + k) = f y \<down> nat (int m + k)\<close> .
+  hence \<open>f x \<down> nat (int m + k) \<down> n = f y \<down> nat (int m + k) \<down> n\<close> by argo
+  thus \<open>f x \<down> n \<down> nat (int m + k) = f y \<down> n \<down> nat (int m + k)\<close>
+    by (simp add: min.commute)
+qed
+
+lemma restriction_shift_restricted [restriction_shift_simpset] :
+  \<open>restriction_shift f k \<Longrightarrow> restriction_shift (\<lambda>x. f x \<down> n) k\<close>
+  using restriction_shift_def restriction_shift_on_restricted by blast
+
+corollary constructive_restricted [restriction_shift_simpset] :
+  \<open>constructive f \<Longrightarrow> constructive (\<lambda>x. f x \<down> n)\<close>
+  by (simp add: constructive_def constructive_on_def restriction_shift_on_restricted)
+
+corollary non_destructive_restricted [restriction_shift_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive (\<lambda>x. f x \<down> n)\<close>
+  by (simp add: non_destructive_def non_destructive_on_def restriction_shift_on_restricted)
+
+
+lemma non_destructive_id [restriction_shift_simpset] :
+  \<open>non_destructive id\<close> \<open>non_destructive (\<lambda>x. x)\<close>
+  by (simp_all add: id_def non_destructiveI)
+
+
+
+interpretation less_eqRS : Restriction \<open>(\<down>)\<close> \<open>(\<le>)\<close> by unfold_locales
+    \<comment> \<open>Just to recover const\<open>less_eqRS.restriction_related_set\<close> and
+   const\<open>less_eqRS.restriction_not_related_set\<close>.\<close>
+
+
+class preorder_restriction_space = restriction + preorder +
+  assumes restriction_0_less_eq [simp] : \<open>x \<down> 0 \<le> y \<down> 0\<close>
+    and mono_restriction_less_eq     : \<open>x \<le> y \<Longrightarrow> x \<down> n \<le> y \<down> n\<close>
+    and ex_not_restriction_less_eq   : \<open>\<not> x \<le> y \<Longrightarrow> \<exists>n. \<not> x \<down> n \<le> y \<down> n\<close>
+begin
+
+sublocale less_eqRS : PreorderRestrictionSpace \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<le>)\<close>
+proof unfold_locales
+  show \<open>x \<down> 0 \<le> y \<down> 0\<close> for x y :: 'a by simp
+next
+  show \<open>x \<le> y \<Longrightarrow> x \<down> n \<le> y \<down> n\<close> for x y :: 'a and n
+    by (fact mono_restriction_less_eq)
+next
+  show \<open>\<not> x \<le> y \<Longrightarrow> \<exists>n. \<not> x \<down> n \<le> y \<down> n\<close> for x y :: 'a
+    by (simp add: ex_not_restriction_less_eq)
+next
+  show \<open>x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z\<close> for x y z :: 'a by (fact order_trans)
+qed
+
+end
+
+
+class order_restriction_space = preorder_restriction_space + order
+begin
+
+subclass restriction_space
+proof unfold_locales
+  show \<open>x \<down> 0 = y \<down> 0\<close> for x y :: 'a by (rule order_antisym) simp_all
+next
+  show \<open>x \<noteq> y \<Longrightarrow> \<exists>n. x \<down> n \<noteq> y \<down> n\<close> for x y :: 'a
+    by (metis ex_not_restriction_less_eq order.eq_iff)
+qed
+
+end
+
+
+context preorder_restriction_space begin
+
+sublocale less_eqRS : Restriction_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'b :: {restriction, ord} \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<le>)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<le>)\<close> ..
+
+text \<open>With this we recover constants like const\<open>less_eqRS.restriction_shift_on\<close>.\<close>
+
+sublocale less_eqRS : PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'b :: preorder_restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<le>)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<le>)\<close> ..
+
+text \<open>With that we recover theorems like @{thm less_eqRS.constructive_restriction_restriction}.\<close>
+
+sublocale less_eqRS : Restriction_2_PreorderRestrictionSpace_2_PreorderRestrictionSpace
+  \<open>(\<down>) :: 'c :: restriction \<Rightarrow> nat \<Rightarrow> 'c\<close> \<open>(=)\<close>
+  \<open>(\<down>) :: 'b :: preorder_restriction_space \<Rightarrow> nat \<Rightarrow> 'b\<close> \<open>(\<le>)\<close>
+  \<open>(\<down>) :: 'a \<Rightarrow> nat \<Rightarrow> 'a\<close> \<open>(\<le>)\<close> ..
+
+text \<open>And with that we recover theorems like @{thm less_eqRS.constructive_on_comp_non_destructive_on}.\<close>
+
+end
+
+
+context order_restriction_space begin
+
+text \<open>From @{thm order_antisym} we can obtain stronger lemmas.\<close>
+
+corollary order_restriction_shift_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow>
+             f x \<down> nat (int n + k) \<le> f y \<down> nat (int n + k))
+   \<Longrightarrow> restriction_shift_on f k A\<close>
+  by (simp add: order_antisym restriction_shift_onI)
+
+corollary order_restriction_shiftI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow>
+             f x \<down> nat (int n + k) \<le> f y \<down> nat (int n + k))
+   \<Longrightarrow> restriction_shift f k\<close>
+  by (simp add: order_antisym restriction_shiftI)
+
+corollary order_non_too_destructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> Suc n = y \<down> Suc n\<rbrakk> \<Longrightarrow>
+             f x \<down> n \<le> f y \<down> n)
+   \<Longrightarrow> non_too_destructive_on f A\<close>
+  by (simp add: order_antisym non_too_destructive_onI)
+
+corollary order_non_too_destructiveI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> Suc n = y \<down> Suc n\<rbrakk> \<Longrightarrow> f x \<down> n \<le> f y \<down> n)
+   \<Longrightarrow> non_too_destructive f\<close>
+  by (simp add: order_antisym non_too_destructiveI)
+
+corollary order_non_destructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>n \<noteq> 0; x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> n \<le> f y \<down> n)
+   \<Longrightarrow> non_destructive_on f A\<close>
+  by (simp add: order_antisym non_destructive_onI)
+
+corollary order_non_destructiveI :
+  \<open>(\<And>x y n. \<lbrakk>n \<noteq> 0; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> n \<le> f y \<down> n)
+   \<Longrightarrow> non_destructive f\<close>
+  by (simp add: order_antisym non_destructiveI)
+
+corollary order_constructive_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> Suc n \<le> f y \<down> Suc n)
+   \<Longrightarrow> constructive_on f A\<close>
+  by (simp add: order_antisym constructive_onI)
+
+corollary order_constructiveI :
+  \<open>(\<And>x y n. \<lbrakk>f x \<noteq> f y; x \<down> n = y \<down> n\<rbrakk> \<Longrightarrow> f x \<down> Suc n \<le> f y \<down> Suc n)
+   \<Longrightarrow> constructive f\<close>
+  by (simp add: order_antisym constructiveI)
+
+
+end
+
+
+
+
+subsection \<open>Definition of the Fixed-Point Operator\<close>
+
+subsubsection \<open>Preliminaries\<close>
+
+paragraph \<open>Chain\<close>
+
+context restriction begin
+
+definition restriction_chain :: \<open>[nat \<Rightarrow> 'a] \<Rightarrow> bool\<close> (\<open>chain\<down>\<close>)
+  where \<open>restriction_chain \<sigma> \<equiv> \<forall>n. \<sigma> (Suc n) \<down> n = \<sigma> n\<close>
+
+lemma restriction_chainI : \<open>(\<And>n. \<sigma> (Suc n) \<down> n = \<sigma> n) \<Longrightarrow> restriction_chain \<sigma>\<close>
+  and restriction_chainD : \<open>restriction_chain \<sigma> \<Longrightarrow> \<sigma> (Suc n) \<down> n = \<sigma> n\<close>
+  by (simp_all add: restriction_chain_def)
+
+end
+
+
+context restriction_space begin
+
+lemma (in restriction_space) restriction_chain_def_bis:
+  \<open>restriction_chain \<sigma> \<longleftrightarrow> (\<forall>n m. n < m \<longrightarrow> \<sigma> m \<down> n = \<sigma> n)\<close>
+proof (rule iffI)
+  show \<open>\<forall>n m. n < m \<longrightarrow> \<sigma> m \<down> n = \<sigma> n \<Longrightarrow> restriction_chain \<sigma>\<close>
+    by (simp add: restriction_chainI)
+next
+  show \<open>restriction_chain \<sigma> \<Longrightarrow> \<forall>n m. n < m \<longrightarrow> \<sigma> m \<down> n = \<sigma> n\<close>
+  proof (intro allI impI)
+    fix n m
+    assume restriction : \<open>restriction_chain \<sigma>\<close>
+    show \<open>n < m \<Longrightarrow> \<sigma> m \<down> n = \<sigma> n\<close>
+    proof (induct \<open>m - n\<close> arbitrary: m)
+      fix m
+      assume \<open>0 = m - n\<close> and \<open>n < m\<close>
+      hence False by simp
+      thus \<open>\<sigma> m \<down> n = \<sigma> n\<close> by simp
+    next
+      fix x m
+      assume prems : \<open>Suc x = m - n\<close> \<open>n < m\<close>
+      assume hyp : \<open>x = m' - n \<Longrightarrow> n < m' \<Longrightarrow> \<sigma> m' \<down> n = \<sigma> n\<close> for m'
+      obtain m' where \<open>m = Suc m'\<close> by (meson less_imp_Suc_add prems(2))
+      from prems(2) \<open>m = Suc m'\<close> consider \<open>n = m'\<close> | \<open>n < m'\<close> by linarith
+      thus \<open>\<sigma> m \<down> n = \<sigma> n\<close>
+      proof cases
+        show \<open>n = m' \<Longrightarrow> \<sigma> m \<down> n = \<sigma> n\<close> 
+          by (simp add: \<open>m = Suc m'\<close> restriction restriction_chainD)
+      next
+        assume \<open>n < m'\<close>
+        have * : \<open>\<sigma> (Suc m') \<down> n = \<sigma> (Suc m') \<down> m' \<down> n\<close> by (simp add: \<open>n < m'\<close>)
+        from \<open>m = Suc m'\<close> prems(1) have ** : \<open>x = m' - n\<close> by linarith
+        show \<open>\<sigma> m \<down> n = \<sigma> n\<close>
+          by (simp add: "*" "**" \<open>m = Suc m'\<close> \<open>n < m'\<close> hyp restriction restriction_chainD)
+      qed
+    qed
+  qed
+qed
+
+
+lemma restricted_restriction_chain_is :
+  \<open>restriction_chain \<sigma> \<Longrightarrow> (\<lambda>n. \<sigma> n \<down> n) = \<sigma>\<close>
+  by (rule ext) (metis min.idem restriction_chainD restriction_restriction)
+
+lemma restriction_chain_def_ter:
+  \<open>restriction_chain \<sigma> \<longleftrightarrow> (\<forall>n m. n \<le> m \<longrightarrow> \<sigma> m \<down> n = \<sigma> n)\<close>
+  by (metis le_eq_less_or_eq restricted_restriction_chain_is restriction_chain_def_bis)
+
+lemma restriction_chain_restrictions : \<open>restriction_chain ((\<down>) x)\<close>
+  by (simp add: restriction_chainI)
+
+end
+
+
+
+paragraph \<open>Iterations\<close>
+
+text \<open>The sequence of restricted images of powers of a constructive function
+      is a const\<open>restriction_chain\<close>.\<close>
+
+context fixes f :: \<open>'a \<Rightarrow> 'a :: restriction_space\<close> begin
+
+lemma restriction_chain_funpow_restricted [simp]:
+  \<open>restriction_chain (\<lambda>n. (f ^^ n) x \<down> n)\<close> if \<open>constructive f\<close>
+proof (rule restriction_chainI)
+  show \<open>(f ^^ Suc n) x \<down> Suc n \<down> n = (f ^^ n) x \<down> n\<close> for n
+  proof (induct n)
+    show \<open>(f ^^ Suc 0) x \<down> Suc 0 \<down> 0 = (f ^^ 0) x \<down> 0\<close> by simp
+  next
+    fix n assume \<open>(f ^^ Suc n) x \<down> Suc n \<down> n = (f ^^ n) x \<down> n\<close>
+    from \<open>constructive f\<close>[THEN constructiveD, OF this[simplified]]
+    show \<open>(f ^^ Suc (Suc n)) x \<down> Suc (Suc n) \<down> Suc n = (f ^^ Suc n) x \<down> Suc n\<close> by simp
+  qed
+qed
+
+lemma constructive_imp_eq_funpow_restricted :
+  \<open>n \<le> k \<Longrightarrow> n \<le> l \<Longrightarrow> (f ^^ k) x \<down> n = (f ^^ l) y \<down> n\<close> if \<open>constructive f\<close>
+proof (induct n arbitrary: k l)
+  show \<open>(f ^^ k) x \<down> 0 = (f ^^ l) y \<down> 0\<close> for k l by simp
+next
+  fix n k l assume hyp : \<open>n \<le> k \<Longrightarrow> n \<le> l \<Longrightarrow> (f ^^ k) x \<down> n = (f ^^ l) y \<down> n\<close> for k l
+  assume \<open>Suc n \<le> k\<close> \<open>Suc n \<le> l\<close>
+  then obtain k' l' where \<open>k = Suc k'\<close> \<open>n \<le> k'\<close> \<open>l = Suc l'\<close> \<open>n \<le> l'\<close>
+    by (metis Suc_le_D not_less_eq_eq)
+  from \<open>constructive f\<close>[THEN constructiveD, OF hyp[OF \<open>n \<le> k'\<close> \<open>n \<le> l'\<close>]]
+  show \<open>(f ^^ k) x \<down> Suc n = (f ^^ l) y \<down> Suc n\<close>
+    by (simp add: \<open>k = Suc k'\<close> \<open>l = Suc l'\<close>)
+qed
+
+end
+
+
+
+paragraph \<open>Limits and Convergence\<close>
+
+context restriction begin
+
+definition restriction_tendsto :: \<open>[nat \<Rightarrow> 'a, 'a] \<Rightarrow> bool\<close> (\<open>((_)/ \<midarrow>\<down>\<rightarrow> (_))\<close> [59, 59] 59)
+  where \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<equiv> \<forall>n. \<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close>
+
+lemma restriction_tendstoI : \<open>(\<And>n. \<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n) \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_tendsto_def)
+
+lemma restriction_tendstoD : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close>
+  by (simp add: restriction_tendsto_def) 
+
+lemma restriction_tendstoE :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>n0. (\<And>k. n0 \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> n) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_tendstoD by blast
+
+end
+
+
+lemma (in restriction_space) restriction_tendsto_unique :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>' \<Longrightarrow> \<Sigma> = \<Sigma>'\<close>
+  by (metis ex_not_restriction_eq restriction_tendstoD nat_le_linear)
+
+
+context restriction begin
+
+lemma restriction_tendsto_const_restricted :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. \<sigma> n \<down> k) \<midarrow>\<down>\<rightarrow> \<Sigma> \<down> k\<close>
+  unfolding restriction_tendsto_def by metis
+
+lemma restriction_tendsto_iff_eventually_in_restriction_eq_set :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<forall>n. \<exists>n0. \<forall>k\<ge>n0. n \<in> restriction_related_set \<Sigma> (\<sigma> k))\<close>
+  by (simp add: restriction_tendsto_def)
+
+
+lemma restriction_tendsto_const : \<open>(\<lambda>n. \<Sigma>) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_tendstoI)
+
+lemma (in restriction_space) restriction_tendsto_restrictions : \<open>(\<lambda>n. \<Sigma> \<down> n) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (metis restriction_tendstoI restriction_chain_def_ter restriction_chain_restrictions)
+
+lemma restriction_tendsto_shift_iff : \<open>(\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof (rule iffI)
+  show \<open>(\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> if \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>[THEN restriction_tendstoD]
+    show \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> (k + l) \<down> n\<close> for n by (meson trans_le_add1)
+  qed
+next
+  show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> if \<open>(\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    fix n
+    from \<open>(\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>[THEN restriction_tendstoD]
+    obtain n0 where \<open>\<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> (k + l) \<down> n\<close> ..
+    hence \<open>\<forall>k\<ge>n0 + l. \<Sigma> \<down> n = \<sigma> k \<down> n\<close>
+      by (metis add.commute add_leD2 add_le_imp_le_left le_iff_add)
+    thus \<open>\<exists>n1. \<forall>k\<ge>n1. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> ..
+  qed
+qed
+
+
+lemma restriction_tendsto_shiftI : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_tendsto_shift_iff)
+
+lemma restriction_tendsto_shiftD : \<open>(\<lambda>n. \<sigma> (n + l)) \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_tendsto_shift_iff)
+
+
+lemma (in restriction_space) restriction_tendsto_restricted_iff_restriction_tendsto :
+  \<open>(\<lambda>n. \<sigma> n \<down> n) \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof (rule iffI)
+  assume * : \<open>(\<lambda>n. \<sigma> n \<down> n) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    fix n
+    from "*" restriction_tendstoE obtain n0 where \<open>n0 \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> k \<down> n\<close> for k by blast
+    hence \<open>max n0 n \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> n\<close> for k by auto
+    thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> by blast
+  qed
+next
+  assume * : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  show \<open>(\<lambda>n. \<sigma> n \<down> n) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    fix n
+    from "*" restriction_tendstoE obtain n0 where \<open>n0 \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> n\<close> for k by blast
+    hence \<open>max n0 n \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> k \<down> n\<close> for k by auto
+    thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> k \<down> n\<close> by blast
+  qed
+qed
+
+lemma restriction_tendsto_subseq :
+  \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> if \<open>strict_mono f\<close> and \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof (rule restriction_tendstoI)
+  fix n
+  have \<open>n \<le> f n\<close> by (simp add: strict_mono_imp_increasing \<open>strict_mono f\<close>)
+  moreover from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_tendstoE obtain n0 where \<open>n0 \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> n\<close> for k by blast
+  ultimately have \<open>\<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> (f k) \<down> n\<close>
+    by (metis le_trans strict_mono_imp_increasing that(1))
+  thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = (\<sigma> \<circ> f) k \<down> n\<close> by auto
+qed
+
+end
+
+
+context restriction begin
+
+definition restriction_convergent :: \<open>(nat \<Rightarrow> 'a) \<Rightarrow> bool\<close> (\<open>convergent\<down>\<close>)
+  where \<open>restriction_convergent \<sigma> \<equiv> \<exists>\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+
+lemma restriction_convergentI : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> restriction_convergent \<sigma>\<close>
+  by (auto simp add: restriction_convergent_def)
+
+lemma restriction_convergentD' : \<open>restriction_convergent \<sigma> \<Longrightarrow> \<exists>\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_convergent_def)
+
+end
+
+
+context restriction_space begin
+
+lemma restriction_convergentD :
+  \<open>restriction_convergent \<sigma> \<Longrightarrow> \<exists>!\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  unfolding restriction_convergent_def using restriction_tendsto_unique by blast
+
+lemma restriction_convergentE : 
+  \<open>restriction_convergent \<sigma> \<Longrightarrow>
+   (\<And>\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>\<Sigma>'. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>' \<Longrightarrow> \<Sigma>' = \<Sigma>) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_convergentD by blast
+
+lemma restriction_tendsto_of_restriction_convergent :
+  \<open>restriction_convergent \<sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> (THE \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+  by (simp add: restriction_convergentD the1I2)
+
+end
+
+
+
+context restriction begin
+
+lemma restriction_convergent_const [simp] : \<open>convergent\<down> (\<lambda>n. \<Sigma>)\<close>
+  unfolding restriction_convergent_def restriction_tendsto_def by blast
+
+lemma (in restriction_space) restriction_convergent_restrictions [simp] :
+  \<open>convergent\<down> (\<lambda>n. \<Sigma> \<down> n)\<close>
+  using restriction_convergent_def restriction_tendsto_restrictions by blast
+
+lemma restriction_convergent_shift_iff :
+  \<open>convergent\<down> (\<lambda>n. \<sigma> (n + l)) \<longleftrightarrow> convergent\<down> \<sigma>\<close>
+  by (simp add: restriction_convergent_def restriction_tendsto_shift_iff)
+
+
+lemma restriction_convergent_shift_shiftI :
+  \<open>convergent\<down> \<sigma> \<Longrightarrow> convergent\<down> (\<lambda>n. \<sigma> (n + l))\<close>
+  by (simp add: restriction_convergent_shift_iff)
+
+lemma restriction_convergent_shift_shiftD :
+  \<open>convergent\<down> (\<lambda>n. \<sigma> (n + l)) \<Longrightarrow> convergent\<down> \<sigma>\<close>
+  by (simp add: restriction_convergent_shift_iff)
+
+
+lemma (in restriction_space) restriction_convergent_restricted_iff_restriction_convergent :
+  \<open>convergent\<down> (\<lambda>n. \<sigma> n \<down> n) \<longleftrightarrow> convergent\<down> \<sigma>\<close>
+  by (simp add: restriction_convergent_def
+      restriction_tendsto_restricted_iff_restriction_tendsto)
+
+
+lemma restriction_convergent_subseq :
+  \<open>strict_mono f \<Longrightarrow> restriction_convergent \<sigma> \<Longrightarrow> restriction_convergent (\<sigma> \<circ> f)\<close>
+  unfolding restriction_convergent_def using restriction_tendsto_subseq by blast
+
+lemma (in restriction_space)
+  convergent_restriction_chain_imp_ex1 : \<open>\<exists>!\<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n\<close>
+  and restriction_tendsto_of_convergent_restriction_chain : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close>
+  if \<open>restriction_convergent \<sigma>\<close> and \<open>restriction_chain \<sigma>\<close>
+proof -
+  from \<open>restriction_convergent \<sigma>\<close> restriction_convergentE obtain \<Sigma>
+    where \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>\<Sigma>'. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>' \<Longrightarrow> \<Sigma>' = \<Sigma>\<close> by blast
+
+  have * : \<open>\<Sigma> \<down> n = \<sigma> n\<close> for n
+  proof -
+    from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_tendstoE obtain n0 where \<open>n0 \<le> k \<Longrightarrow> \<Sigma> \<down> n = \<sigma> k \<down> n\<close> for k by blast
+    hence \<open>\<Sigma> \<down> n = \<sigma> (max n0 n) \<down> n\<close> by simp
+    thus \<open>\<Sigma> \<down> n = \<sigma> n\<close>
+      by (simp add: restriction_chain_def_ter[THEN iffD1, OF \<open>restriction_chain \<sigma>\<close>])
+  qed
+  have ** : \<open>\<forall>n. \<Sigma>' \<down> n = \<sigma> n \<Longrightarrow> \<Sigma>' = \<Sigma>\<close> for \<Sigma>'
+    by (metis "*" ex_not_restriction_eq)
+  from "*" "**" show \<open>\<exists>!\<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n\<close> by blast
+  from theI'[OF this] "**" have \<open>\<Sigma> = (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close> by simp
+  with \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> show \<open>\<sigma> \<midarrow>\<down>\<rightarrow> (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close> by simp
+qed
+
+end
+
+
+
+subsubsection \<open>Fixed-Point Operator\<close>
+
+text \<open>Our definition only makes sense if such a fixed point exists and is unique.
+      We will therefore directly add a completeness assumption,
+      and define the fixed-point operator within this context.
+      It will only be valid when the function term\<open>f\<close> is const\<open>constructive\<close>.\<close>
+
+class complete_restriction_space = restriction_space +
+  assumes restriction_chain_imp_restriction_convergent : \<open>chain\<down> \<sigma> \<Longrightarrow> convergent\<down> \<sigma>\<close>
+
+definition (in complete_restriction_space)
+  restriction_fix :: \<open>('a \<Rightarrow> 'a) \<Rightarrow> 'a\<close> (* (binder \<open>\<upsilon> \<close> 10) *)
+  \<comment> \<open>We will use a syntax rather than a binder to be compatible with the product.\<close>
+  where \<open>restriction_fix (\<lambda>x. f x) \<equiv> THE \<Sigma>. (\<lambda>n. (f ^^ n) undefined) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+
+syntax "_restriction_fix" :: \<open>[pttrn, 'a \<Rightarrow> 'a] \<Rightarrow> 'a\<close>
+  (\<open>(\<open>indent=3 notation=\<open>binder restriction_fix\<close>\<close>\<upsilon> _./ _)\<close> [0, 10] 10)
+syntax_consts "_restriction_fix" \<rightleftharpoons> restriction_fix
+translations "\<upsilon> x. f" \<rightleftharpoons> "CONST restriction_fix (\<lambda>x. f)"
+print_translation \<open>
+  [(const_syntax\<open>restriction_fix\<close>, fn ctxt => fn [Abs abs] =>
+      let val (x, t) = Syntax_Trans.atomic_abs_tr' ctxt abs
+      in Syntax.const syntax_const\<open>_restriction_fix\<close> $ x $ t end)]
+\<close>  \<comment> \<open>To avoid eta-contraction of body\<close>
+
+
+context complete_restriction_space begin
+
+text \<open>The following result is quite similar to the Banach's fixed point theorem.\<close>
+
+lemma restriction_chain_imp_ex1 : \<open>\<exists>!\<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n\<close>
+  and restriction_tendsto_of_restriction_chain : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close>
+  if \<open>restriction_chain \<sigma>\<close>
+  by (simp_all add: convergent_restriction_chain_imp_ex1
+      restriction_tendsto_of_convergent_restriction_chain
+      restriction_chain_imp_restriction_convergent \<open>restriction_chain \<sigma>\<close>)
+
+
+
+lemma restriction_chain_is :
+  \<open>\<sigma> = (\<down>) (THE \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+  \<open>\<sigma> = (\<down>) (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close> if \<open>restriction_chain \<sigma>\<close>
+proof -
+  from restriction_tendsto_of_restriction_chain[OF \<open>restriction_chain \<sigma>\<close>]
+  have \<open>\<sigma> \<midarrow>\<down>\<rightarrow> (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close> .
+  with restriction_tendsto_of_restriction_convergent
+    restriction_convergentI restriction_tendsto_unique
+  have * : \<open>(THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n) = (THE \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close> by blast
+
+  from theI'[OF restriction_chain_imp_ex1, OF \<open>restriction_chain \<sigma>\<close>]
+  show \<open>\<sigma> = (\<down>) (THE \<Sigma>. \<forall>n. \<Sigma> \<down> n = \<sigma> n)\<close> by (intro ext) simp
+  with "*" show \<open>\<sigma> = (\<down>) (THE \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close> by auto
+qed
+
+end
+
+
+context 
+  fixes f :: \<open>'a \<Rightarrow> 'a :: complete_restriction_space\<close>
+  assumes \<open>constructive f\<close>
+begin
+
+lemma ex1_restriction_fix :
+  \<open>\<exists>!\<Sigma>. \<forall>x. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+proof -
+  let ?\<sigma> = \<open>\<lambda>x n. (f ^^ n) x \<down> n\<close>
+  from constructive_imp_eq_funpow_restricted[OF \<open>constructive f\<close>]
+  have \<open>?\<sigma> x = ?\<sigma> y\<close> for x y by blast
+  moreover have \<open>restriction_chain (?\<sigma> x)\<close> for x by (simp add: \<open>constructive f\<close>)
+  ultimately obtain \<Sigma> where \<open>\<forall>x. ?\<sigma> x \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (metis (full_types) restriction_chain_is(1) restriction_tendsto_restrictions)
+  with restriction_tendsto_unique have \<open>\<exists>!\<Sigma>. \<forall>x. ?\<sigma> x \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  thus \<open>\<exists>!\<Sigma>. \<forall>x. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (simp add: restriction_tendsto_restricted_iff_restriction_tendsto)
+qed
+
+lemma ex1_restriction_fix_bis :
+  \<open>\<exists>!\<Sigma>. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  using ex1_restriction_fix restriction_tendsto_unique by blast
+
+
+lemma restriction_fix_def_bis :
+  \<open>(\<upsilon> x. f x) = (THE \<Sigma>. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+proof -
+  have \<open>(\<lambda>\<Sigma>. (\<lambda>n. (f ^^ n) undefined) \<midarrow>\<down>\<rightarrow> \<Sigma>) = (\<lambda>\<Sigma>. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+    by (metis ex1_restriction_fix restriction_tendsto_unique)
+  from arg_cong[where f = The, OF this, folded restriction_fix_def]
+  show \<open>(\<upsilon> x. f x) = (THE \<Sigma>. (\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close> .
+qed
+
+
+lemma funpow_restriction_tendsto_restriction_fix : \<open>(\<lambda>n. (f ^^ n) x) \<midarrow>\<down>\<rightarrow> (\<upsilon> x. f x)\<close>
+  by (metis ex1_restriction_fix restriction_convergentI
+      restriction_fix_def_bis restriction_tendsto_of_restriction_convergent)
+
+
+lemma restriction_restriction_fix_is : \<open>(\<upsilon> x. f x) \<down> n = (f ^^ n) x \<down> n\<close>
+proof (rule restriction_tendsto_unique)
+  from funpow_restriction_tendsto_restriction_fix
+  show \<open>(\<lambda>k. (f ^^ k) x \<down> n) \<midarrow>\<down>\<rightarrow> (\<upsilon> x. f x) \<down> n\<close>
+    by (simp add: restriction_tendsto_def)
+next
+  from restriction_tendsto_restrictions
+  have \<open>(\<lambda>k. (f ^^ n) x \<down> n \<down> k) \<midarrow>\<down>\<rightarrow> (f ^^ n) x \<down> n\<close> .
+  then obtain n0 where * : \<open>\<forall>k\<ge>n0. (f ^^ n) x \<down> n = (f ^^ n) x \<down> n \<down> k\<close>
+    by (metis restriction_restriction min_def)
+  show \<open>(\<lambda>k. (f ^^ k) x \<down> n) \<midarrow>\<down>\<rightarrow> (f ^^ n) x \<down> n\<close>
+  proof (rule restriction_tendstoI)
+    fix m
+    from * have \<open>\<forall>k\<ge>n + n0 + m. (f ^^ n) x \<down> n \<down> m = (f ^^ k) x \<down> n \<down> m\<close>
+      by (simp add: \<open>constructive f\<close> constructive_imp_eq_funpow_restricted)
+    thus \<open>\<exists>n0. \<forall>k\<ge>n0. (f ^^ n) x \<down> n \<down> m = (f ^^ k) x \<down> n \<down> m\<close> ..
+  qed
+qed
+
+
+lemma restriction_fix_eq : \<open>(\<upsilon> x. f x) = f (\<upsilon> x. f x)\<close>
+proof (subst restriction_fix_def, intro the1_equality restriction_tendstoI)
+  show \<open>\<exists>!\<Sigma>. (\<lambda>n. (f ^^ n) undefined) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (simp add: ex1_restriction_fix_bis)
+next
+  have \<open>n \<le> k \<Longrightarrow> f (restriction_fix f) \<down> n = (f ^^ k) undefined \<down> n\<close> for n k
+    by (cases n, simp, cases k, simp_all)
+      (meson \<open>constructive f\<close>[THEN constructiveD] restriction_related_le
+        restriction_restriction_fix_is)
+  thus \<open>\<exists>n0. \<forall>k\<ge>n0. f (restriction_fix f) \<down> n = (f ^^ k) undefined \<down> n\<close> for n by blast
+qed
+
+
+lemma restriction_fix_unique : \<open>f x = x \<Longrightarrow> (\<upsilon> x. f x) = x\<close>
+  by (metis (no_types, opaque_lifting) restriction_fix_eq \<open>constructive f\<close>
+      constructiveD dual_order.refl take_lemma_restriction)
+
+
+lemma restriction_fix_def_ter : \<open>(\<upsilon> x. f x) = (THE x. f x = x)\<close>
+  by (metis (mono_tags, lifting) restriction_fix_eq restriction_fix_unique the_equality)
+
+
+
+
+end
+
+
+
+(*<*)
+end
+  (*>*)
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Fun.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Fun.thy
@@ -0,0 +1,250 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Functions towards a Restriction Space\<close>
+
+subsection \<open>Restriction Space\<close>
+
+(*<*)
+theory Restriction_Spaces_Fun
+  imports Restriction_Spaces_Classes
+begin
+  (*>*)
+
+instantiation \<open>fun\<close> :: (type, restriction) restriction
+begin
+
+definition restriction_fun :: \<open>['a \<Rightarrow> 'b, nat, 'a] \<Rightarrow> 'b\<close>
+  where \<open>f \<down> n \<equiv> (\<lambda>x. f x \<down> n)\<close>
+
+instance by intro_classes (simp add: restriction_fun_def)
+
+end 
+
+
+instance \<open>fun\<close> :: (type, restriction_space) restriction_space
+proof (intro_classes, unfold restriction_fun_def)
+  show \<open>(\<lambda>x. f x \<down> 0) = (\<lambda>x. g x \<down> 0)\<close>
+    for f g :: \<open>'a \<Rightarrow> 'b\<close> by (rule ext) simp
+next
+  fix f g :: \<open>'a \<Rightarrow> 'b\<close> assume \<open>f \<noteq> g\<close>
+  then obtain x where \<open>f x \<noteq> g x\<close> by fast
+  with ex_not_restriction_related obtain n
+    where \<open>f x \<down> n \<noteq> g x \<down> n\<close> by blast
+  hence \<open>(\<lambda>x. f x \<down> n) \<noteq> (\<lambda>x. g x \<down> n)\<close> by meson
+  thus \<open>\<exists>n. (\<lambda>x. f x \<down> n) \<noteq> (\<lambda>x. g x \<down> n)\<close> ..
+qed
+
+
+instance \<open>fun\<close> :: (type, preorder_restriction_space) preorder_restriction_space
+proof intro_classes
+  show \<open>f \<down> 0 \<le> g \<down> 0\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (simp add: le_fun_def restriction_fun_def)
+next
+  show \<open>f \<le> g \<Longrightarrow> f \<down> n \<le> g \<down> n\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close> and n
+    by (simp add: restriction_fun_def le_fun_def mono_restriction_less_eq)
+next
+  show \<open>\<not> f \<le> g \<Longrightarrow> \<exists>n. \<not> f \<down> n \<le> g \<down> n\<close> for f g :: \<open>'a \<Rightarrow> 'b\<close>
+    by (metis ex_not_restriction_less_eq le_funD le_funI restriction_fun_def)
+qed
+
+instance \<open>fun\<close> :: (type, order_restriction_space) order_restriction_space ..
+
+
+
+
+subsection \<open>Restriction shift Maps\<close>
+
+lemma restriction_shift_on_fun_iff :
+  \<open>restriction_shift_on f k A \<longleftrightarrow> (\<forall>z. restriction_shift_on (\<lambda>x. f x z) k A)\<close>
+proof (intro iffI allI)
+  show \<open>restriction_shift_on (\<lambda>x. f x z) k A\<close> if \<open>restriction_shift_on f k A\<close> for z
+  proof (rule restriction_shift_onI)
+    fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down> n = y \<down> n\<close>
+    from restriction_shift_onD[OF \<open>restriction_shift_on f k A\<close> this]
+    show \<open>f x z \<down> nat (int n + k) = f y z \<down> nat (int n + k)\<close>
+      by (unfold restriction_fun_def) (blast dest!: fun_cong)
+  qed
+next
+  show \<open>restriction_shift_on f k A\<close> if \<open>\<forall>z. restriction_shift_on (\<lambda>x. f x z) k A\<close>
+  proof (rule restriction_shift_onI)
+    fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down> n = y \<down> n\<close>
+    with \<open>\<forall>z. restriction_shift_on (\<lambda>x. f x z) k A\<close> restriction_shift_onD
+    have \<open>f x z \<down> nat (int n + k) = f y z \<down> nat (int n + k)\<close> for z by blast
+    thus \<open>f x \<down> nat (int n + k) = f y \<down> nat (int n + k)\<close>
+      by (simp add: restriction_fun_def)
+  qed
+qed
+
+lemma restriction_shift_fun_iff : \<open>restriction_shift f k \<longleftrightarrow> (\<forall>z. restriction_shift (\<lambda>x. f x z) k)\<close>
+  by (unfold restriction_shift_def, fact restriction_shift_on_fun_iff)
+
+
+lemma non_too_destructive_on_fun_iff:
+  \<open>non_too_destructive_on f A \<longleftrightarrow> (\<forall>z. non_too_destructive_on (\<lambda>x. f x z) A)\<close>
+  by (simp add: non_too_destructive_on_def restriction_shift_on_fun_iff)
+
+lemma non_too_destructive_fun_iff:
+  \<open>non_too_destructive f \<longleftrightarrow> (\<forall>z. non_too_destructive (\<lambda>x. f x z))\<close>
+  by (unfold restriction_shift_def non_too_destructive_def)
+    (fact non_too_destructive_on_fun_iff)
+
+
+lemma non_destructive_on_fun_iff:
+  \<open>non_destructive_on f A \<longleftrightarrow> (\<forall>z. non_destructive_on (\<lambda>x. f x z) A)\<close>
+  by (simp add: non_destructive_on_def restriction_shift_on_fun_iff)
+
+lemma non_destructive_fun_iff:
+  \<open>non_destructive f \<longleftrightarrow> (\<forall>z. non_destructive (\<lambda>x. f x z))\<close>
+  unfolding non_destructive_def by (fact non_destructive_on_fun_iff)
+
+
+lemma constructive_on_fun_iff:
+  \<open>constructive_on f A \<longleftrightarrow> (\<forall>z. constructive_on (\<lambda>x. f x z) A)\<close>
+  by (simp add: constructive_on_def restriction_shift_on_fun_iff)
+
+lemma constructive_fun_iff:
+  \<open>constructive f \<longleftrightarrow> (\<forall>z. constructive (\<lambda>x. f x z))\<close>
+  unfolding constructive_def by (fact constructive_on_fun_iff)
+
+
+
+lemma restriction_shift_fun [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>(\<And>z. restriction_shift (\<lambda>x. f x z) k) \<Longrightarrow> restriction_shift f k\<close>
+  and non_too_destructive_fun [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>(\<And>z. non_too_destructive (\<lambda>x. f x z)) \<Longrightarrow> non_too_destructive f\<close>
+  and non_destructive_fun [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>(\<And>z. non_destructive (\<lambda>x. f x z)) \<Longrightarrow> non_destructive f\<close>
+  and constructive_fun [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>(\<And>z. constructive (\<lambda>x. f x z)) \<Longrightarrow> constructive f\<close>
+  by (simp_all add: restriction_shift_fun_iff non_too_destructive_fun_iff
+      non_destructive_fun_iff constructive_fun_iff)
+
+
+
+subsection \<open>Limits and Convergence\<close>
+
+lemma reached_dist_funE :
+  fixes f g :: \<open>'a \<Rightarrow> 'b :: restriction_space\<close> assumes \<open>f \<noteq> g\<close>
+  obtains x where \<open>f x \<noteq> g x\<close> \<open>Sup (restriction_related_set f g) = Sup (restriction_related_set (f x) (g x))\<close>
+    \<comment>\<open>Morally, we say here that the distance between two functions is reached.
+   But we did not introduce the concept of distance.\<close>
+proof -
+  let ?n = \<open>Sup (restriction_related_set f g)\<close>
+  from Sup_in_restriction_related_set[OF \<open>f \<noteq> g\<close>]
+  have \<open>?n \<in> restriction_related_set f g\<close> .
+  with restriction_related_le have \<open>\<forall>m\<le>?n. f \<down> m = g \<down> m\<close> by blast
+  moreover have \<open>f \<down> Suc ?n \<noteq> g \<down> Suc ?n\<close>
+    using cSup_upper[OF _ bdd_above_restriction_related_set_iff[THEN iffD2, OF \<open>f \<noteq> g\<close>], of \<open>Suc ?n\<close>]
+    by (metis (mono_tags, lifting) dual_order.refl mem_Collect_eq not_less_eq_eq restriction_related_le)
+  ultimately obtain x where * : \<open>\<forall>m\<le>?n. f x \<down> m = g x \<down> m\<close> \<open>f x \<down> Suc ?n \<noteq> g x \<down> Suc ?n\<close>
+    unfolding restriction_fun_def by meson
+  from "*"(2) have \<open>f x \<noteq> g x\<close> by auto
+  moreover from "*" have \<open>?n = Sup (restriction_related_set (f x) (g x))\<close>
+    by (metis (no_types, lifting) \<open>\<forall>m\<le>?n. f \<down> m = g \<down> m\<close>
+        \<open>f \<down> Suc ?n \<noteq> g \<down> Suc ?n\<close> not_less_eq_eq restriction_related_le)
+  ultimately show thesis using that by blast
+qed
+
+
+lemma reached_restriction_related_set_funE : 
+  fixes f g :: \<open>'a \<Rightarrow> 'b :: restriction_space\<close>
+  obtains x where \<open>restriction_related_set f g = restriction_related_set (f x) (g x)\<close>
+proof (cases \<open>f = g\<close>)
+  from that show \<open>f = g \<Longrightarrow> thesis\<close> by simp
+next
+  from that show \<open>f \<noteq> g \<Longrightarrow> thesis\<close>
+    by (elim reached_dist_funE) (metis (full_types) restriction_related_set_is_atMost)
+qed
+
+
+
+lemma restriction_chain_fun_iff :
+  \<open>restriction_chain \<sigma> \<longleftrightarrow> (\<forall>z. restriction_chain (\<lambda>n. \<sigma> n z))\<close>
+proof (intro iffI allI)
+  show \<open>restriction_chain \<sigma> \<Longrightarrow> restriction_chain (\<lambda>n. \<sigma> n z)\<close> for z
+    by (auto simp add: restriction_chain_def restriction_fun_def dest!: fun_cong)
+next
+  show \<open>\<forall>z. restriction_chain (\<lambda>n. \<sigma> n z) \<Longrightarrow> restriction_chain \<sigma>\<close>
+    by (simp add: restriction_chain_def restriction_fun_def)
+qed
+
+
+
+lemma restriction_tendsto_fun_imp : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. \<sigma> n z) \<midarrow>\<down>\<rightarrow> \<Sigma> z\<close>
+  by (simp add: restriction_tendsto_def restriction_fun_def) meson
+
+
+lemma restriction_convergent_fun_imp :
+  \<open>restriction_convergent \<sigma> \<Longrightarrow> restriction_convergent (\<lambda>n. \<sigma> n z)\<close>
+  by (metis restriction_convergent_def restriction_tendsto_fun_imp)
+
+
+subsection \<open>Completeness\<close>
+
+instance \<open>fun\<close> :: (type, complete_restriction_space) complete_restriction_space
+proof intro_classes
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'b :: complete_restriction_space\<close>
+  assume \<open>restriction_chain \<sigma>\<close>
+  hence * : \<open>restriction_chain (\<lambda>n. \<sigma> n x)\<close> for x
+    by (simp add: restriction_chain_fun_iff)
+  from restriction_chain_imp_restriction_convergent[OF this]
+  have ** : \<open>restriction_convergent (\<lambda>n. \<sigma> n x)\<close> for x .
+  then obtain \<Sigma> where *** : \<open>(\<lambda>n. \<sigma> n x) \<midarrow>\<down>\<rightarrow> \<Sigma> x\<close> for x
+    by (meson restriction_convergent_def)
+  from "*" have **** : \<open>(\<lambda>n. \<sigma> n x \<down> n) = (\<lambda>n. \<sigma> n x)\<close> for x
+    by (simp add: restricted_restriction_chain_is)
+  have \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  proof (rule restriction_tendstoI)
+    fix n
+    have \<open>\<forall>k\<ge>n. \<Sigma> x \<down> n = \<sigma> k x \<down> n\<close> for x
+      by (metis "*" "**" "***" "****" restriction_related_le restriction_chain_is(1)
+          restriction_tendsto_of_restriction_convergent restriction_tendsto_unique)
+    hence \<open>\<forall>k\<ge>n. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> by (simp add: restriction_fun_def)
+    thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> by blast
+  qed
+
+  thus \<open>restriction_convergent \<sigma>\<close> by (fact restriction_convergentI)
+qed
+
+
+
+(*<*)
+end
+  (*>*)
+
+
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Induction.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Induction.thy
@@ -0,0 +1,405 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Induction in Restriction Space\<close>
+
+(*<*)
+theory  Restriction_Spaces_Induction
+  imports Restriction_Spaces_Topology
+begin
+  (*>*)
+
+subsection \<open>Admissibility\<close>
+
+named_theorems restriction_adm_simpset \<comment> \<open>For future automation.\<close>
+
+subsubsection \<open>Definition\<close>
+
+text \<open>We start by defining the notion of admissible predicate.
+      The idea is that if this predicates holds for each value
+      of a convergent sequence, it also holds for its limit.\<close>
+
+context restriction begin
+
+definition restriction_adm :: \<open>('a \<Rightarrow> bool) \<Rightarrow> bool\<close> (\<open>adm\<down>\<close>)
+  where \<open>restriction_adm P \<equiv> \<forall>\<sigma> \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longrightarrow> (\<forall>n. P (\<sigma> n)) \<longrightarrow> P \<Sigma>\<close>
+
+lemma restriction_admI :
+  \<open>(\<And>\<sigma> \<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>n. P (\<sigma> n)) \<Longrightarrow> P \<Sigma>) \<Longrightarrow> restriction_adm P\<close>
+  by (simp add: restriction_adm_def)
+
+lemma restriction_admD :
+  \<open>\<lbrakk>restriction_adm P; \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>; \<And>n. P (\<sigma> n)\<rbrakk> \<Longrightarrow> P \<Sigma>\<close>
+  by (simp add: restriction_adm_def)
+
+
+subsubsection \<open>Properties\<close>
+
+lemma restriction_adm_const [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. t)\<close>
+  by (simp add: restriction_admI)
+
+lemma restriction_adm_conj [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. P x) \<Longrightarrow> adm\<down> (\<lambda>x. Q x) \<Longrightarrow> adm\<down> (\<lambda>x. P x \<and> Q x)\<close>
+  by (fast intro: restriction_admI elim: restriction_admD)
+
+lemma restriction_adm_all [restriction_adm_simpset] :
+  \<open>(\<And>y. adm\<down> (\<lambda>x. P x y)) \<Longrightarrow> adm\<down> (\<lambda>x. \<forall>y. P x y)\<close>
+  by (fast intro: restriction_admI elim: restriction_admD)
+
+lemma restriction_adm_ball [restriction_adm_simpset] :
+  \<open>(\<And>y. y \<in> A \<Longrightarrow> adm\<down> (\<lambda>x. P x y)) \<Longrightarrow> adm\<down> (\<lambda>x. \<forall>y\<in>A. P x y)\<close>
+  by (fast intro: restriction_admI elim: restriction_admD)
+
+lemma restriction_adm_disj [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. P x \<or> Q x)\<close> if \<open>adm\<down> (\<lambda>x. P x)\<close> \<open>adm\<down> (\<lambda>x. Q x)\<close>
+proof (rule restriction_admI)
+  fix \<sigma> \<Sigma>
+  assume * : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. P (\<sigma> n) \<or> Q (\<sigma> n)\<close>
+  from "*"(2) have ** : \<open>(\<forall>i. \<exists>j\<ge>i. P (\<sigma> j)) \<or> (\<forall>i. \<exists>j\<ge>i. Q (\<sigma> j))\<close>
+    by (meson nat_le_linear)
+
+  { fix P assume $ : \<open>adm\<down> (\<lambda>x. P x)\<close> \<open>\<forall>i. \<exists>j\<ge>i. P (\<sigma> j)\<close>
+    define f where \<open>f i = (LEAST j. i \<le> j \<and> P (\<sigma> j))\<close> for i 
+    have f1: \<open>\<And>i. i \<le> f i\<close> and f2: \<open>\<And>i. P (\<sigma> (f i))\<close>
+      using LeastI_ex [OF "$"(2)[rule_format]] by (simp_all add: f_def)
+    have f3 : \<open>(\<lambda>n. \<sigma> (f n)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    proof (rule restriction_tendstoI)
+      fix n
+      from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_tendstoD obtain n0 where \<open>\<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> by blast
+      hence \<open>\<forall>k\<ge>max n0 n. \<Sigma> \<down> n = \<sigma> (f k) \<down> n\<close> by (meson f1 le_trans max.boundedE)
+      thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> (f k) \<down> n\<close> by blast
+    qed
+    have \<open>P \<Sigma>\<close> by (fact restriction_admD[OF "$"(1) f3 f2])
+  }
+
+  with "**" \<open>adm\<down> (\<lambda>x. P x)\<close> \<open>adm\<down> (\<lambda>x. Q x)\<close> show \<open>P \<Sigma> \<or> Q \<Sigma>\<close> by blast
+qed
+
+lemma restriction_adm_imp [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. \<not> P x) \<Longrightarrow> adm\<down> (\<lambda>x. Q x) \<Longrightarrow> adm\<down> (\<lambda>x. P x \<longrightarrow> Q x)\<close>
+  by (subst imp_conv_disj) (rule restriction_adm_disj)
+
+
+lemma restriction_adm_iff [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. P x \<longrightarrow> Q x) \<Longrightarrow> adm\<down> (\<lambda>x. Q x \<longrightarrow> P x) \<Longrightarrow> adm\<down> (\<lambda>x. P x \<longleftrightarrow> Q x)\<close>
+  by (subst iff_conv_conj_imp) (rule restriction_adm_conj)
+
+lemma restriction_adm_if_then_else [restriction_adm_simpset]:
+  \<open>\<lbrakk>P \<Longrightarrow> adm\<down> (\<lambda>x. Q x); \<not> P \<Longrightarrow> adm\<down> (\<lambda>x. R x)\<rbrakk> \<Longrightarrow>
+   adm\<down> (\<lambda>x. if P then Q x else R x)\<close>
+  by (simp add: restriction_adm_def)
+
+end
+
+
+
+text \<open>The notion of continuity is of course strongly related to the notion of admissibility.\<close>
+
+lemma restriction_adm_eq [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. f x = g x)\<close> if \<open>cont\<down> f\<close> and \<open>cont\<down> g\<close>
+for f g :: \<open>'a :: restriction \<Rightarrow> 'b :: restriction_space\<close>
+proof (rule restriction_admI)
+  fix \<sigma> \<Sigma> assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> and \<open>\<And>n. f (\<sigma> n) = g (\<sigma> n)\<close>
+  from restriction_contD[OF \<open>cont\<down> f\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>] have \<open>(\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> .
+  hence \<open>(\<lambda>n. g (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> by (unfold \<open>\<And>n. f (\<sigma> n) = g (\<sigma> n)\<close>)
+  moreover from restriction_contD[OF \<open>cont\<down> g\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>] have \<open>(\<lambda>n. g (\<sigma> n)) \<midarrow>\<down>\<rightarrow> g \<Sigma>\<close> .
+  ultimately show \<open>f \<Sigma> = g \<Sigma>\<close> by (fact restriction_tendsto_unique)
+qed
+
+lemma restriction_adm_subst [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. P (t x))\<close> if \<open>cont\<down> (\<lambda>x. t x)\<close> and \<open>adm\<down> P\<close>
+proof (rule restriction_admI)
+  fix \<sigma> \<Sigma> assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. P (t (\<sigma> n))\<close>
+  from restriction_contD[OF \<open>cont\<down> (\<lambda>x. t x)\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>]
+  have \<open>(\<lambda>n. t (\<sigma> n)) \<midarrow>\<down>\<rightarrow> t \<Sigma>\<close> .
+  from restriction_admD[OF \<open>restriction_adm P\<close> \<open>(\<lambda>n. t (\<sigma> n)) \<midarrow>\<down>\<rightarrow> t \<Sigma>\<close> \<open>\<And>n. P (t (\<sigma> n))\<close>]
+  show \<open>P (t \<Sigma>)\<close> .
+qed
+
+
+
+lemma restriction_adm_prod_domainD [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>x. P (x, y))\<close> and \<open>adm\<down> (\<lambda>y. P (x, y))\<close> if \<open>adm\<down> P\<close>
+proof -
+  show \<open>adm\<down> (\<lambda>x. P (x, y))\<close>
+  proof (rule restriction_admI)
+    show \<open>P (\<Sigma>, y)\<close> if \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. P (\<sigma> n, y)\<close> for \<sigma> \<Sigma>
+    proof (rule restriction_admD[OF \<open>adm\<down> P\<close> _ \<open>\<And>n. P (\<sigma> n, y)\<close>])
+      show \<open>(\<lambda>n. (\<sigma> n, y)) \<midarrow>\<down>\<rightarrow> (\<Sigma>, y)\<close>
+        by (simp add: restriction_tendsto_prod_iff restriction_tendsto_const \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>)
+    qed
+  qed
+next
+  show \<open>adm\<down> (\<lambda>y. P (x, y))\<close>
+  proof (rule restriction_admI)
+    show \<open>P (x, \<Sigma>)\<close> if \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<And>n. P (x, \<sigma> n)\<close> for \<sigma> \<Sigma>
+    proof (rule restriction_admD[OF \<open>adm\<down> P\<close> _ \<open>\<And>n. P (x, \<sigma> n)\<close>])
+      show \<open>(\<lambda>n. (x, \<sigma> n)) \<midarrow>\<down>\<rightarrow> (x, \<Sigma>)\<close>
+        by (simp add: restriction_tendsto_prod_iff restriction_tendsto_const \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>)
+    qed
+  qed
+qed
+
+
+
+lemma restriction_adm_restriction_shift_on [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. restriction_shift_on f k A)\<close>
+proof (rule restriction_admI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'b\<close> and \<Sigma> :: \<open>'a \<Rightarrow> 'b\<close>
+  assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> and hyp : \<open>restriction_shift_on (\<sigma> n) k A\<close> for n
+  show \<open>restriction_shift_on \<Sigma> k A\<close>
+  proof (rule restriction_shift_onI)
+    fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down> n = y \<down> n\<close>
+    from hyp[THEN restriction_shift_onD, OF this]
+    have * : \<open>\<sigma> m x \<down> nat (int n + k) = \<sigma> m y \<down> nat (int n + k)\<close> for m .
+    show \<open>\<Sigma> x \<down> nat (int n + k) = \<Sigma> y \<down> nat (int n + k)\<close>
+    proof (rule restriction_tendsto_unique)
+      show \<open>(\<lambda>m. \<sigma> m x \<down> nat (int n + k)) \<midarrow>\<down>\<rightarrow> \<Sigma> x \<down> nat (int n + k)\<close>
+        by (simp add: \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_tendsto_const_restricted restriction_tendsto_fun_imp)
+    next
+      show \<open>(\<lambda>m. \<sigma> m x \<down> nat (int n + k)) \<midarrow>\<down>\<rightarrow> \<Sigma> y \<down> nat (int n + k)\<close>
+        by (simp add: "*" \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_tendsto_const_restricted restriction_tendsto_fun_imp)
+    qed
+  qed
+qed
+
+lemma restriction_adm_constructive_on [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. constructive_on f A)\<close>
+  by (simp add: constructive_on_def restriction_adm_restriction_shift_on)
+
+lemma restriction_adm_non_destructive_on [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. non_destructive_on f A)\<close>
+  by (simp add: non_destructive_on_def restriction_adm_restriction_shift_on)
+
+
+lemma restriction_adm_restriction_cont_at [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. cont\<down> f at a)\<close>
+proof (rule restriction_admI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a \<Rightarrow> 'b\<close> and \<Sigma> :: \<open>'a \<Rightarrow> 'b\<close>
+  assume \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> and hyp : \<open>cont\<down> (\<sigma> n) at a\<close> for n
+  show \<open>cont\<down> \<Sigma> at a\<close>
+  proof (rule restriction_cont_atI)
+    fix \<gamma> assume \<open>\<gamma> \<midarrow>\<down>\<rightarrow> a\<close>
+    from hyp[THEN restriction_cont_atD, OF this, THEN restriction_tendstoD]
+    have \<open>\<exists>n0. \<forall>k\<ge>n0. \<sigma> m a \<down> n = \<sigma> m (\<gamma> k) \<down> n\<close> for m n .
+    moreover from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>[THEN restriction_tendstoD]
+    have \<open>\<exists>n0. \<forall>k\<ge>n0. \<Sigma> \<down> n = \<sigma> k \<down> n\<close> for n .
+    ultimately show \<open>(\<lambda>n. \<Sigma> (\<gamma> n)) \<midarrow>\<down>\<rightarrow> \<Sigma> a\<close>
+      by (intro restriction_tendstoI) (metis restriction_fun_def)
+  qed
+qed
+
+lemma restriction_adm_restriction_cont_on [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. cont\<down> f on A)\<close>
+  unfolding restriction_cont_on_def
+  by (intro restriction_adm_ball restriction_adm_restriction_cont_at)
+
+
+
+corollary restriction_adm_restriction_shift [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. restriction_shift f k)\<close>
+  and    restriction_adm_constructive [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. constructive f)\<close>
+  and    restriction_adm_non_destructive [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. non_destructive f)\<close>
+  and    restriction_adm_restriction_cont [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>f. cont\<down> f)\<close>
+  by (simp_all add: restriction_adm_simpset restriction_shift_def
+      constructive_def non_destructive_def)
+
+
+
+lemma (in restriction) restriction_adm_mem_restriction_closed [restriction_adm_simpset] :
+  \<open>closed\<down> K \<Longrightarrow> adm\<down> (\<lambda>x. x \<in> K)\<close>
+  by (auto intro!: restriction_admI dest: restriction_closed_sequentialD)
+
+lemma (in restriction_space) restriction_adm_mem_restriction_compact [restriction_adm_simpset] :
+  \<open>compact\<down> K \<Longrightarrow> adm\<down> (\<lambda>x. x \<in> K)\<close>
+  by (simp add: restriction_adm_mem_restriction_closed restriction_compact_imp_restriction_closed)
+
+lemma (in restriction_space) restriction_adm_mem_finite [restriction_adm_simpset] :
+  \<open>finite S \<Longrightarrow> adm\<down> (\<lambda>x. x \<in> S)\<close>
+  by (simp add: finite_imp_restriction_compact restriction_adm_mem_restriction_compact)
+
+
+lemma restriction_adm_restriction_tendsto [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>\<sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+  by (intro restriction_admI restriction_tendstoI)
+    (metis (no_types, opaque_lifting) restriction_fun_def restriction_tendsto_def)
+
+lemma restriction_adm_lim [restriction_adm_simpset] :
+  \<open>adm\<down> (\<lambda>\<Sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+  by (metis restriction_admI restriction_openD restriction_open_restriction_cball
+      restriction_tendsto_iff_restriction_cball_characterization)
+
+
+lemma restriction_restriction_cont_on [restriction_cont_simpset] :
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> (\<lambda>x. f x \<down> n) on A\<close>
+  by (rule restriction_cont_onI)
+    (simp add: restriction_cont_onD restriction_tendsto_const_restricted)
+
+lemma restriction_cont_on_id [restriction_cont_simpset] : \<open>cont\<down> (\<lambda>x. x) on A\<close>
+  by (simp add: restriction_cont_onI)
+
+lemma restriction_cont_on_const [restriction_cont_simpset] : \<open>cont\<down> (\<lambda>x. c) on A\<close>
+  by (simp add: restriction_cont_onI restriction_tendstoI)
+
+lemma restriction_cont_on_fun [restriction_cont_simpset] : \<open>cont\<down> (\<lambda>f. f x) on A\<close>
+  by (rule restriction_cont_onI) (simp add: restriction_tendsto_fun_imp)
+
+lemma restriction_cont2cont_on_fun [restriction_cont_simpset] :
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> (\<lambda>x. f x y) on A\<close>
+  by (rule restriction_cont_onI)
+    (metis restriction_cont_onD restriction_tendsto_fun_imp)
+
+
+
+
+
+
+
+subsection \<open>Induction\<close>
+
+text \<open>Now that we have the concept of admissibility,
+      we can formalize an induction rule for fixed points.
+      Considering a const\<open>constructive\<close> function term\<open>f\<close> of type typ\<open>'a \<Rightarrow> 'a\<close>
+      (where typ\<open>'a\<close> is instance of the class class\<open>complete_restriction_space\<close>)
+      and a predicate term\<open>P\<close> which is admissible, and assuming that :
+      item term\<open>P\<close> holds for a certain element term\<open>x\<close>
+      item for any element term\<open>x\<close>, if term\<open>P\<close> holds for term\<open>x\<close> then it still holds for term\<open>f x\<close>
+      
+      we can have that term\<open>P\<close> holds for the fixed point term\<open>\<upsilon> x. P x\<close>.\<close>
+
+lemma restriction_fix_ind' [case_names constructive adm steps] :
+  \<open>constructive f \<Longrightarrow> adm\<down> P \<Longrightarrow> (\<And>n. P ((f ^^ n) x)) \<Longrightarrow> P (\<upsilon> x. f x)\<close>
+  using restriction_admD funpow_restriction_tendsto_restriction_fix by blast
+
+lemma restriction_fix_ind [case_names constructive adm base step] :
+  \<open>P (\<upsilon> x. f x)\<close> if \<open>constructive f\<close> \<open>adm\<down> P\<close> \<open>P x\<close> \<open>\<And>x. P x \<Longrightarrow> P (f x)\<close>
+proof (induct rule: restriction_fix_ind')
+  show \<open>constructive f\<close> by (fact \<open>constructive f\<close>)
+next
+  show \<open>restriction_adm P\<close> by (fact \<open>restriction_adm P\<close>)
+next
+  show \<open>P ((f ^^ n) x)\<close> for n
+    by (induct n) (simp_all add: \<open>P x\<close> \<open>\<And>x. P x \<Longrightarrow> P (f x)\<close>)
+qed
+
+lemma restriction_fix_ind2 [case_names constructive adm base0 base1 step] :
+  \<open>P (\<upsilon> x. f x)\<close> if \<open>constructive f\<close> \<open>adm\<down> P\<close> \<open>P x\<close> \<open>P (f x)\<close>
+  \<open>\<And>x. \<lbrakk>P x; P (f x)\<rbrakk> \<Longrightarrow> P (f (f x))\<close>
+proof (induct rule: restriction_fix_ind')
+  show \<open>constructive f\<close> by (fact \<open>constructive f\<close>)
+next
+  show \<open>restriction_adm P\<close> by (fact \<open>restriction_adm P\<close>)
+next
+  show \<open>P ((f ^^ n) x)\<close> for n
+    by (induct n rule: induct_nat_012) (simp_all add: that(3-5))
+qed
+
+
+text \<open>We can rewrite the fixed point over a product to
+      obtain this parallel fixed point induction rule. \<close>
+
+lemma parallel_restriction_fix_ind [case_names constructiveL constructiveR adm base step] :
+  fixes f :: \<open>'a :: complete_restriction_space \<Rightarrow> 'a\<close>
+    and g :: \<open>'b :: complete_restriction_space \<Rightarrow> 'b\<close>
+  assumes constructive : \<open>constructive f\<close> \<open>constructive g\<close>
+    and adm  : \<open>restriction_adm (\<lambda>p. P (fst p) (snd p))\<close>
+    and base : \<open>P x y\<close> and step : \<open>\<And>x y. P x y \<Longrightarrow> P (f x) (g y)\<close>
+  shows \<open>P (\<upsilon> x. f x) (\<upsilon> y. g y)\<close>
+proof -
+  define F where \<open>F \<equiv> \<lambda>(x, y). (f x, g y)\<close>
+  define Q where \<open>Q \<equiv> \<lambda>(x, y). P x y\<close>
+
+  have \<open>P (\<upsilon> x. f x) (\<upsilon> y. g y) = Q (\<upsilon> p. F p)\<close>
+    by (simp add: F_def Q_def constructive restriction_fix_indep_prod_is)
+  also have \<open>Q (\<upsilon> p. F p)\<close> 
+  proof (induct F rule : restriction_fix_ind)
+    show \<open>constructive F\<close>
+      by (simp add: F_def constructive_prod_codomain_iff constructive_prod_domain_iff constructive constructive_const)
+  next
+    show \<open>restriction_adm Q\<close>
+      by (unfold Q_def) (metis (mono_tags, lifting) adm case_prod_beta restriction_adm_def)
+  next
+    show \<open>Q (x, y)\<close> by (simp add: Q_def base)
+  next
+    show \<open>Q p \<Longrightarrow> Q (F p)\<close> for p by (simp add: Q_def F_def step split_beta)
+  qed
+  finally show \<open>P (\<upsilon> x. f x) (\<upsilon> y. g y)\<close> .
+qed
+
+
+text \<open>k-steps induction\<close>
+
+lemma restriction_fix_ind_k_steps [case_names constructive adm base_k_steps step] :
+  assumes \<open>constructive f\<close>
+    and \<open>adm\<down> P\<close>
+    and \<open>\<forall>i < k. P ((f ^^ i) x)\<close>
+    and \<open>\<And>x. \<forall>i < k. P ((f ^^ i) x) \<Longrightarrow> P ((f ^^ k) x)\<close>
+  shows \<open>P (\<upsilon> x. f x)\<close>
+proof (rule restriction_fix_ind')
+  show \<open>constructive f\<close> by (fact \<open>constructive f\<close>)
+next
+  show \<open>adm\<down> P\<close> by (fact \<open>adm\<down> P\<close>)
+next
+  have nat_k_induct :
+    \<open>P n\<close> if \<open>\<forall>i<k. P i\<close> and \<open>\<forall>n0. (\<forall>i<k. P (n0 + i)) \<longrightarrow> P (n0 + k)\<close> for k n :: nat and P
+  proof (induct rule: nat_less_induct)
+    fix n assume \<open>\<forall>m<n. P m\<close>
+    show \<open>P n\<close>
+    proof (cases \<open>n < k\<close>)
+      from that(1) show \<open>n < k \<Longrightarrow> P n\<close> by blast
+    next
+      from \<open>\<forall>m<n. P m\<close> that(2)[rule_format, of \<open>n - k\<close>]
+      show \<open>\<not> n < k \<Longrightarrow> P n\<close> by auto
+    qed
+  qed
+  show \<open>P ((f ^^ i) x)\<close> for i
+  proof (induct rule: nat_k_induct)
+    show \<open>\<forall>i<k. P ((f ^^ i) x)\<close> by (simp add: assms(3))
+  next
+    show \<open>\<forall>n0. (\<forall>i<k. P ((f ^^ (n0 + i)) x)) \<longrightarrow> P ((f ^^ (n0 + k)) x)\<close>
+      by (smt (verit, del_insts) add.commute assms(4) funpow_add o_apply)
+  qed
+qed
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Locales.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Locales.thy
@@ -0,0 +1,772 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Locales factorizing the proof Work\<close>
+
+(*<*)
+theory Restriction_Spaces_Locales
+  imports Main
+begin
+  (*>*)
+
+named_theorems restriction_shift_simpset
+named_theorems restriction_shift_introset \<comment> \<open>Useful for future automation.\<close>
+
+
+text \<open>In order to factorize the proof work, we first work with locales and then with classes.\<close>
+
+
+subsection \<open>Basic Notions for Restriction\<close>
+
+locale Restriction =
+  fixes restriction :: \<open>['a, nat] \<Rightarrow> 'a\<close>  (infixl \<open>\<down>\<close> 60)
+    and relation    :: \<open>['a, 'a] \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>\<close> 50)
+  assumes restriction_restriction [simp] : \<open>x \<down> n \<down> m = x \<down> min n m\<close>
+begin
+
+
+abbreviation restriction_related_set  :: \<open>'a \<Rightarrow> 'a \<Rightarrow> nat set\<close>
+  where \<open>restriction_related_set x y  \<equiv> {n. x \<down> n \<lessapprox> y \<down> n}\<close>
+
+abbreviation restriction_not_related_set :: \<open>'a \<Rightarrow> 'a \<Rightarrow> nat set\<close>
+  where \<open>restriction_not_related_set x y \<equiv> {n. \<not> x \<down> n \<lessapprox> y \<down> n}\<close>
+
+lemma restriction_related_set_Un_restriction_not_related_set :
+  \<open>restriction_related_set x y \<union> restriction_not_related_set x y = UNIV\<close> by blast
+
+lemma disjoint_restriction_related_set_restriction_not_related_set :
+  \<open>restriction_related_set x y \<inter> restriction_not_related_set x y = {}\<close> by blast
+
+lemma \<open>bdd_below (restriction_related_set x y)\<close> by (fact bdd_below_bot)
+
+lemma \<open>bdd_below (restriction_not_related_set x y)\<close> by (fact bdd_below_bot)
+
+
+end
+
+
+locale PreorderRestrictionSpace = Restriction +
+  assumes restriction_0_related   [simp] : \<open>x \<down> 0 \<lessapprox> y \<down> 0\<close>
+    and mono_restriction_related   : \<open>  x \<lessapprox> y \<Longrightarrow> x \<down> n \<lessapprox> y \<down> n\<close>
+    and ex_not_restriction_related : \<open>\<not> x \<lessapprox> y \<Longrightarrow> \<exists>n. \<not> x \<down> n \<lessapprox> y \<down> n\<close>
+    and related_trans : \<open>x \<lessapprox> y \<Longrightarrow> y \<lessapprox> z \<Longrightarrow> x \<lessapprox> z\<close>
+begin
+
+lemma exists_restriction_related [simp] : \<open>\<exists>n. x \<down> n \<lessapprox> y \<down> n\<close>
+  using restriction_0_related by blast
+
+lemma all_restriction_related_iff_related : \<open>(\<forall>n. x \<down> n \<lessapprox> y \<down> n) \<longleftrightarrow> x \<lessapprox> y\<close>
+  using mono_restriction_related ex_not_restriction_related by blast
+
+lemma restriction_related_le : \<open>x \<down> n \<lessapprox> y \<down> n\<close> if \<open>n \<le> m\<close> and \<open>x \<down> m \<lessapprox> y \<down> m\<close>
+proof -
+  from mono_restriction_related[OF \<open>x \<down> m \<lessapprox> y \<down> m\<close>] have \<open>x \<down> m \<down> n \<lessapprox> y \<down> m \<down> n\<close> .
+  also have \<open>(\<lambda>x. x \<down> m \<down> n) = (\<lambda>x. x \<down> n)\<close> by (simp add: \<open>n \<le> m\<close>)
+  finally show \<open>x \<down> n \<lessapprox> y \<down> n\<close> .
+qed
+
+corollary restriction_related_pred : \<open>x \<down> Suc n \<lessapprox> y \<down> Suc n \<Longrightarrow> x \<down> n \<lessapprox> y \<down> n\<close>
+  by (metis le_add2 plus_1_eq_Suc restriction_related_le)
+
+lemma all_ge_restriction_related_iff_related : \<open>(\<forall>n\<ge>m. x \<down> n \<lessapprox> y \<down> n) \<longleftrightarrow> x \<lessapprox> y\<close>
+  by (metis all_restriction_related_iff_related nle_le restriction_related_le)
+
+
+lemma take_lemma_restriction : \<open>x \<lessapprox> y\<close>
+  if \<open>\<And>n. \<lbrakk>\<And>k. k \<le> n \<Longrightarrow> x \<down> k \<lessapprox> y \<down> k\<rbrakk> \<Longrightarrow> x \<down> Suc n \<lessapprox> y \<down> Suc n\<close>
+proof (subst all_restriction_related_iff_related[symmetric], intro allI)
+  show \<open>x \<down> n \<lessapprox> y \<down> n\<close> for n
+    by (induct n rule: full_nat_induct)
+      (metis not_less_eq_eq restriction_0_related restriction_related_le that zero_induct)
+qed
+
+lemma ex_not_restriction_related_optimized :
+  \<open>\<exists>!n. \<not> x \<down> Suc n \<lessapprox> y \<down> Suc n \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m)\<close> if \<open>\<not> x \<lessapprox> y\<close>
+proof (rule ex1I)
+  let ?S = \<open>{n. \<not> x \<down> Suc n \<lessapprox> y \<down> Suc n \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m)}\<close>
+  let ?n = \<open>Inf {n. \<not> x \<down> Suc n \<lessapprox> y \<down> Suc n \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m)}\<close>
+  from restriction_related_le[of _ _ x y] take_lemma_restriction[of x y] \<open>\<not> x \<lessapprox> y\<close>
+  have \<open>?S \<noteq> {}\<close> by auto
+  from Inf_nat_def1[OF this] have \<open>?n \<in> ?S\<close> .
+  thus \<open>\<not> x \<down> Suc ?n \<lessapprox> y \<down> Suc ?n \<and> (\<forall>m\<le>?n. x \<down> m \<lessapprox> y \<down> m)\<close> by blast
+  thus \<open>\<not> x \<down> Suc n \<lessapprox> y \<down> Suc n \<and> (\<forall>m\<le>n. x \<down> m \<lessapprox> y \<down> m) \<Longrightarrow> n = ?n\<close> for n
+    by (meson not_less_eq_eq order_antisym_conv)
+qed
+
+
+
+
+lemma nonempty_restriction_related_set : \<open>restriction_related_set x y \<noteq> {}\<close>
+  using restriction_0_related by blast
+
+lemma non_UNIV_restriction_not_related_set : \<open>restriction_not_related_set x y \<noteq> UNIV\<close>
+  using restriction_0_related by blast
+
+lemma UNIV_restriction_related_set_iff : \<open>restriction_related_set x y = UNIV \<longleftrightarrow> x \<lessapprox> y\<close>
+  using all_restriction_related_iff_related by blast
+
+lemma empty_restriction_not_related_set_iff: \<open>restriction_not_related_set x y = {} \<longleftrightarrow> x \<lessapprox> y\<close>
+  by (simp add: all_restriction_related_iff_related)
+
+
+lemma finite_restriction_related_set_iff :
+  \<open>finite (restriction_related_set x y) \<longleftrightarrow> \<not> x \<lessapprox> y\<close>
+proof (rule iffI)
+  assume \<open>finite (restriction_related_set x y)\<close>
+  obtain n where \<open>\<not> x \<down> n \<lessapprox> y \<down> n\<close> 
+    using \<open>finite (restriction_related_set x y)\<close> by fastforce
+  with mono_restriction_related show \<open>\<not> x \<lessapprox> y\<close> by blast
+next
+  assume \<open>\<not> x \<lessapprox> y\<close>
+  then obtain n where \<open>\<forall>m>n. \<not> x \<down> m \<lessapprox> y \<down> m\<close>
+    by (meson all_restriction_related_iff_related less_le_not_le restriction_related_le)
+  hence \<open>restriction_related_set x y \<subseteq> {0..n}\<close>
+    by (simp add: subset_iff) (meson linorder_not_le)
+  thus \<open>finite (restriction_related_set x y)\<close>
+    by (simp add: subset_eq_atLeast0_atMost_finite)
+qed
+
+lemma infinite_restriction_not_related_set_iff :
+  \<open>infinite (restriction_not_related_set x y) \<longleftrightarrow> \<not> x \<lessapprox> y\<close>
+  by (metis empty_restriction_not_related_set_iff finite_restriction_related_set_iff
+      finite.emptyI finite_Collect_not infinite_UNIV_char_0)
+
+lemma bdd_above_restriction_related_set_iff :
+  \<open>bdd_above (restriction_related_set x y) \<longleftrightarrow> \<not> x \<lessapprox> y\<close>
+  by (simp add: bdd_above_nat finite_restriction_related_set_iff)
+
+
+
+context fixes x y assumes \<open>\<not> x \<lessapprox> y\<close> begin
+
+lemma Sup_in_restriction_related_set :
+  \<open>Sup (restriction_related_set x y) \<in> restriction_related_set x y\<close>
+  using Max_in[OF finite_restriction_related_set_iff[THEN iffD2, OF \<open>\<not> x \<lessapprox> y\<close>]
+      nonempty_restriction_related_set]
+    cSup_eq_Max[OF finite_restriction_related_set_iff[THEN iffD2, OF \<open>\<not> x \<lessapprox> y\<close>]
+      nonempty_restriction_related_set]
+  by argo
+
+lemma Inf_in_restriction_not_related_set :
+  \<open>Inf (restriction_not_related_set x y) \<in> restriction_not_related_set x y\<close>
+  by (metis \<open>\<not> x \<lessapprox> y\<close> Inf_nat_def1 finite.emptyI infinite_restriction_not_related_set_iff)
+
+lemma Inf_restriction_not_related_set_eq_Suc_Sup_restriction_related_set :
+  \<open>Inf (restriction_not_related_set x y) = Suc (Sup (restriction_related_set x y))\<close>
+proof -
+  let ?S_eq  = \<open>restriction_related_set  x y\<close>
+  let ?S_neq = \<open>restriction_not_related_set x y\<close>
+  from Inf_in_restriction_not_related_set have \<open>Inf ?S_neq \<in> ?S_neq\<close> by blast
+  from Sup_in_restriction_related_set have \<open>Sup ?S_eq \<in> ?S_eq\<close> by blast
+  hence \<open>Suc (Sup ?S_eq) \<notin> ?S_eq\<close>
+    by (metis Suc_n_not_le_n \<open>\<not> x \<lessapprox> y\<close> bdd_above_restriction_related_set_iff cSup_upper)
+  with restriction_related_set_Un_restriction_not_related_set
+  have \<open>Suc (Sup ?S_eq) \<in> ?S_neq\<close> by auto
+  show \<open>Inf ?S_neq = Suc (Sup ?S_eq)\<close>
+  proof (rule order_antisym)
+    show \<open>Inf ?S_neq \<le> Suc (Sup ?S_eq)\<close>
+      by (fact wellorder_Inf_le1[OF \<open>Suc (Sup ?S_eq) \<in> ?S_neq\<close>])
+  next
+    from \<open>Inf ?S_neq \<in> ?S_neq\<close> \<open>Sup ?S_eq \<in> ?S_eq\<close> show \<open>Suc (Sup ?S_eq) \<le> Inf ?S_neq\<close>
+      by (metis (mono_tags, lifting) mem_Collect_eq not_less_eq_eq restriction_related_le)
+  qed
+qed
+
+end
+
+
+lemma restriction_related_set_is_atMost :
+  \<open>restriction_related_set x y =
+   (if x \<lessapprox> y then UNIV else {..Sup (restriction_related_set x y)})\<close>
+proof (split if_split, intro conjI impI)
+  show \<open>x \<lessapprox> y \<Longrightarrow> restriction_related_set x y = UNIV\<close> 
+    by (simp add: UNIV_restriction_related_set_iff)
+next
+  assume \<open>\<not> x \<lessapprox> y\<close>
+  hence * : \<open>Sup (restriction_related_set x y) \<in> restriction_related_set x y\<close>
+    by (fact Sup_in_restriction_related_set)
+  show \<open>restriction_related_set x y = {..Sup (restriction_related_set x y)} \<close>
+  proof (intro subset_antisym subsetI)
+    show \<open>n \<in> restriction_related_set x y \<Longrightarrow> n \<in> {..Sup (restriction_related_set x y)}\<close> for n
+      by (simp add: \<open>\<not> x \<lessapprox> y\<close> finite_restriction_related_set_iff le_cSup_finite)
+  next
+    from "*" show \<open>n \<in> {..Sup (restriction_related_set x y)} \<Longrightarrow>
+                   n \<in> restriction_related_set x y\<close> for n
+      by simp (meson mem_Collect_eq restriction_related_le)
+  qed
+qed
+
+lemma restriction_not_related_set_is_atLeast :
+  \<open>restriction_not_related_set x y =
+   (if x \<lessapprox> y then {} else {Inf (restriction_not_related_set x y)..})\<close>
+proof (split if_split, intro conjI impI)
+  from empty_restriction_not_related_set_iff
+  show \<open>x \<lessapprox> y \<Longrightarrow> restriction_not_related_set x y = {}\<close> by blast
+next
+  assume \<open>\<not> x \<lessapprox> y\<close>
+  have \<open>restriction_not_related_set x y = UNIV - restriction_related_set x y\<close> by auto
+  also have \<open>\<dots> = UNIV - {..Sup (restriction_related_set x y)}\<close>
+    by (subst restriction_related_set_is_atMost) (simp add: \<open>\<not> x \<lessapprox> y\<close>)
+  also have \<open>\<dots> = {Suc (Sup (restriction_related_set x y))..}\<close> by auto
+  also have \<open>Suc (Sup (restriction_related_set x y)) = Inf (restriction_not_related_set x y)\<close>
+    by (simp add: \<open>\<not> x \<lessapprox> y\<close> flip: Inf_restriction_not_related_set_eq_Suc_Sup_restriction_related_set)
+  finally show \<open>restriction_not_related_set x y = {Inf (restriction_not_related_set x y)..}\<close> .
+qed
+
+
+end
+
+
+
+
+subsection \<open>Restriction shift Maps\<close>
+
+locale Restriction_2_PreorderRestrictionSpace =
+  R1 : Restriction \<open>(\<down>1)\<close> \<open>(\<lessapprox>1)\<close> +
+  PRS2 : PreorderRestrictionSpace \<open>(\<down>2)\<close> \<open>(\<lessapprox>2)\<close>
+  for restriction1 :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a\<close>  (infixl \<open>\<down>1\<close> 60)
+    and relation1    :: \<open>'a \<Rightarrow> 'a \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>1\<close> 50)
+    and restriction2 :: \<open>'b \<Rightarrow> nat \<Rightarrow> 'b\<close>  (infixl \<open>\<down>2\<close> 60)
+    and relation2    :: \<open>'b \<Rightarrow> 'b \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>2\<close> 50)
+begin
+
+subsubsection \<open>Definition\<close>
+
+text \<open>This notion is a generalization of constructive map and non-destructive map.\<close>
+
+definition restriction_shift_on :: \<open>['a \<Rightarrow> 'b, int, 'a set] \<Rightarrow> bool\<close>
+  where \<open>restriction_shift_on f k A \<equiv>
+         \<forall>x\<in>A. \<forall>y\<in>A. \<forall>n. x \<down>1 n \<lessapprox>1 y \<down>1 n \<longrightarrow> f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k)\<close>
+
+definition restriction_shift :: \<open>['a \<Rightarrow> 'b, int] \<Rightarrow> bool\<close>
+  where \<open>restriction_shift f k \<equiv> restriction_shift_on f k UNIV\<close>
+
+
+lemma restriction_shift_onI :
+  \<open>(\<And>x y n. \<lbrakk>x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow>
+             f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k))
+   \<Longrightarrow> restriction_shift_on f k A\<close>
+  unfolding restriction_shift_on_def
+  by (metis PRS2.all_restriction_related_iff_related)
+
+corollary restriction_shiftI :
+  \<open>(\<And>x y n. \<lbrakk>\<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow>
+             f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k))
+   \<Longrightarrow> restriction_shift f k\<close>
+  by (unfold restriction_shift_def, rule restriction_shift_onI)
+
+
+
+lemma restriction_shift_onD :
+  \<open>\<lbrakk>restriction_shift_on f k A; x \<in> A; y \<in> A; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk>
+   \<Longrightarrow> f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k)\<close>
+  by (unfold restriction_shift_on_def) blast
+
+lemma restriction_shiftD :
+  \<open>\<lbrakk>restriction_shift f k; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k)\<close>
+  unfolding restriction_shift_def using restriction_shift_onD by blast
+
+lemma restriction_shift_on_subset :
+  \<open>restriction_shift_on f k B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> restriction_shift_on f k A\<close>
+  by (simp add: restriction_shift_on_def subset_iff)
+
+lemma restriction_shift_imp_restriction_shift_on [restriction_shift_simpset] :
+  \<open>restriction_shift f k \<Longrightarrow> restriction_shift_on f k A\<close>
+  unfolding restriction_shift_def using restriction_shift_on_subset by blast
+
+
+lemma restriction_shift_on_imp_restriction_shift_on_le [restriction_shift_simpset] :
+  \<open>restriction_shift_on f l A\<close> if \<open>l \<le> k\<close> and \<open>restriction_shift_on f k A\<close>
+proof (rule restriction_shift_onI)
+  fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down>1 n \<lessapprox>1 y \<down>1 n\<close>
+  from \<open>restriction_shift_on f k A\<close>[THEN restriction_shift_onD, OF this]
+  have \<open>f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k)\<close> .
+  moreover have \<open>nat (int n + l) \<le> nat (int n + k)\<close> by (simp add: nat_mono \<open>l \<le> k\<close>)
+  ultimately show \<open>f x \<down>2 nat (int n + l) \<lessapprox>2 f y \<down>2 nat (int n + l)\<close>
+    using PRS2.restriction_related_le by blast
+qed
+
+corollary restriction_shift_imp_restriction_shift_le [restriction_shift_simpset] :
+  \<open>l \<le> k \<Longrightarrow> restriction_shift f k \<Longrightarrow> restriction_shift f l\<close>
+  unfolding restriction_shift_def
+  by (fact restriction_shift_on_imp_restriction_shift_on_le)
+
+
+
+lemma restriction_shift_on_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> restriction_shift_on (f x) k A;
+    \<And>x. \<not> P x \<Longrightarrow> restriction_shift_on (g x) k A\<rbrakk> \<Longrightarrow>
+    restriction_shift_on (\<lambda>y. if P x then f x y else g x y) k A\<close>
+  by (rule restriction_shift_onI) (auto dest: restriction_shift_onD)
+
+corollary restriction_shift_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> restriction_shift (f x) k;
+    \<And>x. \<not> P x \<Longrightarrow> restriction_shift (g x) k\<rbrakk> \<Longrightarrow>
+    restriction_shift (\<lambda>y. if P x then f x y else g x y) k\<close>
+  unfolding restriction_shift_def by (fact restriction_shift_on_if_then_else)
+
+
+
+subsubsection \<open>Three particular Cases\<close>
+
+text \<open>The shift is most often equal to term\<open>0 :: int\<close>, term\<open>1 :: int\<close> or term\<open>- 1 :: int\<close>.
+      We provide extra support in these three cases.\<close>
+
+paragraph \<open>Non-too-destructive Map\<close>
+
+definition non_too_destructive_on :: \<open>['a \<Rightarrow> 'b, 'a set] \<Rightarrow> bool\<close>
+  where \<open>non_too_destructive_on f A \<equiv> restriction_shift_on f (- 1) A\<close>
+
+definition non_too_destructive :: \<open>['a \<Rightarrow> 'b] \<Rightarrow> bool\<close>
+  where \<open>non_too_destructive f \<equiv> non_too_destructive_on f UNIV\<close>
+
+
+lemma non_too_destructive_onI :
+  \<open>non_too_destructive_on f A\<close>
+  if \<open>\<And>n x y. \<lbrakk>x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 Suc n \<lessapprox>1 y \<down>1 Suc n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<close>
+proof (unfold non_too_destructive_on_def, rule restriction_shift_onI)
+  fix x y n
+  show \<open>\<lbrakk>x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk>
+        \<Longrightarrow> f x \<down>2 nat (int n + - 1) \<lessapprox>2 f y \<down>2 nat (int n + - 1)\<close>
+    by (cases \<open>n < 1\<close>) (simp_all add: Suc_nat_eq_nat_zadd1 that)
+qed
+
+lemma non_too_destructiveI :
+  \<open>\<lbrakk>\<And>n x y. \<lbrakk>\<not> f x \<lessapprox>2 f y; x \<down>1 Suc n \<lessapprox>1 y \<down>1 Suc n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<rbrakk>
+   \<Longrightarrow> non_too_destructive f\<close>
+  by (unfold non_too_destructive_def, rule non_too_destructive_onI)
+
+lemma non_too_destructive_onD :
+  \<open>\<lbrakk>non_too_destructive_on f A; x \<in> A; y \<in> A; x \<down>1 Suc n \<lessapprox>1 y \<down>1 Suc n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<close>
+  unfolding non_too_destructive_on_def using restriction_shift_onD by fastforce
+
+lemma non_too_destructiveD :
+  \<open>\<lbrakk>non_too_destructive f; x \<down>1 Suc n \<lessapprox>1 y \<down>1 Suc n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<close>
+  unfolding non_too_destructive_def using non_too_destructive_onD by simp
+
+lemma non_too_destructive_on_subset :
+  \<open>non_too_destructive_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> non_too_destructive_on f A\<close>
+  by (meson non_too_destructive_on_def restriction_shift_on_subset)
+
+lemma non_too_destructive_imp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>non_too_destructive f \<Longrightarrow> non_too_destructive_on f A\<close>
+  unfolding non_too_destructive_def using non_too_destructive_on_subset by auto
+
+
+corollary non_too_destructive_on_if_then_else [restriction_shift_simpset, restriction_shift_introset]:
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> non_too_destructive_on (f x) A; \<And>x. \<not> P x \<Longrightarrow> non_too_destructive_on (g x) A\<rbrakk>
+     \<Longrightarrow> non_too_destructive_on (\<lambda>y. if P x then f x y else g x y) A\<close>
+  and non_too_destructive_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> non_too_destructive (f x); \<And>x. \<not> P x \<Longrightarrow> non_too_destructive (g x)\<rbrakk>
+     \<Longrightarrow> non_too_destructive (\<lambda>y. if P x then f x y else g x y)\<close>
+  by (auto simp add: non_too_destructive_def non_too_destructive_on_def
+      intro: restriction_shift_on_if_then_else)
+
+
+
+paragraph \<open>Non-destructive Map\<close>
+
+definition non_destructive_on :: \<open>['a \<Rightarrow> 'b, 'a set] \<Rightarrow> bool\<close>
+  where \<open>non_destructive_on f A \<equiv> restriction_shift_on f 0 A\<close>
+
+definition non_destructive :: \<open>['a \<Rightarrow> 'b] \<Rightarrow> bool\<close>
+  where \<open>non_destructive f \<equiv> non_destructive_on f UNIV\<close>
+
+
+lemma non_destructive_onI :
+  \<open>\<lbrakk>\<And>n x y. \<lbrakk>n \<noteq> 0; x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<rbrakk>
+   \<Longrightarrow> non_destructive_on f A\<close>
+  by (unfold non_destructive_on_def, rule restriction_shift_onI)
+    (metis PRS2.restriction_0_related add.right_neutral nat_int)
+
+lemma non_destructiveI :
+  \<open>\<lbrakk>\<And>n x y. \<lbrakk>n \<noteq> 0; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<rbrakk>
+   \<Longrightarrow> non_destructive f\<close> by (unfold non_destructive_def, rule non_destructive_onI)
+
+lemma non_destructive_onD :
+  \<open>\<lbrakk>non_destructive_on f A; x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<close>
+  by (simp add: non_destructive_on_def restriction_shift_on_def)
+
+lemma non_destructiveD : \<open>\<lbrakk>non_destructive f; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 n \<lessapprox>2 f y \<down>2 n\<close>
+  by (simp add: non_destructive_def non_destructive_on_def restriction_shift_on_def)
+
+lemma non_destructive_on_subset :
+  \<open>non_destructive_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> non_destructive_on f A\<close>
+  by (meson non_destructive_on_def restriction_shift_on_subset)
+
+lemma non_destructive_imp_non_destructive_on [restriction_shift_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_destructive_on f A\<close>
+  unfolding non_destructive_def using non_destructive_on_subset by auto
+
+lemma non_destructive_on_imp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>non_destructive_on f A \<Longrightarrow> non_too_destructive_on f A\<close>
+  unfolding non_destructive_on_def non_too_destructive_on_def
+  by (rule restriction_shift_on_imp_restriction_shift_on_le[of \<open>- 1\<close> 0 f A, simplified])
+
+corollary non_destructive_imp_non_too_destructive [restriction_shift_simpset] :
+  \<open>non_destructive f \<Longrightarrow> non_too_destructive f\<close>
+  by (unfold non_destructive_def non_too_destructive_def)
+    (fact non_destructive_on_imp_non_too_destructive_on)
+
+
+corollary non_destructive_on_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> non_destructive_on (f x) A; \<And>x. \<not> P x \<Longrightarrow> non_destructive_on (g x) A\<rbrakk>
+     \<Longrightarrow> non_destructive_on (\<lambda>y. if P x then f x y else g x y) A\<close>
+  and non_destructive_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> non_destructive (f x); \<And>x. \<not> P x \<Longrightarrow> non_destructive (g x)\<rbrakk>
+     \<Longrightarrow> non_destructive (\<lambda>y. if P x then f x y else g x y)\<close>
+  by (auto simp add: non_destructive_def non_destructive_on_def
+      intro: restriction_shift_on_if_then_else)
+
+
+
+paragraph \<open>Constructive Map\<close>
+
+definition constructive_on :: \<open>['a \<Rightarrow> 'b, 'a set] \<Rightarrow> bool\<close>
+  where \<open>constructive_on f A \<equiv> restriction_shift_on f 1 A\<close>
+
+definition constructive :: \<open>['a \<Rightarrow> 'b] \<Rightarrow> bool\<close>
+  where \<open>constructive f \<equiv> constructive_on f UNIV\<close>
+
+
+lemma constructive_onI :
+  \<open>\<lbrakk>\<And>n x y. \<lbrakk>x \<in> A; y \<in> A; \<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 Suc n \<lessapprox>2 f y \<down>2 Suc n\<rbrakk>
+   \<Longrightarrow> constructive_on f A\<close>
+  by (simp add: Suc_as_int constructive_on_def restriction_shift_onI)
+
+lemma constructiveI :
+  \<open>\<lbrakk>\<And>n x y. \<lbrakk>\<not> f x \<lessapprox>2 f y; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 Suc n \<lessapprox>2 f y \<down>2 Suc n\<rbrakk>
+   \<Longrightarrow> constructive f\<close> by (unfold constructive_def, rule constructive_onI)
+
+lemma constructive_onD :
+  \<open>\<lbrakk>constructive_on f A; x \<in> A; y \<in> A; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 Suc n \<lessapprox>2 f y \<down>2 Suc n\<close>
+  unfolding constructive_on_def by (metis Suc_as_int restriction_shift_onD)
+
+lemma constructiveD : \<open>\<lbrakk>constructive f; x \<down>1 n \<lessapprox>1 y \<down>1 n\<rbrakk> \<Longrightarrow> f x \<down>2 Suc n \<lessapprox>2 f y \<down>2 Suc n\<close>
+  unfolding constructive_def using constructive_onD by blast
+
+lemma constructive_on_subset :
+  \<open>constructive_on f B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> constructive_on f A\<close>
+  by (meson constructive_on_def restriction_shift_on_subset)
+
+lemma constructive_imp_constructive_on [restriction_shift_simpset] :
+  \<open>constructive f \<Longrightarrow> constructive_on f A\<close>
+  unfolding constructive_def using constructive_on_subset by auto
+
+
+lemma constructive_on_imp_non_destructive_on [restriction_shift_simpset] :
+  \<open>constructive_on f A \<Longrightarrow> non_destructive_on f A\<close>
+  by (rule non_destructive_onI)
+    (meson PRS2.restriction_related_pred constructive_onD)
+
+corollary constructive_imp_non_destructive [restriction_shift_simpset] :
+  \<open>constructive f \<Longrightarrow> non_destructive f\<close>
+  unfolding constructive_def non_destructive_def
+  by (fact constructive_on_imp_non_destructive_on)
+
+
+corollary constructive_on_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> constructive_on (f x) A; \<And>x. \<not> P x \<Longrightarrow> constructive_on (g x) A\<rbrakk>
+     \<Longrightarrow> constructive_on (\<lambda>y. if P x then f x y else g x y) A\<close>
+  and constructive_if_then_else [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> constructive (f x); \<And>x. \<not> P x \<Longrightarrow> constructive (g x)\<rbrakk>
+     \<Longrightarrow> constructive (\<lambda>y. if P x then f x y else g x y)\<close>
+  by (auto simp add: constructive_def constructive_on_def
+      intro: restriction_shift_on_if_then_else)
+
+
+end
+
+
+
+subsubsection \<open>Properties\<close>
+
+locale PreorderRestrictionSpace_2_PreorderRestrictionSpace =
+  PRS1 : PreorderRestrictionSpace \<open>(\<down>1)\<close> \<open>(\<lessapprox>1)\<close> +
+  PRS2 : PreorderRestrictionSpace \<open>(\<down>2)\<close> \<open>(\<lessapprox>2)\<close>
+  for restriction1 :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a\<close>  (infixl \<open>\<down>1\<close> 60)
+    and relation1    :: \<open>'a \<Rightarrow> 'a \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>1\<close> 50)
+    and restriction2 :: \<open>'b \<Rightarrow> nat \<Rightarrow> 'b\<close>  (infixl \<open>\<down>2\<close> 60)
+    and relation2    :: \<open>'b \<Rightarrow> 'b \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>2\<close> 50)
+begin
+
+sublocale Restriction_2_PreorderRestrictionSpace by unfold_locales
+
+lemma restriction_shift_on_restriction_restriction :
+  \<open>f (x \<down>1 n) \<down>2 nat (int n + k) \<lessapprox>2 f x \<down>2 nat (int n + k)\<close>
+  if \<open>restriction_shift_on f k A\<close> \<open>x \<down>1 n \<in> A\<close> \<open>x \<in> A\<close> \<open>x \<down>1 n \<lessapprox>1 x \<down>1 n\<close>
+    \<comment>\<open>the last assumption is trivial if term\<open>(\<lessapprox>1)\<close> is reflexive\<close>
+  by (rule restriction_shift_onD
+      [OF \<open>restriction_shift_on f k A\<close> \<open>x \<down>1 n \<in> A\<close> \<open>x \<in> A\<close>])
+    (simp add: \<open>x \<down>1 n \<lessapprox>1 x \<down>1 n\<close>)
+
+corollary restriction_shift_restriction_restriction :
+  \<open>f (x \<down>1 n) \<down>2 nat (int n + k) \<lessapprox>2 f x \<down>2 nat (int n + k)\<close>
+  if \<open>restriction_shift f k\<close> and \<open>x \<down>1 n \<lessapprox>1 x \<down>1 n\<close>
+  by (rule restriction_shiftD[OF \<open>restriction_shift f k\<close>])
+    (simp add: \<open>x \<down>1 n \<lessapprox>1 x \<down>1 n\<close>)
+
+
+corollary constructive_on_restriction_restriction :
+  \<open>\<lbrakk>constructive_on f A; x \<down>1 n \<in> A; x \<in> A; x \<down>1 n \<lessapprox>1 x \<down>1 n\<rbrakk>
+   \<Longrightarrow> f (x \<down>1 n) \<down>2 Suc n \<lessapprox>2 f x \<down>2 Suc n\<close>
+  using restriction_shift_on_restriction_restriction
+    restriction_shift_restriction_restriction Suc_as_int
+  unfolding constructive_on_def by presburger
+
+corollary constructive_restriction_restriction :
+  \<open>constructive f \<Longrightarrow> x \<down>1 n \<lessapprox>1 x \<down>1 n \<Longrightarrow> f (x \<down>1 n) \<down>2 Suc n \<lessapprox>2 f x \<down>2 Suc n\<close>
+  by (simp add: constructive_def constructive_on_restriction_restriction)
+
+
+corollary non_destructive_on_restriction_restriction :
+  \<open>\<lbrakk>non_destructive_on f A; x \<down>1 n \<in> A; x \<in> A; x \<down>1 n \<lessapprox>1 x \<down>1 n\<rbrakk>
+   \<Longrightarrow> f (x \<down>1 n) \<down>2 n \<lessapprox>2 f x \<down>2 n\<close>
+  using restriction_shift_on_restriction_restriction
+    restriction_shift_restriction_restriction
+  unfolding non_destructive_on_def by (metis add.commute add_0 nat_int)
+
+corollary non_destructive_restriction_restriction :
+  \<open>non_destructive f \<Longrightarrow> x \<down>1 n \<lessapprox>1 x \<down>1 n \<Longrightarrow> f (x \<down>1 n) \<down>2 n \<lessapprox>2 f x \<down>2 n\<close>
+  by (simp add: non_destructive_def non_destructive_on_restriction_restriction)
+
+
+corollary non_too_destructive_on_restriction_restriction :
+  \<open>\<lbrakk>non_too_destructive_on f A; x \<down>1 Suc n \<in> A; x \<in> A; x \<down>1 Suc n \<lessapprox>1 x \<down>1 Suc n\<rbrakk>
+   \<Longrightarrow> f (x \<down>1 Suc n) \<down>2 n \<lessapprox>2 f x \<down>2 n\<close>
+  using restriction_shift_on_restriction_restriction
+    restriction_shift_restriction_restriction
+  unfolding non_too_destructive_on_def by fastforce
+
+corollary non_too_destructive_restriction_restriction :
+  \<open>non_too_destructive f \<Longrightarrow> x \<down>1 Suc n \<lessapprox>1 x \<down>1 Suc n \<Longrightarrow> f (x \<down>1 Suc n) \<down>2 n \<lessapprox>2 f x \<down>2 n\<close>
+  by (simp add: non_too_destructive_def non_too_destructive_on_restriction_restriction)
+
+
+end
+
+
+locale Restriction_2_PreorderRestrictionSpace_2_PreorderRestrictionSpace =
+  R2PRS1 : Restriction_2_PreorderRestrictionSpace \<open>(\<down>1)\<close> \<open>(\<lessapprox>1)\<close> \<open>(\<down>2)\<close> \<open>(\<lessapprox>2)\<close> +
+  PRS2 : PreorderRestrictionSpace \<open>(\<down>3)\<close> \<open>(\<lessapprox>3)\<close>
+  for restriction1 :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a\<close>  (infixl \<open>\<down>1\<close> 60)
+    and relation1    :: \<open>'a \<Rightarrow> 'a \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>1\<close> 50)
+    and restriction2 :: \<open>'b \<Rightarrow> nat \<Rightarrow> 'b\<close>  (infixl \<open>\<down>2\<close> 60)
+    and relation2    :: \<open>'b \<Rightarrow> 'b \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>2\<close> 50)
+    and restriction3 :: \<open>'c \<Rightarrow> nat \<Rightarrow> 'c\<close>  (infixl \<open>\<down>3\<close> 60)
+    and relation3    :: \<open>'c \<Rightarrow> 'c \<Rightarrow> bool\<close> (infixl \<open>\<lessapprox>3\<close> 50)
+begin
+
+interpretation R2PRS2 : Restriction_2_PreorderRestrictionSpace \<open>(\<down>1)\<close> \<open>(\<lessapprox>1)\<close> \<open>(\<down>3)\<close> \<open>(\<lessapprox>3)\<close>
+  by unfold_locales
+
+interpretation PRS2PRS3 : PreorderRestrictionSpace_2_PreorderRestrictionSpace \<open>(\<down>2)\<close> \<open>(\<lessapprox>2)\<close> \<open>(\<down>3)\<close> \<open>(\<lessapprox>3)\<close>
+  by unfold_locales
+
+
+
+theorem restriction_shift_on_comp_restriction_shift_on [restriction_shift_simpset] :
+  \<open>R2PRS2.restriction_shift_on (\<lambda>x. g (f x)) (k + l) A\<close>
+  if \<open>f ` A \<subseteq> B\<close> \<open>PRS2PRS3.restriction_shift_on g l B\<close> \<open>R2PRS1.restriction_shift_on f k A\<close>
+proof (rule R2PRS2.restriction_shift_onI)
+  fix x y n assume \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>x \<down>1 n \<lessapprox>1 y \<down>1 n\<close>
+  from \<open>R2PRS1.restriction_shift_on f k A\<close>[THEN R2PRS1.restriction_shift_onD, OF this]
+  have \<open>f x \<down>2 nat (int n + k) \<lessapprox>2 f y \<down>2 nat (int n + k)\<close> .
+  moreover from \<open>x \<in> A\<close> \<open>y \<in> A\<close> \<open>f ` A \<subseteq> B\<close> have \<open>f x \<in> B\<close> \<open>f y \<in> B\<close> by auto
+  ultimately have \<open>g (f x) \<down>3 nat (int (nat (int n + k)) + l) \<lessapprox>3
+                   g (f y) \<down>3 nat (int (nat (int n + k)) + l)\<close>
+    using \<open>PRS2PRS3.restriction_shift_on g l B\<close>[THEN PRS2PRS3.restriction_shift_onD] by blast
+  moreover have \<open>nat (int n + (k + l)) \<le> nat (int (nat (int n + k)) + l)\<close>
+    by (simp add: nat_mono)
+  ultimately show \<open>g (f x) \<down>3 nat (int n + (k + l)) \<lessapprox>3 g (f y) \<down>3 nat (int n + (k + l))\<close>
+    by (metis PRS2.restriction_related_le)
+qed
+
+corollary restriction_shift_comp_restriction_shift_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.restriction_shift g l \<Longrightarrow> R2PRS1.restriction_shift_on f k A \<Longrightarrow>
+   R2PRS2.restriction_shift_on (\<lambda>x. g (f x)) (k + l) A\<close>
+  using PRS2PRS3.restriction_shift_imp_restriction_shift_on
+    restriction_shift_on_comp_restriction_shift_on by blast
+
+corollary restriction_shift_comp_restriction_shift [restriction_shift_simpset] :
+  \<open>PRS2PRS3.restriction_shift g l \<Longrightarrow> R2PRS1.restriction_shift f k \<Longrightarrow>
+   R2PRS2.restriction_shift (\<lambda>x. g (f x)) (k + l)\<close>
+  by (simp add: R2PRS1.restriction_shift_imp_restriction_shift_on
+      R2PRS2.restriction_shift_def restriction_shift_comp_restriction_shift_on)
+
+
+
+corollary non_destructive_on_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.non_destructive_on g B; R2PRS1.non_destructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: R2PRS1.non_destructive_on_def R2PRS2.non_destructive_on_def
+      R2PRS2.restriction_shift_on_def R2PRS1.restriction_shift_on_def)
+    (meson PRS2.mono_restriction_related PRS2PRS3.non_destructive_onD image_subset_iff)
+
+corollary non_destructive_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.non_destructive_on f A \<Longrightarrow>
+   R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: PRS2PRS3.non_destructiveD R2PRS1.non_destructive_on_def
+      R2PRS2.non_destructive_on_def R2PRS2.restriction_shift_on_def
+      R2PRS1.restriction_shift_on_def)
+
+corollary non_destructive_comp_non_destructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.non_destructive f \<Longrightarrow>
+   R2PRS2.non_destructive (\<lambda>x. g (f x))\<close>
+  by (simp add: PRS2PRS3.non_destructiveD R2PRS1.non_destructiveD
+      R2PRS2.non_destructive_def R2PRS2.non_destructive_onI)
+
+
+corollary constructive_on_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.constructive_on g B; R2PRS1.non_destructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.constructive_on (\<lambda>x. g (f x)) A\<close>
+  by (metis PRS2PRS3.constructive_on_def R2PRS1.non_destructive_on_def
+      R2PRS2.constructive_on_def add.commute add_cancel_left_right
+      restriction_shift_on_comp_restriction_shift_on)
+
+corollary constructive_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.constructive g \<Longrightarrow> R2PRS1.non_destructive_on f A \<Longrightarrow>
+   R2PRS2.constructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: R2PRS1.Restriction_2_PreorderRestrictionSpace_axioms
+      PRS2PRS3.constructiveD R2PRS1.non_destructive_on_def R2PRS2.constructive_onI
+      Restriction_2_PreorderRestrictionSpace.restriction_shift_on_def)
+
+corollary constructive_comp_non_destructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.constructive g \<Longrightarrow> R2PRS1.non_destructive f \<Longrightarrow>
+   R2PRS2.constructive (\<lambda>x. g (f x))\<close>
+  by (simp add: PRS2PRS3.constructiveD R2PRS1.non_destructiveD R2PRS2.constructiveI)
+
+
+corollary non_destructive_on_comp_constructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.non_destructive_on g B; R2PRS1.constructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.constructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: PRS2PRS3.non_destructive_onD R2PRS1.constructive_onD
+      R2PRS2.constructive_onI image_subset_iff)
+
+corollary non_destructive_comp_constructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.constructive_on f A \<Longrightarrow>
+   R2PRS2.constructive_on (\<lambda>x. g (f x)) A\<close>
+  using PRS2PRS3.non_destructive_def non_destructive_on_comp_constructive_on by blast
+
+corollary non_destructive_comp_constructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.constructive f \<Longrightarrow>
+   R2PRS2.constructive (\<lambda>x. g (f x))\<close>
+  using PRS2PRS3.non_destructiveD R2PRS1.constructiveD R2PRS2.constructiveI by presburger
+
+
+corollary non_too_destructive_on_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.non_too_destructive_on g B; R2PRS1.non_destructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.non_too_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (metis PRS2PRS3.non_too_destructive_on_def R2PRS1.non_destructive_on_def
+      R2PRS2.non_too_destructive_on_def add.commute
+      add.right_neutral restriction_shift_on_comp_restriction_shift_on)
+
+corollary non_too_destructive_comp_non_destructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_too_destructive g \<Longrightarrow> R2PRS1.non_destructive_on f A \<Longrightarrow>
+   R2PRS2.non_too_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (metis PRS2PRS3.non_too_destructive_imp_non_too_destructive_on
+      non_too_destructive_on_comp_non_destructive_on top_greatest)
+
+corollary non_too_destructive_comp_non_destructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_too_destructive g \<Longrightarrow> R2PRS1.non_destructive f \<Longrightarrow>
+   R2PRS2.non_too_destructive (\<lambda>x. g (f x))\<close>
+  by (simp add: PRS2PRS3.non_too_destructiveD R2PRS1.non_destructiveD R2PRS2.non_too_destructiveI)
+
+
+corollary non_destructive_on_comp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.non_destructive_on g B; R2PRS1.non_too_destructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.non_too_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: PRS2PRS3.non_destructive_onD R2PRS1.non_too_destructive_onD
+      R2PRS2.non_too_destructive_onI image_subset_iff)
+
+corollary non_destructive_comp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.non_too_destructive_on f A \<Longrightarrow>
+   R2PRS2.non_too_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (simp add: PRS2PRS3.non_destructiveD R2PRS1.non_too_destructive_onD R2PRS2.non_too_destructive_onI)
+
+corollary non_destructive_comp_non_too_destructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_destructive g \<Longrightarrow> R2PRS1.non_too_destructive f \<Longrightarrow>
+   R2PRS2.non_too_destructive (\<lambda>x. g (f x))\<close>
+  using R2PRS1.non_too_destructive_imp_non_too_destructive_on R2PRS2.non_too_destructive_def
+    non_destructive_comp_non_too_destructive_on by blast
+
+
+corollary non_too_destructive_on_comp_constructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.non_too_destructive_on g B; R2PRS1.constructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  using restriction_shift_on_comp_restriction_shift_on[of f A B g \<open>- 1\<close> 1]
+  by (simp add: PRS2PRS3.non_too_destructive_on_def
+      R2PRS2.non_destructive_on_def R2PRS1.constructive_on_def )
+
+corollary non_too_destructive_comp_constructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_too_destructive g \<Longrightarrow> R2PRS1.constructive_on f A \<Longrightarrow>
+   R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  by (metis PRS2PRS3.non_too_destructive_imp_non_too_destructive_on
+      non_too_destructive_on_comp_constructive_on top_greatest)
+
+corollary non_too_destructive_comp_constructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.non_too_destructive g \<Longrightarrow> R2PRS1.constructive f \<Longrightarrow>
+   R2PRS2.non_destructive (\<lambda>x. g (f x))\<close>
+  by (simp add: PRS2PRS3.non_too_destructiveD R2PRS1.constructiveD R2PRS2.non_destructiveI)
+
+
+corollary constructive_on_comp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>\<lbrakk>f ` A \<subseteq> B; PRS2PRS3.constructive_on g B; R2PRS1.non_too_destructive_on f A\<rbrakk> \<Longrightarrow>
+    R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  using restriction_shift_on_comp_restriction_shift_on[of f A B g 1 \<open>- 1\<close>] 
+  by (simp add: R2PRS1.non_too_destructive_on_def
+      PRS2PRS3.constructive_on_def R2PRS2.non_destructive_on_def)
+
+corollary constructive_comp_non_too_destructive_on [restriction_shift_simpset] :
+  \<open>PRS2PRS3.constructive g \<Longrightarrow> R2PRS1.non_too_destructive_on f A \<Longrightarrow>
+   R2PRS2.non_destructive_on (\<lambda>x. g (f x)) A\<close>
+  using PRS2PRS3.constructive_imp_constructive_on constructive_on_comp_non_too_destructive_on by blast
+
+corollary constructive_comp_non_too_destructive [restriction_shift_simpset] :
+  \<open>PRS2PRS3.constructive g \<Longrightarrow> R2PRS1.non_too_destructive f \<Longrightarrow>
+   R2PRS2.non_destructive (\<lambda>x. g (f x))\<close>
+  by (metis R2PRS1.non_too_destructive_imp_non_too_destructive_on R2PRS2.non_destructiveI
+      R2PRS2.non_destructive_onD UNIV_I constructive_comp_non_too_destructive_on)
+
+
+
+end
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Prod.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Prod.thy
@@ -0,0 +1,438 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Product over Restriction Spaces\<close>
+
+(*<*)
+theory Restriction_Spaces_Prod
+  imports Restriction_Spaces_Classes
+begin
+  (*>*)
+
+subsection \<open>Restriction Space\<close>
+
+instantiation prod :: (restriction, restriction) restriction
+begin
+
+definition restriction_prod :: \<open>'a \<times> 'b \<Rightarrow> nat \<Rightarrow> 'a \<times> 'b\<close>
+  where \<open>p \<down> n \<equiv> (fst p \<down> n, snd p \<down> n)\<close>
+
+instance by intro_classes (simp add: restriction_prod_def)
+
+end
+
+
+
+instance prod :: (restriction_space, restriction_space) restriction_space
+proof intro_classes
+  show \<open>p \<down> 0 = q \<down> 0\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (simp add: restriction_prod_def)
+next
+  show \<open>p \<noteq> q \<Longrightarrow> \<exists>n. p \<down> n \<noteq> q \<down> n\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (simp add: restriction_prod_def)
+      (meson ex_not_restriction_related prod.expand)
+qed
+
+
+
+instantiation prod :: (preorder_restriction_space, preorder_restriction_space) preorder_restriction_space
+begin
+
+text \<open>
+We might want to use lexicographic order :
+  item term\<open>p \<le> q \<equiv> fst p < fst q \<or> fst p = fst q \<and> snd p \<le> snd q\<close>
+  item term\<open>p < q \<equiv> fst p < fst q \<or> (fst p = fst q \<and> snd p < snd q)\<close>
+
+but this is wrong since it is incompatible with term\<open>p \<down> 0 \<le> q \<down> 0\<close>,
+term\<open>\<not> p \<le> q \<Longrightarrow> \<exists>n. \<not> p \<down> n \<le> q \<down> n\<close> and term\<open>p \<le> q \<Longrightarrow> p \<down> n \<le> q \<down> n\<close>.\<close>
+
+definition less_eq_prod :: \<open>'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool\<close>
+  where \<open>p \<le> q \<equiv> fst p \<le> fst q \<and> snd p \<le> snd q\<close>
+
+definition less_prod :: \<open>'a \<times> 'b \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool\<close>
+  where \<open>p < q \<equiv> fst p \<le> fst q \<and> snd p < snd q \<or> fst p < fst q \<and> snd p \<le> snd q\<close>
+
+instance
+proof intro_classes
+  show \<open>p < q \<longleftrightarrow> p \<le> q \<and> \<not> q \<le> p\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (auto simp add: less_eq_prod_def less_prod_def less_le_not_le)
+next
+  show \<open>p \<le> p\<close> for p :: \<open>'a \<times> 'b\<close>
+    by (simp add: less_eq_prod_def)
+next
+  show \<open>p \<le> q \<Longrightarrow> q \<le> r \<Longrightarrow> p \<le> r\<close> for p q r :: \<open>'a \<times> 'b\<close>
+    by (auto simp add: less_eq_prod_def intro: order_trans)
+next
+  show \<open>p \<down> 0 \<le> q \<down> 0\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (simp add: less_eq_prod_def restriction_prod_def)
+next
+  show \<open>p \<le> q \<Longrightarrow> p \<down> n \<le> q \<down> n\<close> for p q :: \<open>'a \<times> 'b\<close> and n
+    by (simp add: less_eq_prod_def restriction_prod_def mono_restriction_less_eq)
+next
+  show \<open>\<not> p \<le> q \<Longrightarrow> \<exists>n. \<not> p \<down> n \<le> q \<down> n\<close> for p q :: \<open>'a \<times> 'b\<close>
+    by (simp add: less_eq_prod_def restriction_prod_def)
+      (meson ex_not_restriction_less_eq)
+qed
+
+end
+
+
+instance prod :: (order_restriction_space, order_restriction_space) order_restriction_space
+  by intro_classes (simp add: less_eq_prod_def order_antisym prod.expand)
+
+
+
+subsection \<open>Restriction shift Maps\<close>
+
+subsubsection \<open>Domain is a Product\<close>
+
+lemma restriction_shift_on_prod_domain_iff : 
+  \<open>restriction_shift_on f k (A \<times> B) \<longleftrightarrow> (\<forall>x\<in>A. restriction_shift_on (\<lambda>y. f (x, y)) k B) \<and>
+                                         (\<forall>y\<in>B. restriction_shift_on (\<lambda>x. f (x, y)) k A)\<close>
+proof (intro iffI conjI ballI)
+  show \<open>restriction_shift_on (\<lambda>y. f (x, y)) k B\<close>
+    if \<open>restriction_shift_on f k (A \<times> B)\<close> and \<open>x \<in> A\<close> for x
+  proof (rule restriction_shift_onI)
+    show \<open>y1 \<in> B \<Longrightarrow> y2 \<in> B \<Longrightarrow> y1 \<down> n = y2 \<down> n \<Longrightarrow>
+          f (x, y1) \<down> nat (int n + k) = f (x, y2) \<down> nat (int n + k)\<close> for y1 y2 n
+      by (rule that(1)[THEN restriction_shift_onD])
+        (use that(2) in \<open>simp_all add: restriction_prod_def\<close>)
+  qed
+next
+  show \<open>restriction_shift_on (\<lambda>x. f (x, y)) k A\<close>
+    if \<open>restriction_shift_on f k (A \<times> B)\<close> and \<open>y \<in> B\<close> for y
+  proof (rule restriction_shift_onI)
+    show \<open>x1 \<in> A \<Longrightarrow> x2 \<in> A \<Longrightarrow> x1 \<down> n = x2 \<down> n \<Longrightarrow>
+          f (x1, y) \<down> nat (int n + k) = f (x2, y) \<down> nat (int n + k)\<close> for x1 x2 n
+      by (rule that(1)[THEN restriction_shift_onD])
+        (use that(2) in \<open>simp_all add: restriction_prod_def\<close>)
+  qed
+next
+  assume assm : \<open>(\<forall>x\<in>A. restriction_shift_on (\<lambda>y. f (x, y)) k B) \<and>
+                 (\<forall>y\<in>B. restriction_shift_on (\<lambda>x. f (x, y)) k A)\<close>
+  show \<open>restriction_shift_on f k (A \<times> B)\<close>
+  proof (rule restriction_shift_onI)
+    fix p q n assume \<open>p \<in> A \<times> B\<close> \<open>q \<in> A \<times> B\<close> \<open>p \<down> n = q \<down> n\<close>
+    then obtain x1 y1 x2 y2
+      where \<open>p = (x1, y1)\<close> \<open>q = (x2, y2)\<close> \<open>x1 \<in> A\<close> \<open>y1 \<in> B\<close>
+        \<open>x2 \<in> A\<close> \<open>y2 \<in> B\<close> \<open>x1 \<down> n = x2 \<down> n\<close> \<open>y1 \<down> n = y2 \<down> n\<close>
+      by (cases p, cases q, simp add: restriction_prod_def)
+    have \<open>p = (x1, y1)\<close> by (fact \<open>p = (x1, y1)\<close>)
+    also have \<open>f (x1, y1) \<down> nat (int n + k) = f (x1, y2) \<down> nat (int n + k)\<close>
+      by (rule assm[THEN conjunct1, rule_format, OF \<open>x1 \<in> A\<close>, THEN restriction_shift_onD])
+        (fact \<open>y1 \<in> B\<close> \<open>y2 \<in> B\<close> \<open>y1 \<down> n = y2 \<down> n\<close>)+
+    also have \<open>f (x1, y2) \<down> nat (int n + k) = f (x2, y2) \<down> nat (int n + k)\<close>
+      by (rule assm[THEN conjunct2, rule_format, OF \<open>y2 \<in> B\<close>, THEN restriction_shift_onD])
+        (fact \<open>x1 \<in> A\<close> \<open>x2 \<in> A\<close> \<open>x1 \<down> n = x2 \<down> n\<close>)+
+    also have \<open>(x2, y2) = q\<close> by (simp add: \<open>q = (x2, y2)\<close>)
+    finally show \<open>f p \<down> nat (int n + k) = f q \<down> nat (int n + k)\<close> .
+  qed
+qed
+
+lemma restriction_shift_prod_domain_iff:
+  \<open>restriction_shift f k \<longleftrightarrow> (\<forall>x. restriction_shift (\<lambda>y. f (x, y)) k) \<and>
+                             (\<forall>y. restriction_shift (\<lambda>x. f (x, y)) k)\<close>
+  unfolding restriction_shift_def
+  by (metis UNIV_I UNIV_Times_UNIV restriction_shift_on_prod_domain_iff)
+
+
+lemma non_too_destructive_on_prod_domain_iff :
+  \<open>non_too_destructive_on f (A \<times> B) \<longleftrightarrow> (\<forall>x\<in>A. non_too_destructive_on (\<lambda>y. f (x, y)) B) \<and>
+                                        (\<forall>y\<in>B. non_too_destructive_on (\<lambda>x. f (x, y)) A)\<close>
+  by (simp add: non_too_destructive_on_def restriction_shift_on_prod_domain_iff)
+
+lemma non_too_destructive_prod_domain_iff :
+  \<open>non_too_destructive f \<longleftrightarrow> (\<forall>x. non_too_destructive (\<lambda>y. f (x, y))) \<and>
+                             (\<forall>y. non_too_destructive (\<lambda>x. f (x, y)))\<close>
+  unfolding non_too_destructive_def
+  by (metis UNIV_I UNIV_Times_UNIV non_too_destructive_on_prod_domain_iff)
+
+
+lemma non_destructive_on_prod_domain_iff :
+  \<open>non_destructive_on f (A \<times> B) \<longleftrightarrow> (\<forall>x\<in>A. non_destructive_on (\<lambda>y. f (x, y)) B) \<and>
+                                    (\<forall>y\<in>B. non_destructive_on (\<lambda>x. f (x, y)) A)\<close>
+  by (simp add: non_destructive_on_def restriction_shift_on_prod_domain_iff)
+
+lemma non_destructive_prod_domain_iff :
+  \<open>non_destructive f \<longleftrightarrow> (\<forall>x. non_destructive (\<lambda>y. f (x, y))) \<and>
+                         (\<forall>y. non_destructive (\<lambda>x. f (x, y)))\<close>
+  unfolding non_destructive_def
+  by (metis UNIV_I UNIV_Times_UNIV non_destructive_on_prod_domain_iff)
+
+
+lemma constructive_on_prod_domain_iff :
+  \<open>constructive_on f (A \<times> B) \<longleftrightarrow> (\<forall>x\<in>A. constructive_on (\<lambda>y. f (x, y)) B) \<and>
+                                 (\<forall>y\<in>B. constructive_on (\<lambda>x. f (x, y)) A)\<close>
+  by (simp add: constructive_on_def restriction_shift_on_prod_domain_iff)
+
+lemma constructive_prod_domain_iff :
+  \<open>constructive f \<longleftrightarrow> (\<forall>x. constructive (\<lambda>y. f (x, y))) \<and>
+                      (\<forall>y. constructive (\<lambda>x. f (x, y)))\<close>
+  unfolding constructive_def
+  by (metis UNIV_I UNIV_Times_UNIV constructive_on_prod_domain_iff)
+
+
+
+lemma restriction_shift_prod_domain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. restriction_shift (\<lambda>y. f (x, y)) k;
+    \<And>y. restriction_shift (\<lambda>x. f (x, y)) k\<rbrakk> \<Longrightarrow> restriction_shift f k\<close>
+  and non_too_destructive_prod_domain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. non_too_destructive (\<lambda>y. f (x, y));
+    \<And>y. non_too_destructive (\<lambda>x. f (x, y))\<rbrakk> \<Longrightarrow> non_too_destructive f\<close>
+  and non_destructive_prod_domain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. non_destructive (\<lambda>y. f (x, y));
+    \<And>y. non_destructive (\<lambda>x. f (x, y))\<rbrakk> \<Longrightarrow> non_destructive f\<close>
+  and constructive_prod_domain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>\<And>x. constructive (\<lambda>y. f (x, y));
+    \<And>y. constructive (\<lambda>x. f (x, y))\<rbrakk> \<Longrightarrow> constructive f\<close>
+  by (simp_all add: restriction_shift_prod_domain_iff non_too_destructive_prod_domain_iff
+      non_destructive_prod_domain_iff constructive_prod_domain_iff)
+
+
+
+subsubsection \<open>Codomain is a Product\<close>
+
+lemma restriction_shift_on_prod_codomain_iff : 
+  \<open>restriction_shift_on f k A \<longleftrightarrow> (restriction_shift_on (\<lambda>x. fst (f x)) k A) \<and>
+                                  (restriction_shift_on (\<lambda>x. snd (f x)) k A)\<close>
+  by (simp add: restriction_shift_on_def restriction_prod_def) blast
+
+
+lemma restriction_shift_prod_codomain_iff:
+  \<open>restriction_shift f k \<longleftrightarrow> (restriction_shift (\<lambda>x. fst (f x)) k) \<and>
+                             (restriction_shift (\<lambda>x. snd (f x)) k)\<close>
+  unfolding restriction_shift_def
+  by (fact restriction_shift_on_prod_codomain_iff)
+
+
+lemma non_too_destructive_on_prod_codomain_iff :
+  \<open>non_too_destructive_on f A \<longleftrightarrow> (non_too_destructive_on (\<lambda>x. fst (f x)) A) \<and>
+                                  (non_too_destructive_on (\<lambda>x. snd (f x)) A)\<close>
+  by (simp add: non_too_destructive_on_def restriction_shift_on_prod_codomain_iff)
+
+lemma non_too_destructive_prod_codomain_iff :
+  \<open>non_too_destructive f \<longleftrightarrow> (non_too_destructive (\<lambda>x. fst (f x))) \<and>
+                             (non_too_destructive (\<lambda>x. snd (f x)))\<close>
+  by (simp add: non_too_destructive_def non_too_destructive_on_prod_codomain_iff)
+
+
+lemma non_destructive_on_prod_codomain_iff :
+  \<open>non_destructive_on f A \<longleftrightarrow> (non_destructive_on (\<lambda>x. fst (f x)) A) \<and>
+                              (non_destructive_on (\<lambda>x. snd (f x)) A)\<close>
+
+  by (simp add: non_destructive_on_def restriction_shift_on_prod_codomain_iff)
+
+lemma non_destructive_prod_codomain_iff :
+  \<open>non_destructive f \<longleftrightarrow> (non_destructive (\<lambda>x. fst (f x))) \<and>
+                         (non_destructive (\<lambda>x. snd (f x)))\<close>
+  by (simp add: non_destructive_def non_destructive_on_prod_codomain_iff)
+
+
+lemma constructive_on_prod_codomain_iff :
+  \<open>constructive_on f A \<longleftrightarrow> (constructive_on (\<lambda>x. fst (f x)) A) \<and>
+                           (constructive_on (\<lambda>x. snd (f x)) A)\<close>
+  by (simp add: constructive_on_def restriction_shift_on_prod_codomain_iff)
+
+lemma constructive_prod_codomain_iff :
+  \<open>constructive f \<longleftrightarrow> (constructive (\<lambda>x. fst (f x))) \<and>
+                      (constructive (\<lambda>x. snd (f x)))\<close>
+  by (simp add: constructive_def constructive_on_prod_codomain_iff)
+
+
+lemma restriction_shift_prod_codomain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>restriction_shift f k; restriction_shift g k\<rbrakk> \<Longrightarrow>
+   restriction_shift (\<lambda>x. (f x, g x)) k\<close>
+  and non_too_destructive_prod_codomain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>non_too_destructive f; non_too_destructive g\<rbrakk> \<Longrightarrow> non_too_destructive (\<lambda>x. (f x, g x))\<close>
+  and non_destructive_prod_codomain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>non_destructive f; non_destructive g\<rbrakk> \<Longrightarrow> non_destructive (\<lambda>x. (f x, g x))\<close>
+  and constructive_prod_codomain [restriction_shift_simpset, restriction_shift_introset] :
+  \<open>\<lbrakk>constructive f; constructive g\<rbrakk> \<Longrightarrow> constructive (\<lambda>x. (f x, g x))\<close>
+  by (simp_all add: restriction_shift_prod_codomain_iff non_too_destructive_prod_codomain_iff
+      non_destructive_prod_codomain_iff constructive_prod_codomain_iff)
+
+
+
+subsection \<open>Limits and Convergence\<close>
+
+lemma restriction_chain_prod_iff :
+  \<open>restriction_chain \<sigma> \<longleftrightarrow> restriction_chain (\<lambda>n. fst (\<sigma> n)) \<and>
+                           restriction_chain (\<lambda>n. snd (\<sigma> n))\<close>
+  by (simp add: restriction_chain_def restriction_prod_def)
+    (metis fst_conv prod.collapse snd_conv)
+
+lemma restriction_tendsto_prod_iff :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<lambda>n. fst (\<sigma> n)) \<midarrow>\<down>\<rightarrow> fst \<Sigma> \<and> (\<lambda>n. snd (\<sigma> n)) \<midarrow>\<down>\<rightarrow> snd \<Sigma>\<close>
+  by (simp add: restriction_tendsto_def restriction_prod_def)
+    (meson nle_le order_trans)
+
+lemma restriction_convergent_prod_iff :
+  \<open>restriction_convergent \<sigma> \<longleftrightarrow> restriction_convergent (\<lambda>n. fst (\<sigma> n)) \<and>
+                                restriction_convergent (\<lambda>n. snd (\<sigma> n))\<close>
+  by (simp add: restriction_convergent_def restriction_tendsto_prod_iff)
+
+lemma funpow_indep_prod_is :
+  \<open>((\<lambda>(x, y). (f x, g y)) ^^ n) (x, y) = ((f ^^ n) x, (g ^^ n) y)\<close>
+  for f g :: \<open>'a \<Rightarrow> 'a\<close>
+  by (induct n) simp_all
+
+
+
+subsection \<open>Completeness\<close>
+
+instance prod :: (complete_restriction_space, complete_restriction_space) complete_restriction_space
+proof intro_classes
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a :: complete_restriction_space \<times> 'b :: complete_restriction_space\<close>
+  assume \<open>chain\<down> \<sigma>\<close>
+  hence \<open>chain\<down> (\<lambda>n. fst (\<sigma> n))\<close> \<open>chain\<down> (\<lambda>n. snd (\<sigma> n))\<close>
+    by (simp_all add: restriction_chain_prod_iff)
+  hence \<open>convergent\<down> (\<lambda>n. fst (\<sigma> n))\<close> \<open>convergent\<down> (\<lambda>n. snd (\<sigma> n))\<close>
+    by (simp_all add: restriction_chain_imp_restriction_convergent)
+  thus \<open>convergent\<down> \<sigma>\<close>
+    by (simp add: restriction_convergent_prod_iff)
+qed    
+
+
+
+subsection \<open>Fixed Point\<close>
+
+lemma restriction_fix_indep_prod_is :
+  \<open>(\<upsilon> (x, y). (f x, g y)) = (\<upsilon> x. f x, \<upsilon> y. g y)\<close>
+  if contructive : \<open>constructive f\<close> \<open>constructive g\<close>
+  for f :: \<open>'a \<Rightarrow> 'a :: complete_restriction_space\<close>
+    and g :: \<open>'b \<Rightarrow> 'b :: complete_restriction_space\<close>
+proof (rule restriction_fix_unique)
+  from contructive show \<open>constructive (\<lambda>(x, y). (f x, g y))\<close>
+    by (simp add: constructive_prod_domain_iff constructive_prod_codomain_iff constructive_const)
+next
+  show \<open>(case (\<upsilon> x. f x, \<upsilon> y. g y) of (x, y) \<Rightarrow> (f x, g y)) = (\<upsilon> x. f x, \<upsilon> y. g y)\<close>
+    by simp (metis restriction_fix_eq contructive)
+qed
+
+
+
+lemma non_destructive_fst : \<open>non_destructive fst\<close>
+  by (rule non_destructiveI) (simp add: restriction_prod_def)
+
+lemma non_destructive_snd : \<open>non_destructive snd\<close>
+  by (rule non_destructiveI) (simp add: restriction_prod_def)
+
+
+
+
+lemma constructive_restriction_fix_right :
+  \<open>constructive (\<lambda>x. \<upsilon> y. f (x, y))\<close> if \<open>constructive f\<close>
+for f :: \<open>'a :: complete_restriction_space \<times> 'b :: complete_restriction_space \<Rightarrow> 'b\<close>
+proof (rule constructiveI)
+  fix n and x x' :: 'a assume \<open>x \<down> n = x' \<down> n\<close>
+  show \<open>(\<upsilon> y. f (x, y)) \<down> Suc n = (\<upsilon> y. f (x', y)) \<down> Suc n\<close>
+  proof (subst (1 2) restriction_restriction_fix_is)
+    show \<open>constructive (\<lambda>y. f (x', y))\<close> and \<open>constructive (\<lambda>y. f (x, y))\<close>
+      using \<open>constructive f\<close> constructive_prod_domain_iff by blast+
+  next
+    from \<open>x \<down> n = x' \<down> n\<close> show \<open>((\<lambda>y. f (x, y)) ^^ Suc n) undefined \<down> Suc n =
+                                ((\<lambda>y. f (x', y)) ^^ Suc n) undefined \<down> Suc n\<close>
+      by (induct n, simp_all)
+        (use \<open>constructive f\<close> constructiveD restriction_0_related in blast,
+          metis (no_types, lifting) \<open>constructive f\<close> constructiveD prod.sel
+          restriction_related_pred restriction_prod_def)
+  qed
+qed
+
+
+lemma constructive_restriction_fix_left :
+  \<open>constructive (\<lambda>y. \<upsilon> x. f (x, y))\<close> if \<open>constructive f\<close>
+for f :: \<open>'a :: complete_restriction_space \<times> 'b :: complete_restriction_space \<Rightarrow> 'a\<close>
+proof (rule constructiveI)
+  fix n and y y' :: 'b assume \<open>y \<down> n = y' \<down> n\<close>
+  show \<open>(\<upsilon> x. f (x, y)) \<down> Suc n = (\<upsilon> x. f (x, y')) \<down> Suc n\<close>
+  proof (subst (1 2) restriction_restriction_fix_is)
+    show \<open>constructive (\<lambda>x. f (x, y'))\<close> and \<open>constructive (\<lambda>x. f (x, y))\<close>
+      using \<open>constructive f\<close> constructive_prod_domain_iff by blast+
+  next
+    from \<open>y \<down> n = y' \<down> n\<close> show \<open>((\<lambda>x. f (x, y)) ^^ Suc n) undefined \<down> Suc n =
+                                ((\<lambda>x. f (x, y')) ^^ Suc n) undefined \<down> Suc n\<close>
+      by (induct n, simp_all)
+        (use restriction_0_related \<open>constructive f\<close> constructiveD in blast,
+          metis (no_types, lifting) \<open>constructive f\<close> constructiveD prod.sel
+          restriction_related_pred restriction_prod_def)
+  qed
+qed
+
+
+
+\<comment> \<open>``Bekic's Theorem'' in HOLCF\<close>
+
+lemma restriction_fix_prod_is :
+  \<open>(\<upsilon> p. f p) = (\<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y)))),
+                 \<upsilon> y. snd (f (\<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y)))), y)))\<close>
+  (is \<open>(\<upsilon> p. f p) = (?x, ?y)\<close>) if \<open>constructive f\<close>
+for f :: \<open>'a :: complete_restriction_space \<times> 'b :: complete_restriction_space \<Rightarrow> 'a \<times> 'b\<close>
+proof (rule restriction_fix_unique[OF \<open>constructive f\<close>])
+  have \<open>constructive (\<lambda>p. snd (f p))\<close>
+    by (fact non_destructive_comp_constructive[OF non_destructive_snd \<open>constructive f\<close>])
+  hence \<open>constructive (\<lambda>x. \<upsilon> y. snd (f (x, y)))\<close>
+    by (fact constructive_restriction_fix_right)
+  hence \<open>non_destructive (\<lambda>x. \<upsilon> y. snd (f (x, y)))\<close>
+    by (fact constructive_imp_non_destructive)
+  hence \<open>non_destructive (\<lambda>x. (x, \<upsilon> y. snd (f (x, y))))\<close>
+    by (fact non_destructive_prod_codomain[OF non_destructive_id(2)])
+  hence \<open>constructive (\<lambda>x. f (x, \<upsilon> y. snd (f (x, y))))\<close>
+    by (fact constructive_comp_non_destructive[OF \<open>constructive f\<close>])
+  hence * : \<open>constructive (\<lambda>x. fst (f (x, \<upsilon> y. snd (f (x, y)))))\<close>
+    by (fact non_destructive_comp_constructive[OF non_destructive_fst])
+
+  have \<open>non_destructive (\<lambda>x. \<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y)))))\<close>
+    by (fact constructive_imp_non_destructive[OF constructive_const])
+  hence \<open>non_destructive (Pair (\<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y))))))\<close>
+    by (fact non_destructive_prod_codomain[OF _ non_destructive_id(2)])
+  hence \<open>constructive (\<lambda>x. f (\<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y)))), x))\<close>
+    by (fact constructive_comp_non_destructive[OF \<open>constructive f\<close>])
+  hence ** : \<open>constructive (\<lambda>x. snd (f (\<upsilon> x. fst (f (x, \<upsilon> y. snd (f (x, y)))), x)))\<close>
+    by (fact non_destructive_comp_constructive[OF non_destructive_snd])
+
+  have \<open>fst (f (?x, ?y)) = ?x\<close>
+    by (rule trans [symmetric, OF restriction_fix_eq[OF "*"]]) simp
+  moreover have \<open>snd (f (?x, ?y)) = ?y\<close>
+    by (rule trans [symmetric, OF restriction_fix_eq[OF "**"]]) simp
+  ultimately show \<open>f (?x, ?y) = (?x, ?y)\<close> by (cases \<open>f (?x, ?y)\<close>) simp
+qed 
+
+
+(*<*)
+end
+  (*>*)
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/Restriction_Spaces_Topology.thy
--- /dev/null
+++ thys/Restriction_Spaces/Restriction_Spaces_Topology.thy
@@ -0,0 +1,841 @@
+(***********************************************************************************
+ * Copyright (c) 2025 Université Paris-Saclay
+ *
+ * Author: Benoît Ballenghien, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ * Author: Benjamin Puyobro, Université Paris-Saclay,
+           IRT SystemX, CNRS, ENS Paris-Saclay, LMF
+ * Author: Burkhart Wolff, Université Paris-Saclay,
+           CNRS, ENS Paris-Saclay, LMF
+ *
+ * All rights reserved.
+ *
+ * Redistribution and use in source and binary forms, with or without
+ * modification, are permitted provided that the following conditions are met:
+ *
+ * * Redistributions of source code must retain the above copyright notice, this
+ *
+ * * Redistributions in binary form must reproduce the above copyright notice,
+ *   this list of conditions and the following disclaimer in the documentation
+ *   and/or other materials provided with the distribution.
+ *
+ * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
+ * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
+ * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
+ * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
+ * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
+ * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
+ * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
+ * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
+ * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
+ * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+ *
+ * SPDX-License-Identifier: BSD-2-Clause
+ ***********************************************************************************)
+
+
+section \<open>Topological Notions\<close>
+
+(*<*)
+theory Restriction_Spaces_Topology
+  imports Restriction_Spaces_Prod Restriction_Spaces_Fun
+begin
+  (*>*)
+
+
+named_theorems restriction_cont_simpset \<comment> \<open>For future automation.\<close>
+
+subsection \<open>Continuity\<close>
+
+context restriction begin
+
+definition restriction_cont_at :: \<open>['b :: restriction \<Rightarrow> 'a, 'b] \<Rightarrow> bool\<close>
+  (\<open>cont\<down> (_) at (_)\<close> [1000, 1000])
+  where \<open>cont\<down> f at \<Sigma> \<equiv> \<forall>\<sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close>
+
+lemma restriction_cont_atI : \<open>(\<And>\<sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>) \<Longrightarrow> cont\<down> f at \<Sigma>\<close>
+  by (simp add: restriction_cont_at_def)
+
+lemma restriction_cont_atD : \<open>cont\<down> f at \<Sigma> \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close>
+  by (simp add: restriction_cont_at_def)
+
+lemma restriction_cont_at_comp [restriction_cont_simpset] :
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> cont\<down> g at (f \<Sigma>) \<Longrightarrow> cont\<down> (\<lambda>x. g (f x)) at \<Sigma>\<close>
+  by (simp add: restriction_cont_at_def restriction_class.restriction_cont_at_def)
+
+lemma restriction_cont_at_if_then_else [restriction_cont_simpset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> cont\<down> (f x) at \<Sigma>; \<And>x. \<not> P x \<Longrightarrow> cont\<down> (g x) at \<Sigma>\<rbrakk>
+   \<Longrightarrow> cont\<down> (\<lambda>y. if P x then f x y else g x y) at \<Sigma>\<close>
+  by (auto intro!: restriction_cont_atI) (blast dest: restriction_cont_atD)+
+
+
+
+definition restriction_open :: \<open>'a set \<Rightarrow> bool\<close> (\<open>open\<down>\<close>)
+  where \<open>open\<down> U \<equiv> \<forall>\<Sigma>\<in>U. \<forall>\<sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longrightarrow> (\<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> U)\<close>
+
+lemma restriction_openI : \<open>(\<And>\<Sigma> \<sigma>. \<Sigma> \<in> U \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> U) \<Longrightarrow> open\<down> U\<close>
+  by (simp add: restriction_open_def)
+
+lemma restriction_openD : \<open>open\<down> U \<Longrightarrow> \<Sigma> \<in> U \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> U\<close>
+  by (simp add: restriction_open_def)
+
+lemma restriction_openE :
+  \<open>open\<down> U \<Longrightarrow> \<Sigma> \<in> U \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>n0. (\<And>n. n0 \<le> k \<Longrightarrow> \<sigma> k \<in> U) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_openD by blast
+
+lemma restriction_open_UNIV  [simp] : \<open>open\<down> UNIV\<close>
+  and restriction_open_empty [simp] : \<open>open\<down> {}\<close>
+  by (simp_all add: restriction_open_def)
+
+
+lemma restriction_open_union :
+  \<open>open\<down> U \<Longrightarrow> open\<down> V \<Longrightarrow> open\<down> (U \<union> V)\<close>
+  by (metis Un_iff restriction_open_def)
+
+lemma restriction_open_Union :
+  \<open>(\<And>i. i \<in> I \<Longrightarrow> open\<down> (U i)) \<Longrightarrow> open\<down> (\<Union>i\<in>I. U i)\<close>
+  by (rule restriction_openI) (metis UN_iff restriction_openD)
+
+lemma restriction_open_inter :
+  \<open>open\<down> (U \<inter> V)\<close> if \<open>open\<down> U\<close> and \<open>open\<down> V\<close>
+proof (rule restriction_openI)
+  fix \<Sigma> \<sigma> assume \<open>\<Sigma> \<in> U \<inter> V\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  from \<open>\<Sigma> \<in> U \<inter> V\<close> have \<open>\<Sigma> \<in> U\<close> and \<open>\<Sigma> \<in> V\<close> by simp_all
+  from \<open>open\<down> U\<close> \<open>\<Sigma> \<in> U\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_openD
+  obtain n0 where \<open>\<forall>k\<ge>n0. \<sigma> k \<in> U\<close> by blast
+  moreover from \<open>open\<down> V\<close> \<open>\<Sigma> \<in> V\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_openD
+  obtain n1 where \<open>\<forall>k\<ge>n1. \<sigma> k \<in> V\<close> by blast
+  ultimately have \<open>\<forall>k\<ge>max n0 n1. \<sigma> k \<in> U \<inter> V\<close> by simp
+  thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> U \<inter> V\<close> by blast
+qed
+
+lemma restriction_open_finite_Inter :
+  \<open>finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> open\<down> (U i)) \<Longrightarrow> open\<down> (\<Inter>i\<in>I. U i)\<close>
+  by (induct I rule: finite_induct)
+    (simp_all add: restriction_open_inter)
+
+
+
+definition restriction_closed :: \<open>'a set \<Rightarrow> bool\<close> (\<open>closed\<down>\<close>)
+  where \<open>closed\<down> S \<equiv> open\<down> (- S)\<close>
+
+lemma restriction_closedI : \<open>(\<And>\<Sigma> \<sigma>. \<Sigma> \<notin> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<notin> S) \<Longrightarrow> closed\<down> S\<close>
+  by (simp add: restriction_closed_def restriction_open_def)
+
+lemma restriction_closedD : \<open>closed\<down> S \<Longrightarrow> \<Sigma> \<notin> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<notin> S\<close>
+  by (simp add: restriction_closed_def restriction_open_def)
+
+lemma restriction_closedE :
+  \<open>closed\<down> S \<Longrightarrow> \<Sigma> \<notin> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>n0. (\<And>n. n0 \<le> k \<Longrightarrow> \<sigma> k \<notin> S) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_closedD by blast
+
+lemma restriction_closed_UNIV  [simp] : \<open>closed\<down> UNIV\<close>
+  and restriction_closed_empty [simp] : \<open>closed\<down> {}\<close>
+  by (simp_all add: restriction_closed_def)
+
+end
+
+
+subsection \<open>Balls\<close>
+
+context restriction begin
+
+definition restriction_cball :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a set\<close> (\<open>\<B>\<down>'(_, _')\<close>)
+  where \<open>\<B>\<down>(a, n) \<equiv> {x. x \<down> n = a \<down> n}\<close>
+
+lemma restriction_cball_mem_iff : \<open>x \<in> \<B>\<down>(a, n) \<longleftrightarrow>  x \<down> n = a \<down> n\<close>
+  and restriction_cball_memI    : \<open>x \<down> n = a \<down> n \<Longrightarrow> x \<in> \<B>\<down>(a, n)\<close>
+  and restriction_cball_memD    : \<open>x \<in> \<B>\<down>(a, n) \<Longrightarrow> x \<down> n = a \<down> n\<close>
+  by (simp_all add: restriction_cball_def)
+
+
+
+abbreviation (iff) restriction_ball :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a set\<close>
+  where \<open>restriction_ball a n \<equiv> \<B>\<down>(a, Suc n)\<close>
+
+lemma \<open>x \<in> restriction_ball a n \<longleftrightarrow>  x \<down> Suc n = a \<down> Suc n\<close>
+  and \<open>x \<down> Suc n = a \<down> Suc n \<Longrightarrow> x \<in> restriction_ball a n\<close>
+  and \<open>x \<in> restriction_ball a n \<Longrightarrow> x \<down> Suc n = a \<down> Suc n\<close>
+  by (simp_all add: restriction_cball_def)
+
+
+
+lemma \<open>a \<in> restriction_ball a n\<close>
+  and center_mem_restriction_cball [simp] : \<open>a \<in> \<B>\<down>(a, n)\<close>
+  by (simp_all add: restriction_cball_memI)
+
+lemma (in restriction_space) restriction_cball_0_is_UNIV [simp] :
+  \<open>\<B>\<down>(a, 0) = UNIV\<close> by (simp add: restriction_cball_def)
+
+
+
+
+lemma every_point_of_restriction_cball_is_centre :
+  \<open>b \<in> \<B>\<down>(a, n) \<Longrightarrow> \<B>\<down>(a, n) = \<B>\<down>(b, n)\<close>
+  by (simp add: restriction_cball_def)
+
+
+lemma \<open>b \<in> restriction_ball a n \<Longrightarrow> restriction_ball a n = restriction_ball b n\<close>
+  by (simp add: every_point_of_restriction_cball_is_centre)
+
+
+definition restriction_sphere :: \<open>'a \<Rightarrow> nat \<Rightarrow> 'a set\<close> (\<open>\<S>\<down>'(_, _')\<close>)
+  where \<open>\<S>\<down>(a, n) \<equiv> {x. x \<down> n = a \<down> n \<and> x \<down> Suc n \<noteq> a \<down> Suc n}\<close>
+
+lemma restriction_sphere_mem_iff : \<open>x \<in> \<S>\<down>(a, n) \<longleftrightarrow> x \<down> n = a \<down> n \<and> x \<down> Suc n \<noteq> a \<down> Suc n\<close>
+  and restriction_sphere_memI    : \<open>x \<down> n = a \<down> n \<Longrightarrow> x \<down> Suc n \<noteq> a \<down> Suc n \<Longrightarrow> x \<in> \<S>\<down>(a, n)\<close>
+  and restriction_sphere_memD1   : \<open>x \<in> \<S>\<down>(a, n) \<Longrightarrow> x \<down> n = a \<down> n\<close>
+  and restriction_sphere_memD2   : \<open>x \<in> \<S>\<down>(a, n) \<Longrightarrow> x \<down> Suc n \<noteq> a \<down> Suc n\<close>
+  by (simp_all add: restriction_sphere_def)
+
+lemma restriction_sphere_is_diff : \<open>\<S>\<down>(a, n) = \<B>\<down>(a, n) - \<B>\<down>(a, Suc n)\<close>
+  by (simp add: set_eq_iff restriction_sphere_mem_iff restriction_cball_mem_iff)
+
+
+
+
+lemma restriction_open_restriction_cball [simp] : \<open>open\<down> \<B>\<down>(a, n)\<close>
+  by (metis restriction_cball_mem_iff restriction_tendstoE restriction_openI)
+
+lemma restriction_closed_restriction_cball [simp] : \<open>closed\<down> \<B>\<down>(a, n)\<close>
+  by (metis restriction_cball_mem_iff restriction_closedI restriction_tendstoE)
+
+lemma restriction_open_Compl_iff : \<open>open\<down> (- S) \<longleftrightarrow> closed\<down> S\<close>
+  by (simp add: restriction_closed_def)
+
+lemma restriction_open_restriction_sphere [simp] : \<open>open\<down> \<S>\<down>(a, n)\<close>
+  by (simp add: restriction_sphere_is_diff Diff_eq
+      restriction_open_Compl_iff restriction_open_inter)
+
+lemma restriction_closed_restriction_sphere : \<open>closed\<down> \<S>\<down>(a, n)\<close>
+  by (simp add: restriction_closed_def restriction_sphere_is_diff)
+    (simp add: restriction_open_union restriction_open_Compl_iff)
+
+end
+
+
+context restriction_space begin
+
+lemma restriction_cball_anti_mono : \<open>n \<le> m \<Longrightarrow> \<B>\<down>(a, m) \<subseteq> \<B>\<down>(a, n)\<close>
+  by (meson restriction_cball_memD restriction_cball_memI restriction_related_le subsetI)
+
+lemma inside_every_cball_iff_eq : \<open>(\<forall>n. x \<in> \<B>\<down>(\<Sigma>, n)) \<longleftrightarrow> x = \<Sigma>\<close>
+  by (simp add: all_restriction_related_iff_related restriction_cball_mem_iff)
+
+
+
+lemma Inf_many_inside_cball_iff_eq : \<open>(\<exists>\<infinity>n. x \<in> \<B>\<down>(\<Sigma>, n)) \<longleftrightarrow> x = \<Sigma>\<close>
+  by (unfold INFM_nat_le)
+    (meson inside_every_cball_iff_eq nle_le restriction_cball_anti_mono subset_eq)
+
+lemma Inf_many_inside_cball_imp_eq : \<open>\<exists>\<infinity>n. x \<in> \<B>\<down>(\<Sigma>, n) \<Longrightarrow> x = \<Sigma>\<close>
+  by (simp add: Inf_many_inside_cball_iff_eq)
+
+
+
+lemma restriction_cballs_disjoint_or_subset :
+  \<open>\<B>\<down>(a, n) \<inter> \<B>\<down>(b, m) = {} \<or> \<B>\<down>(a, n) \<subseteq> \<B>\<down>(b, m) \<or> \<B>\<down>(b, m) \<subseteq> \<B>\<down>(a, n)\<close>
+proof (unfold disj_imp, intro impI)
+  assume \<open>\<B>\<down>(a, n) \<inter> \<B>\<down>(b, m) \<noteq> {}\<close> \<open>\<not> \<B>\<down>(a, n) \<subseteq> \<B>\<down>(b, m)\<close>
+  from \<open>\<B>\<down>(a, n) \<inter> \<B>\<down>(b, m) \<noteq> {}\<close> obtain x where \<open>x \<in> \<B>\<down>(a, n)\<close> \<open>x \<in> \<B>\<down>(b, m)\<close> by blast
+  with every_point_of_restriction_cball_is_centre
+  have \<open>\<B>\<down>(a, n) = \<B>\<down>(x, n)\<close> \<open>\<B>\<down>(b, m) = \<B>\<down>(x, m)\<close> by auto
+  with \<open>\<not> \<B>\<down>(a, n) \<subseteq> \<B>\<down>(b, m)\<close> show \<open>\<B>\<down>(b, m) \<subseteq> \<B>\<down>(a, n)\<close>
+    by (metis nle_le restriction_cball_anti_mono)
+qed
+
+
+
+lemma equal_restriction_to_cball :
+  \<open>a \<notin> \<B>\<down>(b, n) \<Longrightarrow> x \<in> \<B>\<down>(b, n) \<Longrightarrow> y \<in> \<B>\<down>(b, n) \<Longrightarrow> x \<down> k = a \<down> k \<longleftrightarrow> y \<down> k = a \<down> k\<close>
+  by (metis nat_le_linear restriction_cball_memD restriction_cball_memI restriction_related_le)
+
+end
+
+
+context restriction begin
+
+lemma restriction_tendsto_iff_restriction_cball_characterization :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longleftrightarrow> (\<forall>n. \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> \<B>\<down>(\<Sigma>, n))\<close>
+  by (metis restriction_cball_mem_iff restriction_tendsto_def)
+
+corollary restriction_tendsto_restriction_cballI : \<open>(\<And>n. \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> \<B>\<down>(\<Sigma>, n)) \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_tendsto_iff_restriction_cball_characterization)
+
+corollary restriction_tendsto_restriction_cballD : \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> \<B>\<down>(\<Sigma>, n)\<close>
+  by (simp add: restriction_tendsto_iff_restriction_cball_characterization)
+
+corollary restriction_tendsto_restriction_cballE :
+  \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<And>n0. (\<And>k. n0 \<le> k \<Longrightarrow> \<sigma> k \<in> \<B>\<down>(\<Sigma>, n)) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_tendsto_restriction_cballD by blast
+
+end
+
+
+
+context restriction begin
+
+theorem restriction_closed_iff_sequential_characterization :
+  \<open>closed\<down> S \<longleftrightarrow> (\<forall>\<Sigma> \<sigma>. range \<sigma> \<subseteq> S \<longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longrightarrow> \<Sigma> \<in> S)\<close>
+proof (intro iffI allI impI)
+  show \<open>restriction_closed S \<Longrightarrow> range \<sigma> \<subseteq> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<Sigma> \<in> S\<close> for \<Sigma> \<sigma>
+    by (meson le_add1 range_subsetD restriction_closedD)
+next
+  assume * : \<open>\<forall>\<Sigma> \<sigma>. range \<sigma> \<subseteq> S \<longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<longrightarrow> \<Sigma> \<in> S\<close>
+  show \<open>closed\<down> S\<close>
+  proof (rule restriction_closedI, rule ccontr)
+    fix \<Sigma> \<sigma> assume \<open>\<Sigma> \<notin> S\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>\<nexists>n0. \<forall>k\<ge>n0. \<sigma> k \<notin> S\<close>
+    from \<open>\<nexists>n0. \<forall>k\<ge>n0. \<sigma> k \<notin> S\<close> INFM_nat_le have \<open>\<exists>\<infinity>k. \<sigma> k \<in> S\<close> by auto
+    from this[THEN extraction_subseqD[of \<open>\<lambda>x. x \<in> S\<close>]]
+    obtain f :: \<open>nat \<Rightarrow> nat\<close> where \<open>strict_mono f\<close> \<open>\<forall>k. \<sigma> (f k) \<in> S\<close> by blast
+    from \<open>\<forall>k. \<sigma> (f k) \<in> S\<close> have \<open>range (\<sigma> \<circ> f) \<subseteq> S\<close> by auto
+    moreover from \<open>strict_mono f\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> have \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+      by (fact restriction_tendsto_subseq)
+    ultimately have \<open>\<Sigma> \<in> S\<close> by (fact "*"[rule_format])
+    with \<open>\<Sigma> \<notin> S\<close> show False ..
+  qed
+qed
+
+
+corollary restriction_closed_sequentialI :
+  \<open>(\<And>\<Sigma> \<sigma>. range \<sigma> \<subseteq> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<Sigma> \<in> S) \<Longrightarrow> closed\<down> S\<close>
+  by (simp add: restriction_closed_iff_sequential_characterization)
+
+corollary restriction_closed_sequentialD :
+  \<open>closed\<down> S \<Longrightarrow> range \<sigma> \<subseteq> S \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> \<Sigma> \<in> S\<close>
+  by (simp add: restriction_closed_iff_sequential_characterization)
+
+end
+
+
+
+context restriction_space begin
+
+theorem restriction_open_iff_restriction_cball_characterization :
+  \<open>open\<down> U \<longleftrightarrow> (\<forall>\<Sigma>\<in>U. \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U)\<close>
+proof (intro iffI ballI)
+  show \<open>open\<down> U \<Longrightarrow> \<Sigma> \<in> U \<Longrightarrow> \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close> for \<Sigma>
+  proof (rule ccontr)
+    assume \<open>open\<down> U\<close> \<open>\<Sigma> \<in> U\<close> \<open>\<nexists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close>
+    from \<open>\<nexists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close> have \<open>\<forall>n. \<exists>\<sigma>. \<sigma> \<in> \<B>\<down>(\<Sigma>, n) \<inter> - U\<close> by auto
+    then obtain \<sigma> where \<open>\<sigma> n \<in> \<B>\<down>(\<Sigma>, n)\<close> \<open>\<sigma> n \<in> - U\<close> for n by (metis IntE)
+    from \<open>\<And>n. \<sigma> n \<in> \<B>\<down>(\<Sigma>, n)\<close> have \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> 
+      by (metis restriction_cball_memD restriction_related_le restriction_tendstoI)
+    moreover from \<open>open\<down> U\<close> have \<open>closed\<down> (- U)\<close>
+      by (simp add: restriction_closed_def)
+    ultimately have \<open>\<Sigma> \<in> - U\<close>
+      using \<open>\<And>n. \<sigma> n \<in> - U\<close> restriction_closedD by blast
+    with \<open>\<Sigma> \<in> U\<close> show False by simp
+  qed
+next
+  show \<open>\<forall>\<Sigma>\<in>U. \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U \<Longrightarrow> open\<down> U\<close>
+    by (metis center_mem_restriction_cball restriction_open_def
+        restriction_open_restriction_cball subset_iff)
+qed
+
+
+corollary restriction_open_restriction_cballI :
+  \<open>(\<And>\<Sigma>. \<Sigma> \<in> U \<Longrightarrow> \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U) \<Longrightarrow> open\<down> U\<close>
+  by (simp add: restriction_open_iff_restriction_cball_characterization)
+
+corollary restriction_open_restriction_cballD :
+  \<open>open\<down> U \<Longrightarrow> \<Sigma> \<in> U \<Longrightarrow> \<exists>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U\<close>
+  by (simp add: restriction_open_iff_restriction_cball_characterization)
+
+corollary restriction_open_restriction_cballE :
+  \<open>open\<down> U \<Longrightarrow> \<Sigma> \<in> U \<Longrightarrow> (\<And>n. \<B>\<down>(\<Sigma>, n) \<subseteq> U \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  using restriction_open_restriction_cballD by blast
+
+end
+
+
+context restriction begin
+
+definition restriction_cont_on :: \<open>['b :: restriction \<Rightarrow> 'a, 'b set] \<Rightarrow> bool\<close>
+  (\<open>cont\<down> (_) on (_)\<close> [1000, 1000])
+  where \<open>cont\<down> f on A \<equiv> \<forall>\<Sigma>\<in>A. cont\<down> f at \<Sigma>\<close>
+
+lemma restriction_cont_onI : \<open>(\<And>\<Sigma> \<sigma>. \<Sigma> \<in> A \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>) \<Longrightarrow> cont\<down> f on A\<close>
+  by (simp add: restriction_cont_on_def restriction_cont_atI)
+
+lemma restriction_cont_onD : \<open>cont\<down> f on A \<Longrightarrow> \<Sigma> \<in> A \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close>
+  by (simp add: restriction_cont_on_def restriction_cont_atD)
+
+lemma restriction_cont_on_comp [restriction_cont_simpset] :
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> g on B \<Longrightarrow> f ` A \<subseteq> B \<Longrightarrow> cont\<down> (\<lambda>x. g (f x)) on A\<close>
+  by (simp add: image_subset_iff restriction_cont_at_comp
+      restriction_cont_on_def restriction_class.restriction_cont_on_def)
+
+lemma restriction_cont_on_if_then_else [restriction_cont_simpset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> cont\<down> (f x) on A; \<And>x. \<not> P x \<Longrightarrow> cont\<down> (g x) on A\<rbrakk>
+   \<Longrightarrow> cont\<down> (\<lambda>y. if P x then f x y else g x y) on A\<close>
+  by (auto intro!: restriction_cont_onI) (blast dest: restriction_cont_onD)+
+
+lemma restriction_cont_on_subset [restriction_cont_simpset] :
+  \<open>cont\<down> f on B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> cont\<down> f on A\<close>
+  by (simp add: restriction_cont_on_def subset_iff)
+
+
+abbreviation restriction_cont :: \<open>['b :: restriction \<Rightarrow> 'a] \<Rightarrow> bool\<close> (\<open>cont\<down>\<close>)
+  where \<open>cont\<down> f \<equiv> cont\<down> f on UNIV\<close>
+
+lemma restriction_contI : \<open>(\<And>\<Sigma> \<sigma>. \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>) \<Longrightarrow> cont\<down> f\<close>
+  by (simp add: restriction_cont_onI)
+
+lemma restriction_contD : \<open>cont\<down> f \<Longrightarrow> \<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma> \<Longrightarrow> (\<lambda>n. f (\<sigma> n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close>
+  by (simp add: restriction_cont_onD)
+
+lemma restriction_cont_comp [restriction_cont_simpset] :
+  \<open>cont\<down> g \<Longrightarrow> cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>x. g (f x))\<close>
+  by (simp add: restriction_cont_on_comp)
+
+lemma restriction_cont_if_then_else [restriction_cont_simpset] :
+  \<open>\<lbrakk>\<And>x. P x \<Longrightarrow> cont\<down> (f x); \<And>x. \<not> P x \<Longrightarrow> cont\<down> (g x)\<rbrakk>
+   \<Longrightarrow> cont\<down> (\<lambda>y. if P x then f x y else g x y)\<close>
+  by (auto intro!: restriction_contI) (blast dest: restriction_contD)+
+
+end
+
+
+
+context restriction_space begin
+
+theorem restriction_cont_at_iff_restriction_cball_characterization :
+  \<open>cont\<down> f at \<Sigma> \<longleftrightarrow> (\<forall>n. \<exists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n))\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+proof (intro iffI allI)
+  show \<open>\<exists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n)\<close> if \<open>cont\<down> f at \<Sigma>\<close> for n
+  proof (rule ccontr)
+    assume \<open>\<nexists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n)\<close>
+    hence \<open>\<forall>k. \<exists>\<psi>. \<psi> \<in> f ` \<B>\<down>(\<Sigma>, k) \<and> \<psi> \<notin> \<B>\<down>(f \<Sigma>, n)\<close> by auto
+    then obtain \<psi> where * : \<open>\<psi> k \<in> f ` \<B>\<down>(\<Sigma>, k)\<close> \<open>\<psi> k \<notin> \<B>\<down>(f \<Sigma>, n)\<close> for k by metis
+    from "*"(1) obtain \<sigma> where ** : \<open>\<sigma> k \<in> \<B>\<down>(\<Sigma>, k)\<close> \<open>\<psi> k = f (\<sigma> k)\<close> for k
+      by (simp add: image_iff) metis
+    have \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+      by (rule restriction_class.restriction_tendsto_restriction_cballI)
+        (use "**"(1) restriction_space_class.restriction_cball_anti_mono in blast)
+    with restriction_cont_atD \<open>restriction_cont_at f \<Sigma>\<close>
+    have \<open>(\<lambda>k. f (\<sigma> k)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> by blast
+    hence \<open>\<psi> \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> by (fold "**"(2))
+    with "*"(2) restriction_tendsto_restriction_cballD show False by blast
+  qed
+next
+  show \<open>\<forall>n. \<exists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n) \<Longrightarrow> cont\<down> f at \<Sigma>\<close>
+    by (intro restriction_cont_atI restriction_tendsto_restriction_cballI)
+      (meson image_iff restriction_class.restriction_tendsto_restriction_cballD subset_eq)
+qed
+
+
+corollary restriction_cont_at_restriction_cballI :
+  \<open>(\<And>n. \<exists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n)) \<Longrightarrow> cont\<down> f at \<Sigma>\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_at_iff_restriction_cball_characterization)
+
+corollary restriction_cont_at_restriction_cballD :
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> \<exists>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n)\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_at_iff_restriction_cball_characterization)
+
+corollary restriction_cont_at_restriction_cballE :
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> (\<And>k. f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n) \<Longrightarrow> thesis) \<Longrightarrow> thesis\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  using restriction_cont_at_restriction_cballD by blast
+
+
+
+theorem restriction_cont_iff_restriction_open_characterization :
+  \<open>cont\<down> f \<longleftrightarrow> (\<forall>U. open\<down> U \<longrightarrow> open\<down> (f -` U))\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+proof (intro iffI allI impI)
+  fix U :: \<open>'a set\<close> assume \<open>cont\<down> f\<close> \<open>open\<down> U\<close>
+  show \<open>open\<down> (f -` U)\<close>
+  proof (rule restriction_space_class.restriction_open_restriction_cballI)
+    fix \<Sigma> assume \<open>\<Sigma> \<in> f -` U\<close>
+    hence \<open>f \<Sigma> \<in> U\<close> by simp
+    with \<open>open\<down> U\<close> restriction_open_restriction_cballD
+    obtain n where \<open>\<B>\<down>(f \<Sigma>, n) \<subseteq> U\<close> by blast
+    moreover obtain k where \<open>f ` \<B>\<down>(\<Sigma>, k) \<subseteq> \<B>\<down>(f \<Sigma>, n)\<close>
+      by (meson UNIV_I \<open>cont\<down> f\<close> restriction_cont_at_restriction_cballE restriction_cont_on_def)
+    ultimately have \<open>\<B>\<down>(\<Sigma>, k) \<subseteq> f -` U\<close> by blast
+    thus \<open>\<exists>k. \<B>\<down>(\<Sigma>, k) \<subseteq> f -` U\<close> ..
+  qed
+next
+  show \<open>\<forall>U. open\<down> U \<longrightarrow> open\<down> (f -` U) \<Longrightarrow> cont\<down> f\<close>
+    by (unfold restriction_cont_on_def, intro ballI restriction_cont_at_restriction_cballI)
+      (simp add: image_subset_iff_subset_vimage restriction_space_class.restriction_open_restriction_cballD)
+qed
+
+corollary restriction_cont_restriction_openI :
+  \<open>(\<And>U. open\<down> U \<Longrightarrow> open\<down> (f -` U)) \<Longrightarrow> cont\<down> f\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_iff_restriction_open_characterization)
+
+corollary restriction_cont_restriction_openD :
+  \<open>cont\<down> f \<Longrightarrow> open\<down> U \<Longrightarrow> open\<down> (f -` U)\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_iff_restriction_open_characterization)
+
+
+theorem restriction_cont_iff_restriction_closed_characterization :
+  \<open>cont\<down> f \<longleftrightarrow> (\<forall>S. closed\<down> S \<longrightarrow> closed\<down> (f -` S))\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (metis boolean_algebra_class.boolean_algebra.double_compl local.restriction_closed_def
+      restriction_class.restriction_closed_def restriction_cont_iff_restriction_open_characterization vimage_Compl)
+
+corollary restriction_cont_restriction_closedI :
+  \<open>(\<And>U. closed\<down> U \<Longrightarrow> closed\<down> (f -` U)) \<Longrightarrow> cont\<down> f\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_iff_restriction_closed_characterization)
+
+corollary restriction_cont_restriction_closedD :
+  \<open>cont\<down> f \<Longrightarrow> closed\<down> U \<Longrightarrow> closed\<down> (f -` U)\<close>
+  for f :: \<open>'b :: restriction_space \<Rightarrow> 'a\<close>
+  by (simp add: restriction_cont_iff_restriction_closed_characterization)
+
+
+theorem restriction_shift_on_restriction_open_imp_restriction_cont_on :
+  \<open>cont\<down> f on U\<close> if \<open>open\<down> U\<close> and \<open>restriction_shift_on f k U\<close>
+proof (intro restriction_cont_onI restriction_tendstoI)
+  fix \<Sigma> \<sigma> and n :: nat assume \<open>\<Sigma> \<in> U\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  with \<open>open\<down> U\<close> obtain n0 where \<open>\<forall>l\<ge>n0. \<sigma> l \<in> U\<close>
+    by (meson restriction_class.restriction_openD)
+  moreover from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>[THEN restriction_class.restriction_tendstoD]
+  obtain n1 where \<open>\<forall>l\<ge>n1. \<Sigma> \<down> nat (int n - k) = \<sigma> l \<down> nat (int n - k)\<close> ..
+  ultimately have \<open>\<forall>l\<ge>max n0 n1. \<sigma> l \<in> U \<and> \<Sigma> \<down> nat (int n - k) = \<sigma> l \<down> nat (int n - k)\<close> by simp
+  with \<open>\<Sigma> \<in> U\<close> \<open>restriction_shift_on f k U\<close> restriction_shift_onD
+  have \<open>\<forall>l\<ge>max n0 n1. f \<Sigma> \<down> nat (int (nat (int n - k)) + k) = f (\<sigma> l) \<down> nat (int (nat (int n - k)) + k)\<close> by blast
+  moreover have \<open>n \<le> nat (int (nat (int n - k)) + k)\<close> by auto
+  ultimately have \<open>\<forall>l\<ge>max n0 n1. f \<Sigma> \<down> n = f (\<sigma> l) \<down> n\<close> by (meson restriction_related_le)
+  thus \<open>\<exists>n2. \<forall>l\<ge>n2. f \<Sigma> \<down> n = f (\<sigma> l) \<down> n\<close> by blast
+qed
+
+corollary restriction_shift_imp_restriction_cont [restriction_cont_simpset] :
+  \<open>restriction_shift f k \<Longrightarrow> cont\<down> f\<close>
+  by (simp add: restriction_shift_def
+      restriction_shift_on_restriction_open_imp_restriction_cont_on)
+
+corollary non_too_destructive_imp_restriction_cont [restriction_cont_simpset] :
+  \<open>non_too_destructive f \<Longrightarrow> cont\<down> f\<close>
+  by (simp add: non_too_destructive_def non_too_destructive_on_def
+      restriction_shift_on_restriction_open_imp_restriction_cont_on)
+
+
+end
+
+
+
+subsection \<open>Compactness\<close>
+
+context restriction begin
+
+definition restriction_compact :: \<open>'a set \<Rightarrow> bool\<close> (\<open>compact\<down>\<close>)
+  where \<open>compact\<down> K \<equiv>
+         \<forall>\<sigma>. range \<sigma> \<subseteq> K \<longrightarrow>
+         (\<exists>f :: nat \<Rightarrow> nat. \<exists>\<Sigma>. \<Sigma> \<in> K \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>)\<close>
+
+
+lemma restriction_compactI :
+  \<open>(\<And>\<sigma>. range \<sigma> \<subseteq> K \<Longrightarrow> \<exists>f :: nat \<Rightarrow> nat. \<exists>\<Sigma>. \<Sigma> \<in> K \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>)
+   \<Longrightarrow> compact\<down> K\<close> by (simp add: restriction_compact_def)
+
+lemma restriction_compactD :
+  \<open>compact\<down> K \<Longrightarrow> range \<sigma> \<subseteq> K \<Longrightarrow>
+   \<exists>f :: nat \<Rightarrow> nat. \<exists>\<Sigma>. \<Sigma> \<in> K \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (simp add: restriction_compact_def)
+
+lemma restriction_compactE :
+  assumes \<open>compact\<down> K\<close> and \<open>range \<sigma> \<subseteq> K\<close>
+  obtains f :: \<open>nat \<Rightarrow> nat\<close> and \<Sigma> where \<open>\<Sigma> \<in> K\<close> \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  by (meson assms restriction_compactD)
+
+lemma restriction_compact_empty [simp] : \<open>compact\<down> {}\<close>
+  by (simp add: restriction_compact_def)
+
+
+
+lemma (in restriction_space) restriction_compact_imp_restriction_closed :
+  \<open>closed\<down> K\<close> if \<open>compact\<down> K\<close>
+proof (rule restriction_closed_sequentialI)
+  fix \<sigma> \<Sigma> assume \<open>range \<sigma> \<subseteq> K\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  from restriction_compactD \<open>compact\<down> K\<close> \<open>range \<sigma> \<subseteq> K\<close> 
+  obtain f and \<Sigma>' where \<open>\<Sigma>' \<in> K\<close> \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>'\<close> by blast
+  from restriction_tendsto_subseq \<open>strict_mono f\<close> \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  have \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  with \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>'\<close> have \<open>\<Sigma>' = \<Sigma>\<close> by (fact restriction_tendsto_unique)
+  with \<open>\<Sigma>' \<in> K\<close> show \<open>\<Sigma> \<in> K\<close> by simp
+qed
+
+
+lemma restriction_compact_union : \<open>compact\<down> (K \<union> L)\<close>
+  if \<open>compact\<down> K\<close> and \<open>compact\<down> L\<close>
+proof (rule restriction_compactI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> _\<close> assume \<open>range \<sigma> \<subseteq> K \<union> L\<close>
+  { fix K L and f :: \<open>nat \<Rightarrow> nat\<close>
+    assume \<open>compact\<down> K\<close> \<open>strict_mono f\<close> \<open>\<sigma> (f n) \<in> K\<close> for n
+    from \<open>(\<And>n. \<sigma> (f n) \<in> K)\<close> have \<open>range (\<sigma> \<circ> f) \<subseteq> K\<close> by auto
+    with \<open>compact\<down> K\<close> restriction_compactD obtain g \<Sigma>
+      where \<open>\<Sigma> \<in> K\<close> \<open>strict_mono g\<close> \<open>(\<sigma> \<circ> f \<circ> g) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+    hence \<open>\<Sigma> \<in> K \<union> L \<and> strict_mono (f \<circ> g) \<and> (\<sigma> \<circ> (f \<circ> g)) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+      by (metis (no_types, lifting) Un_iff \<open>strict_mono f\<close> comp_assoc
+          monotone_on_o subset_UNIV)
+    hence \<open>\<exists>f \<Sigma>. \<Sigma> \<in> K \<union> L \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  } note * = this
+  have \<open>(\<exists>\<infinity>n. \<sigma> n \<in> K) \<or> (\<exists>\<infinity>n. \<sigma> n \<in> L)\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> ((\<exists>\<infinity>n. \<sigma> n \<in> K) \<or> (\<exists>\<infinity>n. \<sigma> n \<in> L))\<close>
+    hence \<open>finite {n. \<sigma> n \<in> K} \<and> finite {n. \<sigma> n \<in> L}\<close>
+      using frequently_cofinite by blast
+    then obtain n where \<open>n \<notin> {n. \<sigma> n \<in> K} \<and> n \<notin> {n. \<sigma> n \<in> L}\<close>
+      by (metis (mono_tags, lifting) INFM_nat_le dual_order.refl
+          frequently_cofinite le_sup_iff mem_Collect_eq)
+    hence \<open>\<sigma> n \<notin> K \<union> L\<close> by simp
+    with \<open>range \<sigma> \<subseteq> K \<union> L\<close> show False by blast
+  qed
+  thus \<open>\<exists>f \<Sigma>. \<Sigma> \<in> K \<union> L \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (elim disjE extraction_subseqE)
+      (use "*" \<open>compact\<down> K\<close> in blast, metis "*" Un_iff \<open>compact\<down> L\<close>)
+qed
+
+lemma restriction_compact_finite_Union :
+  \<open>\<lbrakk>finite I; \<And>i. i \<in> I \<Longrightarrow> compact\<down> (K i)\<rbrakk> \<Longrightarrow> compact\<down> (\<Union>i\<in>I. K i)\<close>
+  by (induct I rule: finite_induct)
+    (simp_all add: restriction_compact_union)
+
+
+lemma (in restriction_space) restriction_compact_Inter :
+  \<open>compact\<down> (\<Inter>i. K i)\<close> if \<open>\<And>i. compact\<down> (K i)\<close>
+proof (rule restriction_compactI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> 'a\<close> assume \<open>range \<sigma> \<subseteq> \<Inter> (range K)\<close>
+  hence \<open>range \<sigma> \<subseteq> K i\<close> for i by blast
+  with \<open>\<And>i. compact\<down> (K i)\<close> restriction_compactD
+  obtain f \<Sigma> where \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  from \<open>\<And>i. compact\<down> (K i)\<close> have \<open>closed\<down> (K i)\<close> for i
+    by (simp add: restriction_compact_imp_restriction_closed)
+  moreover from \<open>\<And>i. range \<sigma> \<subseteq> K i\<close> have \<open>range (\<sigma> \<circ> f) \<subseteq> K i\<close> for i by auto
+  ultimately have \<open>\<Sigma> \<in> K i\<close> for i
+    by (meson \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_closed_sequentialD)
+  with \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  show \<open>\<exists>f \<Sigma>. \<Sigma> \<in> \<Inter> (range K) \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+qed
+
+lemma finite_imp_restriction_compact : \<open>compact\<down> K\<close> if \<open>finite K\<close>
+proof (rule restriction_compactI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> _\<close> assume \<open>range \<sigma> \<subseteq> K\<close>
+  have \<open>\<exists>\<Sigma>\<in>K. \<exists>\<infinity>n. \<sigma> n = \<Sigma>\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> (\<exists>\<Sigma>\<in>K. \<exists>\<infinity>n. \<sigma> n = \<Sigma>)\<close>
+    hence \<open>\<forall>\<Sigma>\<in>K. finite {n. \<sigma> n = \<Sigma>}\<close> by (simp add: frequently_cofinite)
+    with \<open>finite K\<close> have \<open>finite (\<Union>\<Sigma>\<in>K. {n. \<sigma> n = \<Sigma>})\<close> by blast
+    also from \<open>range \<sigma> \<subseteq> K\<close> have \<open>(\<Union>\<Sigma>\<in>K. {n. \<sigma> n = \<Sigma>}) = UNIV\<close> by auto
+    finally show False by simp
+  qed
+  then obtain \<Sigma> where \<open>\<Sigma> \<in> K\<close> \<open>\<exists>\<infinity>n. \<sigma> n = \<Sigma>\<close> ..
+  from extraction_subseqD[of _ \<sigma>, OF \<open>\<exists>\<infinity>n. \<sigma> n = \<Sigma>\<close>]
+  obtain f :: \<open>nat \<Rightarrow> nat\<close> where \<open>strict_mono f\<close> \<open>\<sigma> (f n) = \<Sigma>\<close> for n by blast
+  from \<open>\<And>n. \<sigma> (f n) = \<Sigma>\<close> have \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+    by (simp add: restriction_tendstoI)
+  with \<open>strict_mono f\<close> \<open>\<Sigma> \<in> K\<close>
+  show \<open>\<exists>f \<Sigma>. \<Sigma> \<in> K \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+qed
+
+
+lemma restriction_compact_restriction_closed_subset : \<open>compact\<down> L\<close>
+  if \<open>L \<subseteq> K\<close> \<open>compact\<down> K\<close> \<open>closed\<down> L\<close>
+proof (rule restriction_compactI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> _\<close> assume \<open>range \<sigma> \<subseteq> L\<close>
+  with \<open>L \<subseteq> K\<close> have \<open>range \<sigma> \<subseteq> K\<close> by blast
+  with \<open>compact\<down> K\<close> restriction_compactD
+  obtain f \<Sigma> where \<open>\<Sigma> \<in> K\<close> \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  from \<open>range \<sigma> \<subseteq> L\<close> have \<open>range (\<sigma> \<circ> f) \<subseteq> L\<close> by auto
+  from restriction_closed_sequentialD \<open>restriction_closed L\<close>
+    \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> \<open>range (\<sigma> \<circ> f) \<subseteq> L\<close> have \<open>\<Sigma> \<in> L\<close> by blast
+  with \<open>strict_mono f\<close> \<open>(\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  show \<open>\<exists>f \<Sigma>. \<Sigma> \<in> L \<and> strict_mono f \<and> (\<sigma> \<circ> f) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+qed
+
+
+lemma restriction_cont_image_of_restriction_compact :
+  \<open>compact\<down> (f ` K)\<close> if \<open>compact\<down> K\<close> and \<open>cont\<down> f on K\<close>
+proof (rule restriction_compactI)
+  fix \<sigma> :: \<open>nat \<Rightarrow> _\<close> assume \<open>range \<sigma> \<subseteq> f ` K\<close>
+  hence \<open>\<forall>n. \<exists>\<gamma>. \<gamma> \<in> K \<and> \<sigma> n = f \<gamma>\<close> by (meson imageE range_subsetD)
+  then obtain \<gamma> :: \<open>nat \<Rightarrow> _\<close> where \<open>range \<gamma> \<subseteq> K\<close> \<open>\<sigma> n = f (\<gamma> n)\<close> for n
+    by (metis image_subsetI)
+  from restriction_class.restriction_compactD[OF \<open>compact\<down> K\<close> \<open>range \<gamma> \<subseteq> K\<close>]
+  obtain g \<Sigma> where \<open>\<Sigma> \<in> K\<close> \<open>strict_mono g\<close> \<open>(\<gamma> \<circ> g) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+  from \<open>cont\<down> f on K\<close> \<open>\<Sigma> \<in> K\<close>
+  have \<open>cont\<down> f at \<Sigma>\<close> by (simp add: restriction_cont_on_def)
+  with \<open>(\<gamma> \<circ> g) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> restriction_cont_atD
+  have \<open>(\<lambda>n. f ((\<gamma> \<circ> g) n)) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> by blast 
+  also have \<open>(\<lambda>n. f ((\<gamma> \<circ> g) n)) = (\<sigma> \<circ> g)\<close>
+    by (simp add: \<open>\<And>n. \<sigma> n = f (\<gamma> n)\<close> comp_def)
+  finally have \<open>(\<sigma> \<circ> g) \<midarrow>\<down>\<rightarrow> f \<Sigma>\<close> .
+  with \<open>\<Sigma> \<in> K\<close> \<open>strict_mono g\<close>
+  show \<open>\<exists>g \<Sigma>. \<Sigma> \<in> f ` K \<and> strict_mono g \<and> (\<sigma> \<circ> g) \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> by blast
+qed
+
+end
+
+
+
+
+
+subsection \<open>Properties for Function and Product\<close>
+
+lemma restriction_cball_fun_is : \<open>\<B>\<down>(f, n) = {g. \<forall>x. g x \<in> \<B>\<down>(f x, n)}\<close>
+  by (simp add: set_eq_iff restriction_cball_mem_iff restriction_fun_def) metis
+
+lemma restriction_cball_prod_is :
+  \<open>\<B>\<down>(\<Sigma>, n) = \<B>\<down>(fst \<Sigma>, n) \<times> \<B>\<down>(snd \<Sigma>, n)\<close>
+  by (simp add: set_eq_iff restriction_cball_def restriction_prod_def)
+
+
+lemma restriction_open_prod_imp_restriction_open_image_fst :
+  \<open>open\<down> (fst ` U)\<close> if \<open>open\<down> U\<close>
+proof (rule restriction_openI)
+  fix \<Sigma> \<sigma> assume \<open>\<Sigma> \<in> fst ` U\<close> and \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  from \<open>\<Sigma> \<in> fst ` U\<close> obtain v where \<open>(\<Sigma>, v) \<in> U\<close> by auto
+  from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> have \<open>(\<lambda>n. (\<sigma> n, v)) \<midarrow>\<down>\<rightarrow> (\<Sigma>, v)\<close>
+    by (simp add: restriction_tendsto_prod_iff restriction_tendsto_const)
+  from restriction_openD[OF \<open>restriction_open U\<close> \<open>(\<Sigma>, v) \<in> U\<close> this]
+  obtain n0 where \<open>\<forall>k\<ge>n0. (\<sigma> k, v) \<in> U\<close> ..
+  thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> fst ` U\<close> by (metis fst_conv imageI)
+qed
+
+lemma restriction_open_prod_imp_restriction_open_image_snd :
+  \<open>open\<down> (snd ` U)\<close> if \<open>open\<down> U\<close>
+proof (rule restriction_openI)
+  fix \<Sigma> \<sigma> assume \<open>\<Sigma> \<in> snd ` U\<close> and \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+  from \<open>\<Sigma> \<in> snd ` U\<close> obtain u where \<open>(u, \<Sigma>) \<in> U\<close> by auto
+  from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> have \<open>(\<lambda>n. (u, \<sigma> n)) \<midarrow>\<down>\<rightarrow> (u, \<Sigma>)\<close>
+    by (simp add: restriction_tendsto_prod_iff restriction_tendsto_const)
+  from restriction_openD[OF \<open>restriction_open U\<close> \<open>(u, \<Sigma>) \<in> U\<close> this]
+  obtain n0 where \<open>\<forall>k\<ge>n0. (u, \<sigma> k) \<in> U\<close> ..
+  thus \<open>\<exists>n0. \<forall>k\<ge>n0. \<sigma> k \<in> snd ` U\<close> by (metis snd_conv imageI)
+qed
+
+lemma restriction_open_prod_iff :
+  \<open>open\<down> (U \<times> V) \<longleftrightarrow> (V = {} \<or> open\<down> U) \<and> (U = {} \<or> open\<down> V)\<close>
+proof (intro iffI conjI)
+  show \<open>open\<down> (U \<times> V) \<Longrightarrow> V = {} \<or> open\<down> U\<close>
+    by (metis fst_image_times restriction_open_prod_imp_restriction_open_image_fst)
+next
+  show \<open>open\<down> (U \<times> V) \<Longrightarrow> U = {} \<or> open\<down> V\<close>
+    by (metis restriction_open_prod_imp_restriction_open_image_snd snd_image_times)
+next
+  assume \<open>(V = {} \<or> open\<down> U) \<and> (U = {} \<or> open\<down> V)\<close>
+  then consider \<open>U = {}\<close> | \<open>V = {}\<close> | \<open>open\<down> U \<and> open\<down> V\<close> by fast
+  thus \<open>open\<down> (U \<times> V)\<close>
+  proof cases
+    show \<open>U = {} \<Longrightarrow> open\<down> (U \<times> V)\<close> by simp
+  next
+    show \<open>V = {} \<Longrightarrow> open\<down> (U \<times> V)\<close> by simp
+  next
+    show \<open>open\<down> (U \<times> V)\<close> if * : \<open>open\<down> U \<and> open\<down> V\<close>
+    proof (rule restriction_openI)
+      fix \<Sigma> \<sigma> assume \<open>\<Sigma> \<in> U \<times> V\<close> and \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close>
+      from \<open>\<Sigma> \<in> U \<times> V\<close> have \<open>fst \<Sigma> \<in> U\<close> \<open>snd \<Sigma> \<in> V\<close> by auto
+      from \<open>\<sigma> \<midarrow>\<down>\<rightarrow> \<Sigma>\<close> have \<open>(\<lambda>n. fst (\<sigma> n)) \<midarrow>\<down>\<rightarrow> fst \<Sigma>\<close> \<open>(\<lambda>n. snd (\<sigma> n)) \<midarrow>\<down>\<rightarrow> snd \<Sigma>\<close>
+        by (simp_all add: restriction_tendsto_prod_iff)
+      from restriction_openD[OF "*"[THEN conjunct1] \<open>fst \<Sigma> \<in> U\<close> \<open>(\<lambda>n. fst (\<sigma> n)) \<midarrow>\<down>\<rightarrow> fst \<Sigma>\<close>]
+      obtain n0 where \<open>\<forall>k\<ge>n0. fst (\<sigma> k) \<in> U\<close> ..
+      moreover from restriction_openD[OF "*"[THEN conjunct2] \<open>snd \<Sigma> \<in> V\<close> \<open>(\<lambda>n. snd (\<sigma> n)) \<midarrow>\<down>\<rightarrow> snd \<Sigma>\<close>]
+      obtain n1 where \<open>\<forall>k\<ge>n1. snd (\<sigma> k) \<in> V\<close> ..
+      ultimately have \<open>\<forall>k\<ge>max n0 n1. \<sigma> k \<in> U \<times> V\<close> by (simp add: mem_Times_iff)
+      thus \<open>\<exists>n2. \<forall>k\<ge>n2. \<sigma> k \<in> U \<times> V\<close> by blast
+    qed
+  qed
+qed
+
+
+
+lemma restriction_cont_at_prod_codomain_iff:
+  \<open>cont\<down> f at \<Sigma> \<longleftrightarrow> cont\<down> (\<lambda>x. fst (f x)) at \<Sigma> \<and> cont\<down> (\<lambda>x. snd (f x)) at \<Sigma>\<close>
+  by (auto simp add: restriction_cont_at_def restriction_tendsto_prod_iff)
+
+lemma restriction_cont_on_prod_codomain_iff:
+  \<open>cont\<down> f on A \<longleftrightarrow> cont\<down> (\<lambda>x. fst (f x)) on A \<and> cont\<down> (\<lambda>x. snd (f x)) on A\<close>
+  by (metis restriction_cont_at_prod_codomain_iff restriction_cont_on_def)
+
+lemma restriction_cont_prod_codomain_iff:
+  \<open>cont\<down> f \<longleftrightarrow> cont\<down> (\<lambda>x. fst (f x)) \<and> cont\<down> (\<lambda>x. snd (f x))\<close>
+  by (fact restriction_cont_on_prod_codomain_iff)
+
+
+lemma restriction_cont_at_prod_codomain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> cont\<down> (\<lambda>x. fst (f x)) at \<Sigma>\<close>
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> cont\<down> (\<lambda>x. snd (f x)) at \<Sigma>\<close>
+  by (simp_all add: restriction_cont_at_prod_codomain_iff)
+
+lemma restriction_cont_on_prod_codomain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> (\<lambda>x. fst (f x)) on A\<close>
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> (\<lambda>x. snd (f x)) on A\<close>
+  by (simp_all add: restriction_cont_on_prod_codomain_iff)
+
+lemma restriction_cont_prod_codomain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>x. fst (f x))\<close>
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>x. snd (f x))\<close>
+  by (simp_all add: restriction_cont_prod_codomain_iff)
+
+
+
+
+lemma restriction_cont_at_fun_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f at A \<Longrightarrow> cont\<down> (\<lambda>x. f x y) at A\<close>
+  by (rule restriction_cont_atI)
+    (metis restriction_cont_atD restriction_tendsto_fun_imp)
+
+lemma restriction_cont_on_fun_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f on A \<Longrightarrow> cont\<down> (\<lambda>x. f x y) on A\<close>
+  by (simp add: restriction_cont_at_fun_imp restriction_cont_on_def)
+
+corollary restriction_cont_fun_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>x. f x y)\<close>
+  by (fact restriction_cont_on_fun_imp) 
+
+
+
+lemma restriction_cont_at_prod_domain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> cont\<down> (\<lambda>x. f (x, snd \<Sigma>)) at (fst \<Sigma>)\<close>
+  \<open>cont\<down> f at \<Sigma> \<Longrightarrow> cont\<down> (\<lambda>y. f (fst \<Sigma>, y)) at (snd \<Sigma>)\<close>
+  for f :: \<open>'a :: restriction_space \<times> 'b :: restriction_space \<Rightarrow> 'c :: restriction_space\<close>
+  by (simp add: restriction_cball_prod_is subset_iff image_iff
+      restriction_cont_at_iff_restriction_cball_characterization,
+      meson center_mem_restriction_cball)+
+
+lemma restriction_cont_on_prod_domain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> (\<lambda>x. f (x, y)) on {x. (x, y) \<in> A}\<close>
+  \<open>cont\<down> (\<lambda>y. f (x, y)) on {y. (x, y) \<in> A}\<close> if \<open>cont\<down> f on A\<close>
+for f :: \<open>'a :: restriction_space \<times> 'b :: restriction_space \<Rightarrow> 'c :: restriction_space\<close>
+proof -
+  show \<open>cont\<down> (\<lambda>x. f (x, y)) on {x. (x, y) \<in> A}\<close>
+  proof (unfold restriction_cont_on_def, rule ballI)
+    fix x assume \<open>x \<in> {x. (x, y) \<in> A}\<close>
+    with \<open>cont\<down> f on A\<close> have \<open>cont\<down> f at (x, y)\<close>
+      unfolding restriction_cont_on_def by simp
+    thus \<open>cont\<down> (\<lambda>x. f (x, y)) at x\<close>
+      by (fact restriction_cont_at_prod_domain_imp[of f \<open>(x, y)\<close>, simplified])
+  qed
+next
+  show \<open>cont\<down> (\<lambda>y. f (x, y)) on {y. (x, y) \<in> A}\<close>
+  proof (unfold restriction_cont_on_def, rule ballI)
+    fix y assume \<open>y \<in> {y. (x, y) \<in> A}\<close>
+    with \<open>cont\<down> f on A\<close> have \<open>cont\<down> f at (x, y)\<close>
+      unfolding restriction_cont_on_def by simp
+    thus \<open>cont\<down> (\<lambda>y. f (x, y)) at y\<close>
+      by (fact restriction_cont_at_prod_domain_imp[of f \<open>(x, y)\<close>, simplified])
+  qed
+qed
+
+lemma restriction_cont_prod_domain_imp [restriction_cont_simpset] :
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>x. f (x, y))\<close>
+  \<open>cont\<down> f \<Longrightarrow> cont\<down> (\<lambda>y. f (x, y))\<close>
+for f :: \<open>'a :: restriction_space \<times> 'b :: restriction_space \<Rightarrow> 'c :: restriction_space\<close>
+  by (metis UNIV_I restriction_cont_at_prod_domain_imp(1) restriction_cont_on_def split_pairs)
+    (metis UNIV_I restriction_cont_at_prod_domain_imp(2) restriction_cont_on_def split_pairs)
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Restriction_Spaces/document/root.tex
--- /dev/null
+++ thys/Restriction_Spaces/document/root.tex
@@ -0,0 +1,90 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{pifont}
+
+\usepackage{mathpartir}
+
+\usepackage{amsmath}
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{Restriction\_Spaces: a Fixed-Point Theory}
+\author{Benoît Ballenghien \and Benjamin Puyobro \and Burkhart Wolff}
+\maketitle
+
+\abstract{
+  Fixed-point constructions are fundamental to defining recursive and co-recursive functions.
+  However, a general axiom $Y f = f (Y f)$ leads to inconsistency, and definitions
+  must therefore be based on theories guaranteeing existence under suitable conditions.
+  In \verb'Isabelle/HOL', such constructions are typically based on sets, well-founded orders
+  or domain-theoretic models such as for example \verb'HOLCF'.
+  In this submission we introduce \verb'Restriction_Spaces', a formalization of spaces
+  equipped with a so-called restriction, denoted by $\downarrow$,
+  satifying three properties:
+  \begin{align*}
+    & x \downarrow 0 = y \downarrow 0\\
+    & x \downarrow n \downarrow m = x \downarrow \mathrm{min} \ n \ m\\
+    & x \not = y \Longrightarrow \exists n. \ x \downarrow n \not = y \downarrow n
+  \end{align*}
+  They turn out to be cartesian closed and admit natural notions of constructiveness and
+  completeness, enabling the definition of a fixed-point operator under verifiable side-conditions.
+  This is achieved in our entry, from topological definitions to induction principles.
+  Additionally, we configure the simplifier so that it can automatically solve both
+  constructiveness and admissibility subgoals, as long as users write higher-order rules
+  for their operators.
+  Since our implementation relies on axiomatic type classes, the resulting library
+  is a fully abstract, flexible and reusable framework.
+}
+
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+%\bibliographystyle{abbrv}
+%\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Collections.thy
--- /dev/null
+++ thys/Shallow_Expressions/Collections.thy
@@ -0,0 +1,126 @@
+section \<open> Collections \<close>
+
+theory Collections
+  imports Substitutions
+begin
+
+text \<open> A lens whose source is a collection type (e.g. a list or map) can be divided into several
+  lenses, corresponding to each of the elements in the collection. This can be used to support
+  update of an individual collection element, such as an array update. Here, we provide the
+  infrastructure to support such collection lenses~\cite{Foster2020-LocalVars}.  \<close>
+
+subsection \<open> Partial Lens Definedness \<close>
+
+text \<open> Partial lenses (e.g. @{term mwb_lens}) are only defined for certain states. For example,
+  the list lens is defined only when the source list is of a sufficient length. Below, we define
+  a predicate that characterises the states in a which such a lens is defined. \<close>
+
+definition lens_defined :: "('a \<Longrightarrow> 's) \<Rightarrow> (bool, 's) expr" where
+[expr_defs]: "lens_defined x = ($v \<in> \<S>\<guillemotleft>x\<guillemotright>)e"
+
+syntax "_lens_defined" :: "svid \<Rightarrow> logic" ("D'(_')")
+translations "_lens_defined x" == "CONST lens_defined x"
+
+expr_constructor lens_defined
+
+subsection \<open> Dynamic Lenses \<close>
+
+text \<open> Dynamics lenses~\cite{Foster2020-LocalVars} are used to model elements of a lens indexed 
+  by some type @{typ "'i"}. The index is typically used to select different elements of a collection. 
+  The lens is ``dynamic'' because the particular index is provided by an expression @{typ "'s \<Rightarrow> 'i"},
+  which can change from state to state. We normally assume that this expression does not
+  refer to the indexed lens itself. \<close>
+
+definition dyn_lens :: "('i \<Rightarrow> ('a \<Longrightarrow> 's)) \<Rightarrow> ('s \<Rightarrow> 'i) \<Rightarrow> ('a \<Longrightarrow> 's)" where
+[lens_defs]: "dyn_lens f x = \<lparr> lens_get = (\<lambda> s. getf (x s) s), lens_put = (\<lambda> s v. putf (x s) s v) \<rparr>"
+
+lemma dyn_lens_mwb [simp]: "\<lbrakk> \<And> i. mwb_lens (f i); \<And> i. $f(i) \<sharp> e \<rbrakk> \<Longrightarrow> mwb_lens (dyn_lens f e)"
+  apply (unfold_locales, auto simp add: expr_defs lens_defs lens_indep.lens_put_irr2)
+  apply (metis lens_override_def mwb_lens_weak weak_lens.put_get)
+  apply (metis lens_override_def mwb_lens.put_put)
+  done
+
+lemma ind_lens_vwb [simp]: "\<lbrakk> \<And> i. vwb_lens (f i); \<And> i. $f(i) \<sharp> e \<rbrakk> \<Longrightarrow> vwb_lens (dyn_lens f e)"
+  by (unfold_locales, auto simp add: lens_defs expr_defs lens_indep.lens_put_irr2 lens_scene_override)
+     (metis mwb_lens_weak vwb_lens_mwb weak_lens.put_get, metis mwb_lens.put_put vwb_lens_mwb)
+
+lemma src_dyn_lens: "\<lbrakk> \<And> i. mwb_lens (f i); \<And> i. $f(i) \<sharp> e \<rbrakk> \<Longrightarrow> \<S>dyn_lens f e = {s. s \<in> \<S>f (e s)}"
+  by (auto simp add: lens_defs expr_defs lens_source_def lens_scene_override unrest)
+     (metis mwb_lens.put_put)+
+
+lemma subst_lookup_dyn_lens [usubst]: "\<lbrakk> \<And> i. f i \<bowtie> x \<rbrakk> \<Longrightarrow> \<langle>subst_upd \<sigma> (dyn_lens f k) e\<rangle>s x = \<langle>\<sigma>\<rangle>s x"
+  by (expr_simp, metis (mono_tags, lifting) lens_indep.lens_put_irr2)
+
+lemma get_upd_dyn_lens [usubst_eval]: "\<lbrakk> \<And> i. f i \<bowtie> x \<rbrakk> \<Longrightarrow> getx (subst_upd \<sigma> (dyn_lens f k) e s) = getx (\<sigma> s)"
+  by (expr_simp, metis lens_indep.lens_put_irr2)
+
+subsection \<open> Overloaded Collection Lens \<close>
+
+text \<open> The following polymorphic constant is used to provide implementations of different
+  collection lenses. Type @{typ "'k"} is the index into the collection. For the list
+  collection lens, the index has type @{typ "nat"}. \<close>
+
+consts collection_lens :: "'k \<Rightarrow> ('a \<Longrightarrow> 's)"
+
+definition [lens_defs]: "fun_collection_lens = fun_lens"
+definition [lens_defs]: "list_collection_lens = list_lens"
+
+adhoc_overloading 
+  collection_lens \<rightleftharpoons> fun_collection_lens and
+  collection_lens \<rightleftharpoons> list_collection_lens  
+
+lemma vwb_fun_collection_lens [simp]: "vwb_lens (fun_collection_lens k)"
+  by (simp add: fun_collection_lens_def fun_vwb_lens)
+
+lemma put_fun_collection_lens [simp]: 
+  "putfun_collection_lens i = (\<lambda>f. fun_upd f i)"
+  by (simp add: fun_collection_lens_def fun_lens_def)
+
+lemma mwb_list_collection_lens [simp]: "mwb_lens (list_collection_lens i)"
+  by (simp add: list_collection_lens_def list_mwb_lens)
+
+lemma source_list_collection_lens: "\<S>list_collection_lens i = {xs. i < length xs}"
+  by (simp add: list_collection_lens_def source_list_lens)
+
+lemma put_list_collection_lens [simp]: 
+  "putlist_collection_lens i = (\<lambda> xs. list_augment xs i)"
+  by (simp add: list_collection_lens_def list_lens_def)
+
+subsection \<open> Syntax for Collection Lenses \<close>
+
+text \<open> We add variable identifier syntax for collection lenses, which allows us to write
+  @{text "x[i]"} for some collection and index. \<close>
+
+abbreviation "dyn_lens_poly f x i \<equiv> dyn_lens (\<lambda> k. ns_alpha x (f k)) i"
+
+syntax
+  "_svid_collection" :: "svid \<Rightarrow> logic \<Rightarrow> svid" ("_[_]" [999, 0] 999)
+
+translations
+  "_svid_collection x e" == "CONST dyn_lens_poly CONST collection_lens x (e)e"
+
+lemma source_ns_alpha: "\<lbrakk> mwb_lens a; mwb_lens x \<rbrakk> \<Longrightarrow> \<S>ns_alpha a x = {s \<in> \<S>a. geta s \<in> \<S>x}"
+  by (simp add: ns_alpha_def source_lens_comp)
+
+lemma defined_vwb_lens [simp]: "vwb_lens x \<Longrightarrow> D(x) = (True)e"
+  by (simp add: lens_defined_def)
+     (metis UNIV_I vwb_lens_iff_mwb_UNIV_src)
+
+lemma defined_list_collection_lens [simp]:
+  "\<lbrakk> vwb_lens x; $x \<sharp> e \<rbrakk> \<Longrightarrow> D(x[e]) = (e < length($x))e"
+  by (simp add: lens_defined_def src_dyn_lens unrest source_ns_alpha source_list_collection_lens)
+     (simp add: lens_defs wb_lens.source_UNIV)
+
+lemma lens_defined_list_code [code_unfold]: 
+  "vwb_lens x \<Longrightarrow> lens_defined (ns_alpha x (list_collection_lens i)) = (\<guillemotleft>i\<guillemotright> < length($x))e"
+  by (simp add: lens_defined_def src_dyn_lens unrest source_ns_alpha source_list_collection_lens)
+     (simp add: lens_defs wb_lens.source_UNIV)
+
+text \<open> The next theorem allows the simplification of a collection lens update. \<close>
+
+lemma get_subst_upd_dyn_lens [simp]: 
+  "mwb_lens x \<Longrightarrow> getx (subst_upd \<sigma> (dyn_lens_poly cl x (e)e) v s)
+                  = lens_put (cl (e (\<sigma> s))) (getx (\<sigma> s)) (v s)"
+  by expr_simp
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Expr_Util.ML
--- /dev/null
+++ thys/Shallow_Expressions/Expr_Util.ML
@@ -0,0 +1,50 @@
+(* Utilities for manipulating shallow expression syntax *)
+
+signature EXPR_UTIL =
+sig
+  val const_or_free: Proof.context -> string -> term
+  val subst_tab: term -> term Symtab.table
+  val tab_subst: Proof.context -> term Symtab.table -> term
+  val log_vars: term -> string list
+end
+
+structure Expr_Util : EXPR_UTIL =
+struct
+local
+  open Syntax
+in
+  (* Extract a table of variable assignments for a substitution *)
+  fun subst_tab' m (Const (@{const_name subst_upd}, _) $ s $ x $ e) = 
+    (case x of
+       Const (n, _) => subst_tab' (Symtab.update (n, e) m) s |
+       Free (n, _) => subst_tab' (Symtab.update (n, e) m) s |
+       _ => m) |
+  subst_tab' m _ = m;
+
+  val subst_tab = subst_tab' Symtab.empty;
+
+  (* Insert a constant or free variable depending on whether x refers to a constant *)
+  fun const_or_free ctx x = 
+    let val consts = (Proof_Context.consts_of ctx)
+        val {const_space, ...} = Consts.dest consts
+        val c = Consts.intern consts x
+    in if (Name_Space.declared const_space c) then Const (c, dummyT) else Free (x, dummyT)
+    end;
+
+  (* Construct a substitution from a table *)
+  fun tab_subst ctx m = 
+    List.foldl (fn ((x, e), s) => const @{const_name subst_upd} $ s $ const_or_free ctx x $ e) 
+               (const @{const_name subst_id}) 
+               (Symtab.dest m);
+
+  (* Extract the "logical variables" from an expression, excluding lenses *)
+  fun log_vars (Const (@{const_name lens_get}, _) $ _ $ _) = Ord_List.make string_ord [] 
+  | log_vars (Bound _) = []
+  | log_vars (Abs (_, _, e)) = log_vars e
+  | log_vars (Const (_, _)) = []
+  | log_vars (Free (n, _)) = Ord_List.make string_ord [n]
+  | log_vars (e $ f) = Ord_List.union string_ord (log_vars e) (log_vars f)
+  | log_vars (Var (_, _)) = [];
+
+end
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Expressions.thy
--- /dev/null
+++ thys/Shallow_Expressions/Expressions.thy
@@ -0,0 +1,252 @@
+section \<open> Expressions \<close>
+
+theory Expressions
+  imports Variables
+  keywords "expr_constructor" "expr_function" :: "thy_decl_block"
+begin
+
+subsection \<open> Types and Constructs \<close>
+
+named_theorems expr_defs and named_expr_defs
+
+text \<open> An expression is represented simply as a function from the state space @{typ "'s"} to
+  the return type @{typ "'a"}, which is the simplest shallow model for Isabelle/HOL. 
+
+  The aim of this theory is to provide transparent conversion between this representation 
+  and a more intuitive expression syntax. For example, an expression @{term "x + y"} where 
+  $x$ and $y$ are both state variables, can be represented by @{term "\<lambda> s. getx s + gety s"} 
+  when both variables are modelled using lenses. Rather than having to write $\lambda$-terms 
+  directly, it is more convenient to hide this threading of state behind a parser. We introduce
+  the expression bracket syntax, @{text "(_)e"} to support this.
+\<close>
+
+type_synonym ('a, 's) expr = "'s \<Rightarrow> 'a"
+
+text \<open> The following constructor is used to syntactically mark functions that actually
+  denote expressions. It is semantically vacuous. \<close>
+
+definition SEXP :: "('s \<Rightarrow> 'a) \<Rightarrow> ('a, 's) expr" ("[_]e") where
+[expr_defs]: "SEXP x = x"
+
+lemma SEXP_apply [simp]: "SEXP e s = (e s)" by (simp add: SEXP_def)
+
+lemma SEXP_idem [simp]: "[[e]e]e = [e]e" by (simp add: SEXP_def)
+
+text \<open> We can create the core constructs of a simple expression language as indicated below. \<close>
+
+abbreviation (input) var :: "('a \<Longrightarrow> 's) \<Rightarrow> ('a, 's) expr" where
+"var x \<equiv> (\<lambda> s. getx s)"
+
+abbreviation (input) lit :: "'a \<Rightarrow> ('a, 's) expr" where
+"lit k \<equiv> (\<lambda> s. k)"
+
+abbreviation (input) uop :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a, 's) expr \<Rightarrow> ('b, 's) expr" where
+"uop f e \<equiv> (\<lambda> s. f (e s))"
+
+abbreviation (input) bop 
+  :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a, 's) expr \<Rightarrow> ('b, 's) expr \<Rightarrow> ('c, 's) expr" where
+"bop f e1 e2 \<equiv> (\<lambda> s. f (e1 s) (e2 s))"
+
+definition taut :: "(bool, 's) expr \<Rightarrow> bool" where
+[expr_defs]: "taut e = (\<forall> s. e s)"
+
+definition expr_select :: "('a, 's) expr \<Rightarrow> ('b \<Longrightarrow> 'a) \<Rightarrow> ('b, 's) expr" where
+[expr_defs, code_unfold]: "expr_select e x = (\<lambda> s. getx (e s))"
+
+definition expr_if :: "('a, 's) expr \<Rightarrow> (bool, 's) expr \<Rightarrow> ('a, 's) expr \<Rightarrow> ('a, 's) expr" where
+[expr_defs, code_unfold]: "expr_if P b Q = (\<lambda> s. if (b s) then P s else Q s)"
+
+subsection \<open> Lifting Parser and Printer \<close>
+
+text \<open> The lifting parser creates a parser directive that converts an expression to a 
+  @{const SEXP} boxed $\lambda$-term that gives it a semantics. A pretty printer converts
+  a boxed $\lambda$-term back to an expression. \<close>
+
+nonterminal sexp
+
+text \<open> We create some syntactic constants and define parse and print translations for them. \<close>
+
+syntax
+  "_sexp_state"      :: "id"
+  "_sexp_quote"      :: "logic \<Rightarrow> logic" ("'(_')e")
+  \<comment> \<open> Convert the expression to a lambda term, but do not box it. \<close>
+  "_sexp_quote_1way" :: "logic \<Rightarrow> logic" ("'(_')e")
+  "_sexp_lit"        :: "logic \<Rightarrow> logic" ("\<guillemotleft>_\<guillemotright>")
+  "_sexp_var"        :: "svid \<Rightarrow> logic" ("$_" [990] 990)
+  "_sexp_evar"       :: "id_position \<Rightarrow> logic" ("@_" [999] 999)
+  "_sexp_evar"       :: "logic \<Rightarrow> logic" ("@'(_')" [999] 999)
+  "_sexp_pqt"        :: "logic \<Rightarrow> sexp" ("[_]e")
+  "_sexp_taut"       :: "logic \<Rightarrow> logic" ("`_`")
+  "_sexp_select"     :: "logic \<Rightarrow> svid \<Rightarrow> logic" ("_:_" [1000, 999] 1000)
+  "_sexp_if"         :: "logic \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" ("(3_ \<triangleleft> _ \<triangleright>/ _)" [52,0,53] 52)
+
+ML_file \<open>Lift_Expr.ML\<close>
+
+text \<open> We create a number of attributes for configuring the way the parser works. \<close> 
+
+declare [[pretty_print_exprs=true]]
+
+
+text \<open> We can toggle pretty printing of $\lambda$ expressions using @{attribute pretty_print_exprs}. \<close>
+
+declare [[literal_variables=false]]
+
+text \<open> Expressions, of course, can contain variables. However, a variable can denote one of
+  three things: (1) a state variable (i.e. a lens); (2) a placeholder for a value (i.e. a
+  HOL literal); and (3) a placeholder for another expression. The attribute @{attribute literal_variables}
+  selects option (2) as the default behaviour when true, and option (3) when false. \<close>
+
+expr_constructor expr_select
+expr_constructor expr_if
+
+text \<open> Some constants should not be lifted, since they are effectively constructors for expressions.
+  The command @{command expr_constructor} allows us to specify such constants to not be lifted.
+  This being the case, the arguments are left unlifted, unless included in a list of numbers
+  before the constant name. The state is passed as the final argument to expression constructors.
+\<close>
+
+parse_translation \<open> 
+  [(@{syntax_const "_sexp_state"}, fn ctx => fn term => Syntax.free Lift_Expr.state_id),
+   (@{syntax_const "_sexp_quote"}
+   , fn ctx => fn terms =>
+      case terms of
+        [Const (@{const_syntax SEXP}, t) $ e] => Const (@{const_name SEXP}, t) $ e |
+        [e] =>
+            Syntax.const @{const_name SEXP} $ Lift_Expr.mk_lift_expr ctx dummyT e),
+   (@{syntax_const "_sexp_quote_1way"}
+   , fn ctx => fn terms =>
+      case terms of
+        [e] => Lift_Expr.mk_lift_expr ctx dummyT e)]
+\<close>
+
+print_translation \<open>
+  [(@{const_syntax "SEXP"}
+   , fn ctx => fn ts =>
+     if not (Config.get ctx Lift_Expr.pretty_print_exprs)
+     then Term.list_comb (Syntax.const @{syntax_const "_sexp_pqt"}, ts)
+     else
+     Syntax.const @{syntax_const "_sexp_quote"} 
+     $ Lift_Expr.print_expr ctx (betapply ((hd ts), Syntax.const @{syntax_const "_sexp_state"})))]
+\<close>
+
+translations
+  "_sexp_var x" => "getx _sexp_state"
+  "_sexp_taut p" == "CONST taut (p)e"
+  "_sexp_select e x" == "CONST expr_select (e)e x"
+  "_sexp_if P b Q" == "CONST expr_if P (b)e Q"
+  "_sexp_var (_svid_tuple (_of_svid_list (x +L y)))" <= "_sexp_var (x +L y)"
+
+text \<open> The main directive is the $e$ subscripted brackets, @{term "(e)e"}. This converts the 
+  expression $e$ to a boxed $\lambda$ term. Essentially, the behaviour is as follows:
+
+  \begin{enumerate}
+    \item a new $\lambda$ abstraction over the state variable $s$ is wrapped around $e$;
+    \item every occurrence of a free lens @{term "$x"} in $e$ is replaced by @{term "getx s"};
+    \item every occurrence of an expression variable @{term "e"} is replaced by @{term "e s"};
+    \item every occurrence of any other free variable is left unchanged.
+  \end{enumerate}
+
+  The pretty printer does this in reverse. \<close>
+
+text \<open> Below is a grammar category for lifted expressions. \<close>
+
+nonterminal sexpr
+
+syntax "_sexpr" :: "logic \<Rightarrow> sexpr" ("_")
+
+parse_translation \<open>
+  [(@{syntax_const "_sexpr"}, fn ctx => fn [e] => 
+    Syntax.const @{const_name SEXP} 
+            $ (lambda (Syntax.free Lift_Expr.state_id) 
+                      (Lift_Expr.lift_expr ctx dummyT (Term_Position.strip_positions e))))]
+\<close>
+
+subsection \<open> Reasoning \<close>
+
+lemma expr_eq_iff: "P = Q \<longleftrightarrow> `P = Q`"
+  by (simp add: taut_def fun_eq_iff)
+
+lemma refine_iff_implies: "P \<le> Q \<longleftrightarrow> `P \<longrightarrow> Q`"
+  by (simp add: le_fun_def taut_def)
+
+lemma taut_True [simp]: "`True` = True"
+  by (simp add: taut_def)
+
+lemma taut_False [simp]: "`False` = False"
+  by (simp add: taut_def)
+
+lemma tautI: "\<lbrakk> \<And> s. P s \<rbrakk> \<Longrightarrow> taut P"
+  by (simp add: taut_def)
+
+named_theorems expr_simps
+
+text \<open> Lemmas to help automation of expression reasoning \<close>
+
+lemma fst_case_sum [simp]: "fst (case p of Inl x \<Rightarrow> (a1 x, a2 x) | Inr x \<Rightarrow> (b1 x, b2 x)) = (case p of Inl x \<Rightarrow> a1 x | Inr x \<Rightarrow> b1 x)"
+  by (simp add: sum.case_eq_if)
+
+lemma snd_case_sum [simp]: "snd (case p of Inl x \<Rightarrow> (a1 x, a2 x) | Inr x \<Rightarrow> (b1 x, b2 x)) = (case p of Inl x \<Rightarrow> a2 x | Inr x \<Rightarrow> b2 x)"
+  by (simp add: sum.case_eq_if)
+
+lemma sum_case_apply [simp]: "(case p of Inl x \<Rightarrow> f x | Inr x \<Rightarrow> g x) y = (case p of Inl x \<Rightarrow> f x y | Inr x \<Rightarrow> g x y)"
+  by (simp add: sum.case_eq_if)
+
+text \<open> Proof methods for simplifying shallow expressions to HOL terms. The first retains the lens
+  structure, and the second removes it when alphabet lenses are present.  \<close>
+
+method expr_lens_simp uses add = 
+  ((simp add: expr_simps)? \<comment> \<open> Perform any possible simplifications retaining the lens structure \<close>
+   ;((simp add: fun_eq_iff prod.case_eq_if expr_defs named_expr_defs lens_defs add)? ; \<comment> \<open> Explode the rest \<close>
+     (simp add: expr_defs named_expr_defs lens_defs add)?))
+
+method expr_simp uses add = (expr_lens_simp add: alpha_splits add)
+
+text \<open> Methods for dealing with tautologies \<close>
+
+method expr_lens_taut uses add = 
+  (rule tautI;
+   expr_lens_simp add: add)
+
+method expr_taut uses add = 
+  (rule tautI;
+   expr_simp add: add; 
+   rename_alpha_vars?)
+
+text \<open> A method for simplifying shallow expressions to HOL terms and applying @{method auto} \<close>
+
+method expr_auto uses add =
+  (expr_simp add: add; 
+   (auto simp add: alpha_splits lens_defs add)?; 
+   (rename_alpha_vars)? \<comment> \<open> Rename any logical variables with v subscripts \<close>
+  )
+
+subsection \<open> Algebraic laws \<close>
+
+lemma expr_if_idem [simp]: "P \<triangleleft> b \<triangleright> P = P"
+  by expr_auto
+
+lemma expr_if_sym: "P \<triangleleft> b \<triangleright> Q = Q \<triangleleft> \<not>b \<triangleright> P"
+  by expr_auto
+
+lemma expr_if_assoc: "(P \<triangleleft> b \<triangleright> Q) \<triangleleft> c \<triangleright> R = P \<triangleleft> b \<and> c \<triangleright> (Q \<triangleleft> c \<triangleright> R)"
+  by expr_auto
+
+lemma expr_if_distr: "P \<triangleleft> b \<triangleright> (Q \<triangleleft> c \<triangleright> R) = (P \<triangleleft> b \<triangleright> Q) \<triangleleft> c \<triangleright> (P \<triangleleft> b \<triangleright> R)"
+  by expr_auto
+
+lemma expr_if_true [simp]: "P \<triangleleft> True \<triangleright> Q = P"
+  by expr_auto
+
+lemma expr_if_false [simp]: "P \<triangleleft> False \<triangleright> Q = Q"
+  by expr_auto
+
+lemma expr_if_reach [simp]: "P \<triangleleft> b \<triangleright> (Q \<triangleleft> b \<triangleright> R) = P \<triangleleft> b \<triangleright> R"
+  by expr_auto
+
+lemma expr_if_disj [simp]: "P \<triangleleft> b \<triangleright> (P \<triangleleft> c \<triangleright> Q) = P \<triangleleft> b \<or> c \<triangleright> Q"
+  by expr_auto
+
+lemma SEXP_expr_if: "[expr_if P b Q]e = expr_if P b Q"
+  by (simp add: SEXP_def)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Expressions_Tests.thy
--- /dev/null
+++ thys/Shallow_Expressions/Expressions_Tests.thy
@@ -0,0 +1,107 @@
+section \<open> Expression Test Cases and Examples \<close>
+
+theory Expressions_Tests
+  imports Expressions Named_Expressions
+begin
+
+text \<open> Some examples of lifted expressions follow. For now, we turn off the pretty printer so that 
+  we can see the results of the parser.\<close>
+
+declare [[pretty_print_exprs=false]]
+
+term "(f + g)e" \<comment> \<open> Lift an expression and insert @{const SEXP} for pretty printing \<close>
+term "(f + g)e" \<comment> \<open> Lift an expression and don't insert @{const SEXP} \<close>
+
+text \<open> The default behaviour of our parser is to recognise identifiers as expression variables.
+  So, the above expression becomes the term @{term "[\<lambda>\<s>. f \<s> + g \<s>]e"}. We can easily change
+  this using the attribute @{attribute literal_variables}: \<close>
+
+declare [[literal_variables]]
+
+term "(f + g)e"
+
+text \<open> Now, @{term f} and @{term g} are both parsed as literals, and so the term is 
+  @{term "[\<lambda>\<s>. f + g]e"}. Alternatively, we could have a lens in the expression, by marking
+  a free variable with a dollar : \<close>
+
+term "($x + g)e"
+
+text \<open> This gives the term @{term "[\<lambda>\<s>. getx \<s> + g]e"}. Although we have default behaviours
+  for parsing, we can use different markup to coerce identifiers to particular variable kinds. \<close>
+
+term "($x + @g)e"
+
+text \<open> This gives @{term "[\<lambda>\<s>. getx \<s> + g \<s>]e"}, the we have requested that @{term "g"} is 
+  treated as an expression variable. We can do similar with literal, as show below. \<close>
+
+term "(f + \<guillemotleft>x\<guillemotright>)e"
+
+text \<open> Some further examples follow. \<close>
+
+term "(\<guillemotleft>f\<guillemotright> (@e))e"
+
+term "(@f + @g)e"
+
+term "(@x)e"
+
+term "($x:y:z)e"
+
+term "(($x:y):z)e"
+
+term "(x::nat)e"
+
+term "(\<forall> x::nat. x > 2)e"
+
+term "SEXP(\<lambda> \<s>. getx \<s> + e \<s> + v)"
+
+term "(v \<in> $xs \<union> ($f) ys \<union> {} \<and> @e)e"
+
+text \<open> We now turn pretty printing back on, so we can see how the user sees expressions. \<close>
+
+declare [[pretty_print_exprs, literal_variables=false]]
+
+term "($x< = $x>)e"
+
+term "($x.1 = $y.2)e"
+
+text \<open> The pretty printer works even when we don't use the parser, as shown below. \<close>
+
+term "[\<lambda> \<s>. getx \<s> + e \<s> + v]e"
+
+text \<open> By default, dollars are printed next to free variables that are lenses. However, we can
+  alter this behaviour with the attribute @{attribute mark_state_variables}: \<close>
+
+declare [[mark_state_variables=false]]
+
+term "($x + e + v)e"
+
+text \<open> This way, the @{term "x"} variable is indistinguishable when printed from the @{term "e"}
+  and @{term "v"}. Usually, this information can be inferred from the types of the entities: \<close>
+
+alphabet st = 
+  x :: int
+
+term "(x + e + v)e"
+
+expression x_is_big over st is "x > 1000"
+
+term "(x_is_big \<longrightarrow> x > 0)e"
+
+text \<open> Here, @{term x} is a lens defined by the @{command alphabet} command, and so the lifting
+  translation treats it as a state variable. This is hidden from the user. \<close>
+
+dataspace testspace =
+   variables z :: int
+
+declare [[literal_variables]]
+
+context testspace
+begin
+
+edefinition "z_is_bigger y = (z > y)"
+
+term "(z_is_bigger (z + 1))e"
+
+end
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Extension.thy
--- /dev/null
+++ thys/Shallow_Expressions/Extension.thy
@@ -0,0 +1,128 @@
+section \<open> Extension and Restriction \<close>
+
+theory Extension
+  imports Substitutions
+begin
+
+text \<open> It is often necessary to coerce an expression into a different state space using a lens,
+  for example when the state space grows to add additional variables. Extension and restriction
+  is provided by @{term aext} and @{term ares} respectively. Here, we provide syntax translations
+  and reasoning support for these.
+\<close>
+
+subsection \<open> Syntax \<close>
+
+syntax 
+  "_aext"      :: "logic \<Rightarrow> svid \<Rightarrow> logic" (infixl "\<up>" 80)
+  "_ares"      :: "logic \<Rightarrow> svid \<Rightarrow> logic" (infixl "\<down>" 80)
+  "_pre"       :: "logic \<Rightarrow> logic" ("_<" [999] 1000)
+  "_post"      :: "logic \<Rightarrow> logic" ("_>" [999] 1000)
+  "_drop_pre"  :: "logic \<Rightarrow> logic" ("_<" [999] 1000)
+  "_drop_post" :: "logic \<Rightarrow> logic" ("_>" [999] 1000)
+
+translations
+  "_aext P a" == "CONST aext P a"
+  "_ares P a" == "CONST ares P a"
+  "_pre P" == "_aext (P)e fstL"
+  "_post P" == "_aext (P)e sndL"
+  "_drop_pre P" == "_ares (P)e fstL"
+  "_drop_post P" == "_ares (P)e sndL"
+
+expr_constructor aext
+expr_constructor ares
+
+named_theorems alpha
+
+subsection \<open> Laws \<close>
+
+lemma aext_var [alpha]: "($x)e \<up> a = ($a:x)e"
+  by (simp add: expr_defs lens_defs)
+
+lemma ares_aext [alpha]: "weak_lens a \<Longrightarrow> P \<up> a \<down> a = P"
+  by (simp add: expr_defs)
+
+lemma aext_ares [alpha]: "\<lbrakk> mwb_lens a; (- $a) \<sharp> P \<rbrakk> \<Longrightarrow> P \<down> a \<up> a = P"
+  unfolding unrest_compl_lens
+  by (auto simp add: expr_defs fun_eq_iff lens_create_def)
+
+lemma expr_pre [simp]: "e< (s1, s2) = (e)e s1"
+  by (simp add: subst_ext_def subst_app_def)
+
+lemma expr_post [simp]: "e> (s1, s2) = (@e)e s2"
+  by (simp add: subst_ext_def subst_app_def)
+
+lemma unrest_aext_expr_lens [unrest]: "\<lbrakk> mwb_lens x; x \<bowtie> a \<rbrakk> \<Longrightarrow> $x \<sharp> (P \<up> a)"
+  by (expr_simp add: lens_indep.lens_put_irr2)
+
+lemma unrest_init_pre [unrest]: "\<lbrakk> mwb_lens x; $x \<sharp> e \<rbrakk> \<Longrightarrow> $x< \<sharp> e<"
+  by expr_auto
+
+lemma unrest_init_post [unrest]: "mwb_lens x \<Longrightarrow> $x< \<sharp> e>"
+  by expr_auto
+
+lemma unrest_fin_pre [unrest]: "mwb_lens x \<Longrightarrow> $x> \<sharp> e<"
+  by expr_auto
+
+lemma unrest_fin_post [unrest]: "\<lbrakk> mwb_lens x; $x \<sharp> e \<rbrakk> \<Longrightarrow> $x> \<sharp> e>"
+  by expr_auto
+
+lemma aext_get_fst [usubst]: "aext (getx) fstL = getns_alpha fstL x"
+  by expr_simp
+
+lemma aext_get_snd [usubst]: "aext (getx) sndL = getns_alpha sndL x"
+  by expr_simp
+
+subsection \<open> Substitutions \<close>
+
+definition subst_aext :: "'a subst \<Rightarrow> ('a \<Longrightarrow> 'b) \<Rightarrow> 'b subst"
+  where [expr_defs]: "subst_aext \<sigma> x = (\<lambda> s. putx s (\<sigma> (getx s)))"
+
+definition subst_ares :: "'b subst \<Rightarrow> ('a \<Longrightarrow> 'b) \<Rightarrow> 'a subst"
+  where [expr_defs]: "subst_ares \<sigma> x = (\<lambda> s. getx (\<sigma> (createx s)))"
+
+syntax 
+  "_subst_aext" :: "logic \<Rightarrow> svid \<Rightarrow> logic" (infixl "\<up>s" 80)
+  "_subst_ares" :: "logic \<Rightarrow> svid \<Rightarrow> logic" (infixl "\<down>s" 80)
+
+translations
+  "_subst_aext P a" == "CONST subst_aext P a"
+  "_subst_ares P a" == "CONST subst_ares P a"
+
+lemma unrest_subst_aext [unrest]: "x \<bowtie> a \<Longrightarrow> $x \<sharp>s (\<sigma> \<up>s a)"
+  by (expr_simp)
+     (metis lens_indep_def lens_override_def lens_scene.rep_eq scene_override.rep_eq)
+
+lemma subst_id_ext [usubst]:
+  "vwb_lens x \<Longrightarrow> [\<leadsto>] \<up>s x = [\<leadsto>]"
+  by expr_auto
+
+lemma upd_subst_ext [alpha]:
+  "vwb_lens x \<Longrightarrow> \<sigma>(y \<leadsto> e) \<up>s x = (\<sigma> \<up>s x)(x:y \<leadsto> e \<up> x)"
+  by expr_auto
+
+lemma apply_subst_ext [alpha]:
+  "vwb_lens x \<Longrightarrow> (\<sigma> \<dagger> e) \<up> x = (\<sigma> \<up>s x) \<dagger> (e \<up> x)"
+  by (expr_auto)
+
+lemma subst_aext_compose [alpha]: "(\<sigma> \<up>s x) \<up>s y = \<sigma> \<up>s y:x"
+  by (expr_simp)
+
+lemma subst_aext_comp [usubst]:
+  "vwb_lens a \<Longrightarrow> (\<sigma> \<up>s a) \<circ>s (\<rho> \<up>s a) = (\<sigma> \<circ>s \<rho>) \<up>s a"
+  by expr_auto
+
+lemma subst_id_res: "mwb_lens a \<Longrightarrow> [\<leadsto>] \<down>s a = [\<leadsto>]"
+  by expr_auto
+
+lemma upd_subst_res_in: 
+  "\<lbrakk> mwb_lens a; x \<subseteq>L a \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<down>s a = (\<sigma> \<down>s a)(x \<restriction> a \<leadsto> e \<down> a)"
+  by (expr_simp, fastforce)
+
+lemma upd_subst_res_out: 
+  "\<lbrakk> mwb_lens a; x \<bowtie> a \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<down>s a = \<sigma> \<down>s a"
+  by (simp add: expr_defs lens_indep_sym)
+
+lemma subst_ext_lens_apply: "\<lbrakk> mwb_lens a; -$a \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> (a\<up> \<circ>s \<sigma>) \<dagger> P = ((\<sigma> \<down>s a) \<dagger> P) \<up> a"
+  by (expr_simp, simp add: lens_override_def scene_override_commute)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Liberation.thy
--- /dev/null
+++ thys/Shallow_Expressions/Liberation.thy
@@ -0,0 +1,61 @@
+subsection \<open> Liberation \<close>
+
+theory Liberation
+  imports Extension
+begin
+
+text \<open> Liberation~\cite{Dongol19} is an operator that removes dependence on a number of variables.
+  It is similar to existential quantification, but is defined over a scene (a variable set). \<close>
+
+subsection \<open> Definition and Syntax \<close>
+
+definition liberate :: "('s \<Rightarrow> bool) \<Rightarrow> 's scene \<Rightarrow> ('s \<Rightarrow> bool)" where
+[expr_defs]: "liberate P x = (\<lambda> s. \<exists> s'. P (s \<oplus>S s' on x))"
+
+syntax
+  "_liberate" :: "logic \<Rightarrow> salpha \<Rightarrow> logic" (infixl "\\" 80)
+
+translations
+  "_liberate P x" == "CONST liberate P x"
+  "_liberate P x" <= "_liberate (P)e x"
+
+expr_constructor liberate (0)
+
+subsection \<open> Laws \<close>
+
+lemma liberate_lens [expr_simps]: 
+  "mwb_lens x \<Longrightarrow> P \\ $x = (\<lambda>s. \<exists>s'. P (s \<triangleleft>x s'))"
+  by (simp add: liberate_def)
+
+lemma liberate_lens': "mwb_lens x \<Longrightarrow> P \\ $x = (\<lambda>s. \<exists>v. P (putx s v))"
+  by (auto simp add: liberate_def lens_defs fun_eq_iff)
+     (metis mwb_lens_weak weak_lens.put_get)
+
+lemma liberate_as_subst: "vwb_lens x \<Longrightarrow> e \\ $x = (\<exists> v. e\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk>)e"
+  by (expr_simp, metis vwb_lens.put_eq)
+
+lemma unrest_liberate: "a \<sharp> P \\ a"
+  by (expr_simp)
+
+lemma unrest_liberate_iff: "(a \<sharp> P) \<longleftrightarrow> (P \\ a = P)"
+  by (expr_simp, metis (full_types) scene_override_overshadow_left)
+
+lemma liberate_none [simp]: "P \\ \<emptyset> = P"
+  by (expr_simp)
+
+lemma liberate_idem [simp]: "P \\ a \\ a = P \\ a"
+  by (expr_simp)
+
+lemma liberate_commute [simp]: "a \<bowtie>S b \<Longrightarrow> P \\ a \\ b = P \\ b \\ a"
+  using scene_override_commute_indep by (expr_auto, fastforce+)
+
+lemma liberate_true [simp]: "(True)e \\ a = (True)e"
+  by (expr_simp)
+
+lemma liberate_false [simp]: "(False)e \\ a = (False)e"
+  by (expr_simp)
+
+lemma liberate_disj [simp]: "(P \<or> Q)e \\ a = (P \\ a \<or> Q \\ a)e"
+  by (expr_auto)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Lift_Expr.ML
--- /dev/null
+++ thys/Shallow_Expressions/Lift_Expr.ML
@@ -0,0 +1,158 @@
+structure ExprFun = Theory_Data
+  (type T = Symset.T
+   val empty = Symset.empty
+   val merge = Symset.merge);
+
+structure ExprFun_Const =
+struct
+
+fun exprfun_const n ctx =  
+      let open Proof_Context; open Syntax
+      in case read_const {proper = true, strict = false} ctx n of
+         Const (c, _) => Local_Theory.background_theory (ExprFun.map (Symset.insert c)) ctx |
+         _ => raise Match
+      end;
+end;
+
+Outer_Syntax.local_theory @{command_keyword "expr_function"} 
+    "declare that certain constants are functions returning expressions"
+    (Scan.repeat1 (Parse.term )
+     >> (fn ns => 
+         (fn ctx => Library.foldl (fn (ctx, n) => ExprFun_Const.exprfun_const n ctx) (ctx, ns))));
+
+structure ExprCtr = Theory_Data
+  (type T = int list Symtab.table
+   val empty = Symtab.empty
+   val merge = Symtab.merge (K true));
+
+structure ExprCtr_Const =
+struct
+
+fun exprctr_const (n, opt) ctx =  
+      let open Proof_Context; open Syntax
+      in case read_const {proper = true, strict = false} ctx n of
+         Const (c, _) => Local_Theory.background_theory (ExprCtr.map (Symtab.update (c, (map Value.parse_int opt)))) ctx |
+         _ => raise Match
+      end;
+end;
+
+Outer_Syntax.local_theory @{command_keyword "expr_constructor"} 
+    "declare that certain constants are expression constructors; the parameter indicates which arguments should not be lifted"
+    (Scan.repeat1 (Parse.term -- Scan.optional (Parse.$$$ "(" |-- Parse.!!! (Scan.repeat1 Parse.number --| Parse.$$$ ")")) [])
+     >> (fn ns => 
+         (fn ctx => Library.foldl (fn (ctx, n) => ExprCtr_Const.exprctr_const n ctx) (ctx, ns))));
+
+signature LIFT_EXPR =
+sig
+val literal_variables: bool Config.T
+val pretty_print_exprs: bool Config.T
+val mark_state_variables: bool Config.T
+val state_id: string
+val lift_expr: Proof.context -> typ -> term -> term
+val read_lift_term: Proof.context -> typ -> string -> term
+val mk_lift_expr: Proof.context -> typ -> term -> term
+val print_expr: Proof.context -> term -> term
+end
+
+structure Lift_Expr: LIFT_EXPR =
+struct
+
+val literal_variables = Attrib.setup_config_bool binding\<open>literal_variables\<close> (K false);
+val pretty_print_exprs = Attrib.setup_config_bool binding\<open>pretty_print_exprs\<close> (K true);
+val mark_state_variables = Attrib.setup_config_bool binding\<open>mark_state_variables\<close> (K true);
+
+val state_id = "\<s>";
+
+fun lift_expr_aux ctx stT (Const (n', t), args) =
+  let 
+    open Syntax
+    val n = (if (Lexicon.is_marked_entity n') then Lexicon.unmark_const n' else n')
+    val args' = map (lift_expr ctx stT) args
+  in if (n = @{const_name lens_get} orelse n = @{const_name SEXP}) 
+     then Term.list_comb (Const (n', t), args)
+     else if (n = @{syntax_const "_sexp_lit"})
+     then Term.list_comb (hd args, tl args')
+     else if (n = @{syntax_const "_sexp_evar"})
+     then Term.list_comb (hd args $ Free (state_id, stT), tl args')
+     else if Symset.member (ExprFun.get (Proof_Context.theory_of ctx)) n then Term.list_comb (Const (n', t), args') $ Free (state_id, stT) 
+     else if (member (op =) (Symtab.keys (ExprCtr.get (Proof_Context.theory_of ctx))) n)
+     (* FIXME: Allow some parameters of an expression constructor to be ignored and not lifted *) 
+     then let val (SOME aopt) = (Symtab.lookup (ExprCtr.get (Proof_Context.theory_of ctx)) n) in
+          Term.list_comb 
+            (Const (n', t)
+            , map_index (fn (i, t) => 
+                         if (member (op =) aopt i) 
+                         then const @{const_name SEXP} $ (lambda (Free (state_id, stT)) (lift_expr ctx stT t)) 
+                         else t) args) $ Free (state_id, stT)
+          end
+     else
+     case (Type_Infer_Context.const_type ctx n) of
+      (* If it has a lens type, we insert the get function *)
+      SOME (Type (type_name\<open>lens_ext\<close>, _)) 
+        => Term.list_comb (const @{const_name lens_get} $ Const (n', t) $ Free (state_id, stT), args') |
+      (* If the type of the constant has an input of type "'st", we assume it's a state and lift it *)
+      SOME (Type (type_name\<open>fun\<close>, [TFree (a, _), _])) 
+        => if a = fst (dest_TFree Lens_Lib.astateT) 
+           then Term.list_comb (Const (n', t), args) $ Free (state_id, stT)
+           else Term.list_comb (Const (n, t), args') |
+      _ => Term.list_comb (Const (n, t), args')
+  end |
+lift_expr_aux ctx stT (Free (n, t), args) = 
+    let open Syntax
+        val consts = (Proof_Context.consts_of ctx)
+        val {const_space, ...} = Consts.dest consts
+        val t' = case (Syntax.check_term ctx (Free (n, t))) of
+                   Free (_, t') => t' | _ => t
+        \<comment> \<open> The name must be internalised in case it needs qualifying. \<close>
+        val c = Consts.intern consts n in
+        \<comment> \<open> If the name refers to a declared constant, then we lift it as a constant. \<close>
+        if (Name_Space.declared const_space c) then
+          lift_expr_aux ctx stT (Const (c, t), args)
+        \<comment> \<open> Otherwise, we leave it alone \<close>
+        else
+          let val trm = 
+                \<comment> \<open> We check whether the free variable has a lens type, and if so lift it as one\<close>
+                case t' of 
+                 Type (type_name\<open>lens_ext\<close>, _) => const @{const_name lens_get} $ Free (n, t') $ Free (state_id, stT) |
+                 \<comment> \<open> Otherwise, we leave it alone, or apply it to the state variable if it's an expression constructor \<close>
+                 _ => if Config.get ctx literal_variables
+                        then Free (n, t) else Free (n, t) $ Free (state_id, stT)
+          in Term.list_comb (trm, map (lift_expr ctx stT) args) end
+  end |
+lift_expr_aux _ _ (e, args) = Term.list_comb (e, args)
+and 
+lift_expr ctx stT (Abs (n, t, e)) = 
+  if (n = state_id) then Abs (n, t, e) else Abs (n, t, lift_expr ctx stT e) |
+lift_expr ctx stT (Const ("_constrain", a) $ e $ t) = (Const ("_constrain", a) $ lift_expr ctx stT e $ t) |
+lift_expr ctx stT (Const ("_constrainAbs", a) $ e $ t) = (Const ("_constrainAbs", a) $ lift_expr ctx stT e $ t) |
+lift_expr ctx stT e = lift_expr_aux ctx stT (Term.strip_comb e)
+
+fun mk_lift_expr ctx stT e = 
+  lambda (Free (state_id, stT)) (lift_expr ctx stT (Term_Position.strip_positions e));
+
+fun print_expr_aux ctx (f, args) =
+  let open Proof_Context
+      fun sexp_evar x = if Config.get ctx literal_variables then Syntax.const "_sexp_evar" $ x else x
+  in
+  case (f, args) of
+    (Const (@{syntax_const "_free"}, _), (Free (n, t) :: Const (@{syntax_const "_sexp_state"}, _) :: r)) => 
+     sexp_evar (Term.list_comb (Syntax.const @{syntax_const "_free"} $ Free (n, t)
+                              , map (print_expr ctx) r)) |
+    (Const (@{const_syntax lens_get}, _), [x, Const (@{syntax_const "_sexp_state"}, _)])
+      => if Config.get ctx mark_state_variables then Syntax.const @{syntax_const "_sexp_var"} $ x else x |
+    (Free (n, t), [Const (@{syntax_const "_sexp_state"}, _)]) 
+      => sexp_evar (Free (n, t)) | 
+    _ => if (length args > 0)
+         then case (List.last args) of
+           Const (@{syntax_const "_sexp_state"}, _) => 
+             Term.list_comb (f, map (print_expr ctx) (take (length args - 1) args)) |
+           _ => Term.list_comb (f, map (print_expr ctx) args)
+         else Term.list_comb (f, map (print_expr ctx) args)
+   end
+and 
+print_expr ctx (Abs (n, t, e)) = Abs (n, t, print_expr ctx e) |
+print_expr ctx e = print_expr_aux ctx (Term.strip_comb e)
+
+fun read_lift_term ctx stT = Syntax.check_term ctx o mk_lift_expr ctx stT o Syntax.parse_term ctx
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Local_State.thy
--- /dev/null
+++ thys/Shallow_Expressions/Local_State.thy
@@ -0,0 +1,117 @@
+section \<open> Local State \<close>
+
+theory Local_State
+  imports "Expressions"
+begin
+
+text \<open> This theory supports ad-hoc extension of an alphabet type with a tuple of lenses constructed
+  using successive applications of @{const fst_lens} and @{const snd_lens}. It effectively allows
+  local variables, since we always add a collection of new variables. \<close>
+
+text \<open> We declare a number of syntax translations to produce lens and product types, to obtain
+  a type for the overall state space, to construct a tuple that denotes the lens vector parameter, 
+  to construct the vector itself, and finally to construct the state declaration. \<close>
+
+syntax
+  "_lensT" :: "type \<Rightarrow> type \<Rightarrow> type" ("LENSTYPE'(_, _')")
+  "_pairT" :: "type \<Rightarrow> type \<Rightarrow> type" ("PAIRTYPE'(_, _')")
+  "_state_type" :: "pttrn \<Rightarrow> type"
+  "_state_tuple" :: "type \<Rightarrow> pttrn \<Rightarrow> logic"
+  "_state_lenses" :: "pttrn \<Rightarrow> logic \<Rightarrow> logic" ("localstate (_)/ over (_)" [0, 10] 10)
+  "_lvar_abs" :: "id \<Rightarrow> type \<Rightarrow> logic \<Rightarrow> logic"
+
+translations
+  (type) "PAIRTYPE('a, 'b)" => (type) "'a \<times> 'b"
+  (type) "LENSTYPE('a, 'b)" => (type) "'a \<Longrightarrow> 'b"
+
+  "_state_type (_constrain x t)" => "t"
+  "_state_type (CONST Product_Type.Pair (_constrain x t) vs)" => "_pairT t (_state_type vs)"
+
+  "_state_tuple st (_constrain x t)" => "_constrain x (_lensT t st)"
+  "_state_tuple st (CONST Pair (_constrain x t) vs)" =>
+    "CONST Product_Type.Pair (_constrain x (_lensT t st)) (_state_tuple st vs)"
+
+ML \<open> 
+signature LIFT_EXPR_LVAR =
+sig
+val lift_expr_lvar: string -> term -> term
+val lift_lvar: string -> term -> term
+end
+
+structure Lift_Expr_LVar: LIFT_EXPR_LVAR =
+struct
+
+fun lift_expr_lvar n (Free (x, t')) =
+  let open Syntax; open Lift_Expr
+  in
+  if x = n then const @{const_name lens_get} $ Free (n, t') $ Bound 0
+  else Free (x, t')
+  end |
+lift_expr_lvar n (Free (x, t') $ Bound 0) =
+  lift_expr_lvar n (Free (x, t')) |
+lift_expr_lvar n (Abs (x, t', trm)) =
+  if x = n then Abs (x, t', trm) else Abs (x, t', lift_expr_lvar n trm) |
+lift_expr_lvar n (t1 $ t2) = lift_expr_lvar n t1 $ lift_expr_lvar n t2 |
+lift_expr_lvar _ trm = trm 
+
+fun lift_lvar n (Const (@{const_name SEXP}, t') $ e)
+  = (Const (@{const_name SEXP}, t') $ lift_expr_lvar n e) |
+lift_lvar n (t1 $ t2) = lift_lvar n t1 $ lift_lvar n t2 |
+lift_lvar n (Abs (x, t', trm)) =
+  if x = n then Abs (x, t', trm) else Abs (x, t', lift_lvar n trm) |
+lift_lvar _ trm = trm 
+
+end
+\<close>
+parse_translation \<open>
+  let
+    open HOLogic; open Syntax;
+    fun lensT s t = Type (@{type_name lens_ext}, [s, t, HOLogic.unitT]);
+    fun lens_comp a b c = Const (@{const_syntax "lens_comp"}, lensT a b --> lensT b c --> lensT a c);
+    fun fst_lens t = Const (@{const_syntax "fst_lens"}, Type (@{type_name lens_ext}, [t, dummyT, unitT]));
+    val snd_lens = Const (@{const_syntax "snd_lens"}, dummyT);
+    fun id_lens t = Const (@{const_syntax "id_lens"},Type (@{type_name lens_ext}, [t, dummyT, unitT]));
+    fun lens_syn_typ t = const @{type_syntax lens_ext} $ t $ const @{type_syntax dummy} $ const @{type_syntax unit};
+    fun constrain t ty = const @{syntax_const "_constrain"} $ t $ ty;
+
+    (* Construct a tuple of n lenses, whose source type is product of the types in ts, and each lens
+       has an element of the type: prod_lens [t0, t1 ... ] 1 : t1 ==> t0 * t1 * ... *)
+    fun prod_lens ts i = 
+      let open Syntax; open Library; fun lens_compf (x, y) = const @{const_name lens_comp} $ x $ y in
+      if (length ts = 1) 
+      then Const (@{const_name id_lens}, lensT (nth ts i) (nth ts i))
+      else if (length ts = i + 1) 
+      then foldl lens_compf (Const (@{const_name "snd_lens"}, lensT (nth ts i) dummyT), replicate (i-1) (const @{const_name "snd_lens"}))
+      else foldl lens_compf (Const (@{const_name "fst_lens"}, lensT (nth ts i) dummyT), replicate i (const @{const_name "snd_lens"}))
+      end;
+
+    (* Construct a tuple of lenses for each of the possible locally declared variables *)
+    fun state_lenses ts sty st = 
+      foldr1 (fn (x, y) => pair_const dummyT dummyT $ x $ y) (map (fn i => lens_comp dummyT sty dummyT $ prod_lens ts i $ st) (upto (0, length ts - 1)));
+
+    fun
+      (* Add up the number of variable declarations in the tuple *)
+      var_decl_num (Const (@{const_syntax "Product_Type.Pair"},_) $ _ $ vs) = var_decl_num vs + 1 |
+      var_decl_num _ = 1;
+
+    fun
+      var_decl_typs (Const (@{const_syntax "Product_Type.Pair"},_) $ (Const ("_constrain", _) $ _ $ typ) $ vs) = Syntax_Phases.decode_typ typ :: var_decl_typs vs |
+      var_decl_typs (Const ("_constrain", _) $ _ $ typ) = [Syntax_Phases.decode_typ typ] |
+      var_decl_typs _ = [];
+
+    fun state_lens ctx [vs, loc] = (state_lenses (var_decl_typs vs) (mk_tupleT (var_decl_typs vs)) loc);
+  in
+  [("_state_lenses", state_lens),
+   (@{syntax_const "_lvar_abs"}
+   , fn ctx => fn terms => 
+     let val [Free (x, _), tty, trm] = terms 
+         val ty = Syntax_Phases.decode_typ tty
+         val lens = Type (@{type_name "lens_ext"}, [ty, dummyT, @{typ "unit"}])
+         val trm' = Lift_Expr_LVar.lift_lvar x trm
+     in Abs (x, lens, abstract_over (free x, trm')) end)]
+  end
+\<close>
+
+term "localstate (x::int, y::int) over 1L"
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Named_Expressions.thy
--- /dev/null
+++ thys/Shallow_Expressions/Named_Expressions.thy
@@ -0,0 +1,80 @@
+section \<open> Named Expression Definitions \<close>
+
+theory Named_Expressions
+  imports Expressions
+  keywords "edefinition" "expression" :: "thy_decl_block" and "over"
+begin
+
+text \<open> Here, we add a command that allows definition of a named expression. It provides a more
+  concise version of @{command definition} and inserts the expression brackets. \<close>
+
+ML \<open>
+structure Expr_Def =
+struct
+
+  fun mk_expr_def_eq ctx term =
+    case (Type.strip_constraints term) of
+      Const (@{const_name "HOL.eq"}, b) $ c $ t => 
+        (fst (dest_Free (fst (Term.strip_comb c))),
+        @{const Trueprop} $ (Const (@{const_name "HOL.eq"}, b) $ c $ (Syntax.const @{const_name SEXP} 
+            $ (lambda (Syntax.free Lift_Expr.state_id) 
+                      (Lift_Expr.lift_expr ctx dummyT (Term_Position.strip_positions t)))))) |
+      _ => raise Match;
+
+  val expr_defs = [[Token.make_string (Binding.name_of @{binding expr_defs}, Position.none)]];
+
+  fun expr_def attr decl term ctx =
+  let val named_expr_defs = @{attributes [named_expr_defs]}
+      val (n, eq) = mk_expr_def_eq ctx term
+      val (thm, ctx0) = Specification.definition 
+                   (Option.map (fn x => fst (Proof_Context.read_var x ctx)) decl) [] [] 
+                   ((fst attr, map (Attrib.check_src ctx) (named_expr_defs @ snd attr)), eq) (snd (Local_Theory.begin_nested ctx))
+      val ctx1 = ExprFun_Const.exprfun_const n (Local_Theory.end_nested ctx0)
+  in (thm, ctx1)
+  end
+
+  fun named_expr n typ stateT expr ctx =
+    let val named_expr_defs = @{attributes [named_expr_defs]}
+        val term = Const (@{const_name "HOL.eq"}, dummyT) $ Syntax.free n $ expr
+        val ctx' = snd (Specification.definition 
+                       (SOME (Binding.name n, SOME (stateT --> typ), Mixfix.NoSyn)) [] [] 
+                       ((Binding.name (n ^ "_def"), named_expr_defs), snd (mk_expr_def_eq ctx term)) (snd (Local_Theory.begin_nested ctx)))
+        (* When adding an expression in a locale, the named recorded below may be the 
+           localised version, which may not work correctly. This may lead to unexpected
+           behaviour when there are two locales each with a constant of the same name. *)
+        val ctx2 = ExprFun_Const.exprfun_const n (Local_Theory.end_nested ctx')
+        in ctx2
+  end
+
+
+  fun named_expr_cmd n t st expr ctx =
+    let val term = Syntax.parse_term ctx expr
+        val stateT = Syntax.read_typ ctx st
+        val typ = Syntax.read_typ ctx t
+        in named_expr n typ stateT term ctx
+  end
+
+end;
+
+val _ =
+let
+  open Expr_Def;
+in
+  Outer_Syntax.local_theory command_keyword\<open>edefinition\<close> "UTP constant definition"
+    (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) --
+      Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, (attr, term)), _), _) =>
+        (fn ctx => snd (expr_def attr decl (Syntax.parse_term ctx term) ctx))))
+end
+
+val _ =
+let
+  open Expr_Def;
+in
+  Outer_Syntax.local_theory command_keyword\<open>expression\<close> "define named expressions"
+    ((((Parse.name -- Scan.optional (@{keyword "::"} |-- Parse.typ) "_" -- Scan.optional (@{keyword "over"} |-- Parse.typ) "_") --| @{keyword "is"}) -- Parse.term)  >> (fn (((n, t), st), expr) => 
+        (named_expr_cmd n t st expr)))
+end;
+
+\<close>
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Quantifiers.thy
--- /dev/null
+++ thys/Shallow_Expressions/Quantifiers.thy
@@ -0,0 +1,127 @@
+subsection \<open> Quantifying Lenses \<close>
+
+theory Quantifiers
+  imports Liberation
+begin
+
+text \<open> We define operators to existentially and universally quantify an expression over a lens. \<close>
+
+subsection \<open> Operators and Syntax \<close>
+
+definition ex_expr :: "('a \<Longrightarrow> 's) \<Rightarrow> (bool, 's) expr \<Rightarrow> (bool, 's) expr" where
+[expr_defs]: "ex_expr x e = (\<lambda> s. (\<exists> v. e (putx s v)))"
+
+definition ex1_expr :: "('a \<Longrightarrow> 's) \<Rightarrow> (bool, 's) expr \<Rightarrow> (bool, 's) expr" where
+[expr_defs]: "ex1_expr x e = (\<lambda> s. (\<exists>! v. e (putx s v)))"
+
+definition all_expr :: "('a \<Longrightarrow> 's) \<Rightarrow> (bool, 's) expr \<Rightarrow> (bool, 's) expr" where
+[expr_defs]: "all_expr x e = (\<lambda> s. (\<forall> v. e (putx s v)))"
+
+expr_constructor ex_expr (1)
+expr_constructor ex1_expr (1)
+expr_constructor all_expr (1)
+
+syntax 
+  "_ex_expr"  :: "svid \<Rightarrow> logic \<Rightarrow> logic" ("\<exists> _ \<Zspot> _" [0, 20] 20)
+  "_ex1_expr" :: "svid \<Rightarrow> logic \<Rightarrow> logic" ("\<exists>1 _ \<Zspot> _" [0, 20] 20)
+  "_all_expr" :: "svid \<Rightarrow> logic \<Rightarrow> logic" ("\<forall> _ \<Zspot> _" [0, 20] 20)
+
+translations
+  "_ex_expr x P" == "CONST ex_expr x P"
+  "_ex1_expr x P" == "CONST ex1_expr x P"
+  "_all_expr x P" == "CONST all_expr x P"
+
+subsection \<open> Laws \<close>
+
+lemma ex_is_liberation: "mwb_lens x \<Longrightarrow> (\<exists> x \<Zspot> P) = (P \\ $x)"
+  by (expr_auto, metis mwb_lens_weak weak_lens.put_get)
+
+lemma ex_unrest_iff: "\<lbrakk> mwb_lens x \<rbrakk> \<Longrightarrow> ($x \<sharp> P) \<longleftrightarrow> (\<exists> x \<Zspot> P) = P"
+  by (simp add: ex_is_liberation unrest_liberate_iff)
+
+lemma ex_unrest: "\<lbrakk> mwb_lens x; $x \<sharp> P \<rbrakk> \<Longrightarrow> (\<exists> x \<Zspot> P) = P"
+  using ex_unrest_iff by blast
+
+lemma unrest_ex_in [unrest]:
+  "\<lbrakk> mwb_lens y; x \<subseteq>L y \<rbrakk> \<Longrightarrow> $x \<sharp> (\<exists> y \<Zspot> P)"
+  by (simp add: ex_expr_def sublens_pres_mwb sublens_put_put unrest_lens)
+
+lemma unrest_ex_out [unrest]:
+  "\<lbrakk> mwb_lens x; $x \<sharp> P; x \<bowtie> y \<rbrakk> \<Longrightarrow> $x \<sharp> (\<exists> y \<Zspot> P)"
+  by (simp add: ex_expr_def unrest_lens, metis lens_indep.lens_put_comm)
+
+lemma subst_ex_out [usubst]: "\<lbrakk> mwb_lens x; $x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> \<sigma> \<dagger> (\<exists> x \<Zspot> P) = (\<exists> x \<Zspot> \<sigma> \<dagger> P)"
+  by (expr_simp)
+
+lemma subst_lit_ex_indep [usubst]:
+  "y \<bowtie> x \<Longrightarrow> \<sigma>(y \<leadsto> \<guillemotleft>v\<guillemotright>) \<dagger> (\<exists> x \<Zspot> P) = \<sigma> \<dagger> (\<exists> x \<Zspot> [y \<leadsto> \<guillemotleft>v\<guillemotright>] \<dagger> P)"
+  by (expr_simp, simp add: lens_indep.lens_put_comm)
+
+lemma subst_ex_in [usubst]:
+  "\<lbrakk> vwb_lens a; x \<subseteq>L a \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<dagger> (\<exists> a \<Zspot> P) = \<sigma> \<dagger> (\<exists> a \<Zspot> P)"
+  by (expr_simp, force)
+
+declare lens_plus_right_sublens [simp]
+
+lemma ex_as_subst: "vwb_lens x \<Longrightarrow> (\<exists> x \<Zspot> e) = (\<exists> v. e\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk>)e"
+  by expr_auto
+
+lemma ex_twice [simp]: "mwb_lens x \<Longrightarrow> (\<exists> x \<Zspot> \<exists> x \<Zspot> P) = (\<exists> x \<Zspot> P)"
+  by (expr_simp)
+
+lemma ex_commute: "x \<bowtie> y \<Longrightarrow> (\<exists> x \<Zspot> \<exists> y \<Zspot> P) = (\<exists> y \<Zspot> \<exists> x \<Zspot> P)"
+  by (expr_auto, metis lens_indep_comm)+
+  
+lemma ex_true [simp]: "(\<exists> x \<Zspot> (True)e) = (True)e"
+  by expr_simp
+
+lemma ex_false [simp]: "(\<exists> x \<Zspot> (False)e) = (False)e"
+  by (expr_simp)
+
+lemma ex_disj [simp]: "(\<exists> x \<Zspot> (P \<or> Q)e) = ((\<exists> x \<Zspot> P) \<or> (\<exists> x \<Zspot> Q))e"
+  by (expr_auto)
+
+lemma ex_plus:
+  "(\<exists> (y,x) \<Zspot> P) = (\<exists> x \<Zspot> \<exists> y \<Zspot> P)"
+  by (expr_auto)
+
+lemma all_as_ex: "(\<forall> x \<Zspot> P) = (\<not> (\<exists> x \<Zspot> \<not> P))e"
+  by (expr_auto)
+
+lemma ex_as_all: "(\<exists> x \<Zspot> P) = (\<not> (\<forall> x \<Zspot> \<not> P))e"
+  by (expr_auto)
+
+subsection \<open> Cylindric Algebra \<close>
+
+lemma ex_C1: "(\<exists> x \<Zspot> (False)e) = (False)e"
+  by (expr_auto)
+
+lemma ex_C2: "wb_lens x \<Longrightarrow> `P \<longrightarrow> (\<exists> x \<Zspot> P)`"
+  by (expr_simp, metis wb_lens.get_put)
+
+lemma ex_C3: "mwb_lens x \<Longrightarrow> (\<exists> x \<Zspot> (P \<and> (\<exists> x \<Zspot> Q)))e = ((\<exists> x \<Zspot> P) \<and> (\<exists> x \<Zspot> Q))e"
+  by (expr_auto)
+
+lemma ex_C4a: "x \<approx>L y \<Longrightarrow> (\<exists> x \<Zspot> \<exists> y \<Zspot> P) = (\<exists> y \<Zspot> \<exists> x \<Zspot> P)"
+  by (expr_simp, metis (mono_tags, lifting) lens.select_convs(2))
+
+lemma ex_C4b: "x \<bowtie> y \<Longrightarrow> (\<exists> x \<Zspot> \<exists> y \<Zspot> P) = (\<exists> y \<Zspot> \<exists> x \<Zspot> P)"
+  using ex_commute by blast
+
+lemma ex_C5:
+  fixes x :: "('a \<Longrightarrow> '\<alpha>)"
+  shows "($x = $x)e = (True)e"
+  by simp
+
+lemma ex_C6:
+  assumes "wb_lens x" "x \<bowtie> y" "x \<bowtie> z"
+  shows "($y = $z)e = (\<exists> x \<Zspot> $y = $x \<and> $x = $z)e"
+  using assms
+  by (expr_simp, metis lens_indep_def)
+
+lemma ex_C7:
+  assumes "weak_lens x" "x \<bowtie> y"
+  shows "((\<exists> x \<Zspot> $x = $y \<and> P) \<and> (\<exists> x \<Zspot> $x = $y \<and> \<not> P))e = (False)e"
+  using assms by (expr_simp, simp add: lens_indep_sym)
+
+end
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/ROOT
--- /dev/null
+++ thys/Shallow_Expressions/ROOT
@@ -0,0 +1,14 @@
+chapter AFP
+
+session "Shallow_Expressions" = "Optics" +
+  options [timeout = 600]
+  sessions
+    "HOL-Library"
+  theories 
+    Shallow_Expressions
+    Expressions_Tests
+    Shallow_Expressions_Examples
+  document_files
+    "root.bib"
+    "root.tex"
+    "document.sty"
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Shallow_Expressions.thy
--- /dev/null
+++ thys/Shallow_Expressions/Shallow_Expressions.thy
@@ -0,0 +1,8 @@
+section \<open> Shallow Expressions Meta-Theory \<close>
+
+theory Shallow_Expressions
+  imports 
+    Variables Expressions Unrestriction Substitutions Extension Liberation Quantifiers Collections
+    Named_Expressions Local_State
+begin end
+
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Shallow_Expressions_Examples.thy
--- /dev/null
+++ thys/Shallow_Expressions/Shallow_Expressions_Examples.thy
@@ -0,0 +1,186 @@
+section \<open> Examples of Shallow Expressions \<close>
+
+theory Shallow_Expressions_Examples
+  imports Shallow_Expressions
+begin
+
+subsection \<open> Basic Expressions and Queries \<close>
+
+text \<open> We define some basic variables using the @{command alphabet} command, process some simple
+  expressions, and then perform some unrestriction queries and substitution transformations. \<close>
+
+declare [[literal_variables]]
+
+alphabet st =
+  v1 :: int
+  v2 :: int
+  v3 :: string
+
+term "(v1 > a)e"
+
+declare [[pretty_print_exprs=false]]
+
+term "(v1 > a)e"
+
+declare [[pretty_print_exprs]]
+
+lemma "$v2 \<sharp> (v1 > 5)e"
+  by unrest
+
+lemma "(v1 > 5)e\<lbrakk>v2/v1\<rbrakk> = (v2 > 5)e"
+  by subst_eval 
+
+text \<open> We sometimes would like to define ``constructors'' for expressions. These are functions that
+  produce expressions, and may also have expressions as arguments. Unlike for other functions,
+  during lifting the state is not passed to the arguments, but is passed to the constructor
+  constant itself. An example is given below: \<close>
+
+definition v1_greater :: "int \<Rightarrow> (bool, st) expr" where
+"v1_greater x = (v1 > x)e"
+
+expr_constructor v1_greater
+
+text \<open> Definition @{const v1_greater} is a constructor for an expression, and so it should not
+  be lifted. Therefore we use the command @{command expr_constructor} to specify this, which
+  modifies the behaviour of the lifting parser, and means that @{term "(v1_greater 7)e"} is 
+  correctly translated. 
+
+  If it is desired that one or more of the arguments is an expression, then this can be specified 
+  using an optional list of numbers. In the example below, the first argument is an expression. \<close>
+
+definition v1_greater' :: "(int, st) expr \<Rightarrow> (bool, st) expr" where
+"v1_greater' x = (v1 > @x)e"
+
+expr_constructor v1_greater' (0)
+
+term "(v1_greater' (v1 + 1))e"
+
+text \<open> We also sometimes wish to have functions that return expressions, whose arguments should be 
+  lifted. We can achieve this using the @{command expr_function} command: \<close>
+
+definition v1_less :: "int \<Rightarrow> (bool, st) expr" where
+"v1_less x = (v1 < x)e"
+
+expr_function v1_less
+
+text \<open> This means, we can parse terms like @{term "(v1_less (v1 + 1))e"} -- notice that this returns
+  an expression and takes an expression as an input. Alternatively, we can achieve the same effect
+  with the @{command edefinition} command, which is like @{command definition}, but uses the expression
+  parse and lifts the arguments as expressions. It is typically used for user-level functions that
+  depend on the state. \<close>
+
+edefinition v1_less' where "v1_less' x = (v1 < x)"
+
+term "(v1_less' (v1 + 1))e"
+
+text \<open> In addition, we can define an expression using the command below, which automatically performs 
+  expression lifting in the defining term. These constants are also set as expression constructors. \<close>
+
+expression v1_is_big over st is "v1 > 100"
+
+expression inc_v1 over "st \<times> st" is "v1> = v1< + 1"
+
+text \<open> Definitional equations for named expressions are collected in the theorem attribute 
+  @{thm named_expr_defs}. \<close>
+
+thm named_expr_defs
+
+subsection \<open> Hierarchical State \<close>
+
+alphabet person =
+  name :: string
+  age  :: nat
+
+alphabet company =
+  adam :: person
+  bella :: person
+  carol :: person
+
+term "($adam:age > $carol:age)e"
+
+term "($adam:name \<noteq> $bella:name)e"
+
+subsection \<open> Program Semantics \<close>
+
+text \<open> We give a predicative semantics to a simple imperative programming language with sequence,
+  conditional, and assignment, using lenses and shallow expressions. We then use these definitions
+  to prove some basic laws of programming. \<close>
+
+declare [[literal_variables=false]]
+
+type_synonym 's prog = "'s \<times> 's \<Rightarrow> bool"
+
+definition seq :: "'s prog \<Rightarrow> 's prog \<Rightarrow> 's prog" (infixr ";;" 85) where
+[expr_defs]: "seq P Q = (\<exists> s. P\<lbrakk>\<guillemotleft>s\<guillemotright>/v>\<rbrakk> \<and> Q\<lbrakk>\<guillemotleft>s\<guillemotright>/v<\<rbrakk>)e"
+
+definition ifthenelse :: "(bool, 's) expr \<Rightarrow> 's prog \<Rightarrow> 's prog \<Rightarrow> 's prog" where
+[expr_defs]: "ifthenelse b P Q = (if b< then P else Q)e"
+
+definition assign :: "('a \<Longrightarrow> 's) \<Rightarrow> ('a, 's) expr \<Rightarrow> 's prog"  where
+[expr_defs]: "assign x e = ($x> = e< \<and> v> \<simeq>\<guillemotleft>x\<guillemotright> v<)e"
+
+syntax 
+  "_assign" :: "svid \<Rightarrow> logic \<Rightarrow> logic" ("_ ::= _" [86, 87] 87)
+  "_ifthenelse" :: "logic \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" ("IF _ THEN _ ELSE _" [0, 0, 84] 84)
+
+text \<open> The syntax translations insert the expression brackets, which means the expressions
+  are lifted, without this being visible to the user. \<close>
+
+translations 
+  "_assign x e" == "CONST assign x (e)e"
+  "_ifthenelse b P Q" == "CONST ifthenelse (b)e P Q"
+
+lemma seq_assoc: "P ;; (Q ;; R) = (P ;; Q) ;; R"
+  by expr_auto
+
+lemma ifthenelse_seq_distr: "(IF B THEN P ELSE Q) ;; R = IF B THEN P ;; R ELSE Q ;; R"
+  by expr_auto
+
+lemma assign_twice:
+  assumes "mwb_lens x"
+  shows "x ::= e ;; x ::= f = x ::= f\<lbrakk>e/x\<rbrakk>"
+  using assms
+  apply expr_simp
+  apply (metis mwb_lens.put_put mwb_lens_weak weak_lens.put_get)
+  done
+
+lemma assign_commute:
+  assumes "mwb_lens x" "mwb_lens y" "x \<bowtie> y" "$y \<sharp> (e)e" "$x \<sharp> (f)e"
+  shows "(x ::= e ;; y ::= f) = (y ::= f ;; x ::= e)"
+  using assms
+  apply expr_simp
+  apply safe
+  apply (metis lens_indep_def mwb_lens_weak weak_lens.put_get)+
+  done
+
+lemma assign_combine:
+  assumes "mwb_lens x" "mwb_lens y" "x \<bowtie> y" "$x \<sharp> (f)e"
+  shows "(x ::= e ;; y ::= f) = (x, y) ::= (e, f)"
+  using assms
+  apply expr_simp
+  apply safe
+  apply (simp_all add: lens_indep.lens_put_comm)
+  apply (metis mwb_lens_weak weak_lens.put_get)
+  done
+
+text \<open> Below, we apply the assignment commutativity law in a small example: \<close>
+
+declare [[literal_variables]]
+
+lemma assign_commute_example: 
+  "adam:name ::= ''Adam'' ;; bella:name ::= ''Bella'' = 
+   bella:name ::= ''Bella'' ;; adam:name ::= ''Adam''"
+proof (rule assign_commute)
+  \<comment> \<open> We show the two variables satisfy the lens axioms \<close>
+  show "mwb_lens (adam:name)v" by simp
+  show "mwb_lens (bella:name)v" by simp
+
+  \<comment> \<open> We show the two variables are independent \<close>
+  show "(adam:name)v \<bowtie> (bella:name)v" by simp
+
+  \<comment> \<open> We show that neither assigned expression depends on the opposite variable \<close>
+  show "$bella:name \<sharp> (''Adam'')e" by unrest
+  show "$adam:name \<sharp> (''Bella'')e" by unrest
+qed
+  
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Substitutions.thy
--- /dev/null
+++ thys/Shallow_Expressions/Substitutions.thy
@@ -0,0 +1,437 @@
+section \<open> Substitutions \<close>
+
+theory Substitutions
+  imports Unrestriction
+begin
+
+subsection \<open> Types and Constants \<close>
+
+text \<open> A substitution is simply a function between two state spaces. Typically,
+  they are used to express mappings from variables to values (e.g. assignments). \<close>
+
+type_synonym ('s1, 's2) psubst = "'s1 \<Rightarrow> 's2"
+type_synonym 's subst = "'s \<Rightarrow> 's"
+
+text \<open> There are different ways of constructing an empty substitution. \<close>
+
+definition subst_id :: "'s subst" ("[\<leadsto>]") 
+  where [expr_defs, code_unfold]: "subst_id = (\<lambda>s. s)"
+
+definition subst_nil :: "('s1, 's2) psubst" ("\<lparr>\<leadsto>\<rparr>") 
+  where [expr_defs, code_unfold]: "\<lparr>\<leadsto>\<rparr> = (\<lambda> s. undefined)"
+
+definition subst_default :: "('s1, 's2::default) psubst" ("\<lblot>\<leadsto>\<rblot>") 
+  where [expr_defs, code_unfold]: "\<lblot>\<leadsto>\<rblot> = (\<lambda> s. default)"
+
+text \<open> We can update a substitution by adding a new variable maplet. \<close>
+
+definition subst_upd :: "('s1, 's2) psubst \<Rightarrow> ('a \<Longrightarrow> 's2) \<Rightarrow> ('a, 's1) expr \<Rightarrow> ('s1, 's2) psubst"
+  where [expr_defs, code_unfold]: "subst_upd \<sigma> x e = (\<lambda> s. putx (\<sigma> s) (e s))"
+
+text \<open> The next two operators extend and restrict the alphabet of a substitution. \<close>
+
+definition subst_ext :: "('s1 \<Longrightarrow> 's2) \<Rightarrow> ('s2, 's1) psubst" ("_\<up>" [999] 999) where
+[expr_defs, code_unfold]: "subst_ext a = geta"
+
+definition subst_res :: "('s1 \<Longrightarrow> 's2) \<Rightarrow> ('s1, 's2) psubst" ("_\<down>" [999] 999) where
+[expr_defs, code_unfold]: "subst_res a = createa"
+
+text \<open> Application of a substitution to an expression is effectively function composition. \<close>
+
+definition subst_app :: "('s1, 's2) psubst \<Rightarrow> ('a, 's2) expr \<Rightarrow> ('a, 's1) expr" 
+  where [expr_defs]: "subst_app \<sigma> e = (\<lambda> s. e (\<sigma> s))"
+
+abbreviation "aext P a \<equiv> subst_app (a\<up>) P"
+abbreviation "ares P a \<equiv> subst_app (a\<down>) P"
+
+text \<open> We can also lookup the expression a variable is mapped to. \<close>
+
+definition subst_lookup :: "('s1, 's2) psubst \<Rightarrow> ('a \<Longrightarrow> 's2) \<Rightarrow> ('a, 's1) expr" ("\<langle>_\<rangle>s")
+  where [expr_defs, code_unfold]: "\<langle>\<sigma>\<rangle>s x = (\<lambda> s. getx (\<sigma> s))"
+
+definition subst_comp :: "('s1, 's2) psubst \<Rightarrow> ('s3, 's1) psubst \<Rightarrow> ('s3, 's2) psubst" (infixl "\<circ>s" 55) 
+    where [expr_defs, code_unfold]: "subst_comp = comp"
+
+definition unrest_usubst :: "'s scene \<Rightarrow> 's subst \<Rightarrow> bool" 
+  where [expr_defs]: "unrest_usubst a \<sigma> = (\<forall> s s'. \<sigma> (s \<oplus>S s' on a) = (\<sigma> s) \<oplus>S s' on a)"
+
+definition par_subst :: "'s subst \<Rightarrow> 's scene \<Rightarrow> 's scene \<Rightarrow> 's subst \<Rightarrow> 's subst"
+  where [expr_defs]: "par_subst \<sigma>1 A B \<sigma>2 = (\<lambda> s. (s \<oplus>S (\<sigma>1 s) on A) \<oplus>S (\<sigma>2 s) on B)"
+
+definition subst_restrict :: "'s subst \<Rightarrow> 's scene \<Rightarrow> 's subst" where 
+[expr_defs]: "subst_restrict \<sigma> a = (\<lambda> s. s \<oplus>S \<sigma> s on a)"
+
+text \<open> Create a substitution that copies the region from the given scene from a given state.
+  This is used primarily in calculating unrestriction conditions. \<close>
+
+definition sset :: "'s scene \<Rightarrow> 's \<Rightarrow> 's subst" 
+  where [expr_defs, code_unfold]: "sset a s' = (\<lambda> s. s \<oplus>S s' on a)"
+
+syntax "_sset" :: "salpha \<Rightarrow> logic \<Rightarrow> logic" ("sset[_, _]")
+translations "_sset a P" == "CONST sset a P"
+
+subsection \<open> Syntax Translations \<close>
+
+nonterminal uexprs and smaplet and smaplets
+
+syntax
+  "_subst_app" :: "logic \<Rightarrow> logic \<Rightarrow> logic" (infix "\<dagger>" 65)
+  "_smaplet"        :: "[svid, logic] => smaplet" ("_ \<leadsto> _")
+  ""                :: "smaplet => smaplets" ("_")
+  "_SMaplets"       :: "[smaplet, smaplets] => smaplets" ("_,/ _")
+  \<comment> \<open> A little syntax utility to extract a list of variable identifiers from a substitution \<close>
+  "_smaplets_svids" :: "smaplets \<Rightarrow> logic"
+  "_SubstUpd"       :: "[logic, smaplets] => logic" ("_/'(_')" [900,0] 900)
+  "_Subst"          :: "smaplets => logic" ("(1[_])")
+  "_PSubst"         :: "smaplets => logic" ("(1\<lparr>_\<rparr>)")
+  "_DSubst"         :: "smaplets \<Rightarrow> logic" ("(1\<lblot>_\<rblot>)")
+  "_psubst"         :: "[logic, svids, uexprs] \<Rightarrow> logic"
+  "_subst"          :: "logic \<Rightarrow> uexprs \<Rightarrow> svids \<Rightarrow> logic" ("(_\<lbrakk>_'/_\<rbrakk>)" [990,0,0] 991)
+  "_uexprs"         :: "[logic, uexprs] => uexprs" ("_,/ _")
+  ""                :: "logic => uexprs" ("_")
+  "_par_subst"      :: "logic \<Rightarrow> salpha \<Rightarrow> salpha \<Rightarrow> logic \<Rightarrow> logic" ("_ [_|_]s _" [100,0,0,101] 101)
+  "_subst_restrict" :: "logic \<Rightarrow> salpha \<Rightarrow> logic" (infixl "\<restriction>s" 85)
+  "_unrest_usubst"  :: "salpha \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (infix "\<sharp>s" 20)
+
+translations
+  "_subst_app \<sigma> e"                    == "CONST subst_app \<sigma> e"
+  "_subst_app \<sigma> e"                    <= "_subst_app \<sigma> (e)e"
+  "_SubstUpd m (_SMaplets xy ms)"     == "_SubstUpd (_SubstUpd m xy) ms"
+  "_SubstUpd m (_smaplet x y)"        == "CONST subst_upd m x (y)e"
+  "_Subst ms"                         == "_SubstUpd [\<leadsto>] ms"
+  "_Subst (_SMaplets ms1 ms2)"        <= "_SubstUpd (_Subst ms1) ms2"
+  "_PSubst ms"                        == "_SubstUpd \<lparr>\<leadsto>\<rparr> ms"
+  "_PSubst (_SMaplets ms1 ms2)"       <= "_SubstUpd (_PSubst ms1) ms2"
+  "_DSubst ms"                        == "_SubstUpd \<lblot>\<leadsto>\<rblot> ms"
+  "_DSubst (_SMaplets ms1 ms2)"       <= "_SubstUpd (_DSubst ms1) ms2"
+  "_SMaplets ms1 (_SMaplets ms2 ms3)" <= "_SMaplets (_SMaplets ms1 ms2) ms3"
+  "_smaplets_svids (_SMaplets (_smaplet x e) ms)" => "x +L (_smaplets_svids ms)"
+  "_smaplets_svids (_smaplet x e)" => "x"
+  "_subst P es vs" => "CONST subst_app (_psubst [\<leadsto>] vs es) P"
+  "_psubst m (_svid_list x xs) (_uexprs v vs)" => "_psubst (_psubst m x v) xs vs"
+  "_psubst m x v"  => "CONST subst_upd m x (v)e"
+  "_subst P v x" <= "CONST subst_app (CONST subst_upd [\<leadsto>] x (v)e) P"
+  "_subst P v x" <= "_subst_app (_Subst (_smaplet x v)) P"
+  "_subst P v x" <= "_subst (_sexp_quote P)  v x"
+  "_subst P v (_svid_tuple (_of_svid_list (x +L y)))" <= "_subst P v (x +L y)"
+  "_par_subst \<sigma>1 A B \<sigma>2" == "CONST par_subst \<sigma>1 A B \<sigma>2"
+  "_subst_restrict \<sigma> a" == "CONST subst_restrict \<sigma> a"
+  "_unrest_usubst x p" == "CONST unrest_usubst x p"
+  "_unrest_usubst (_salphaset (_salphamk (x +L y))) P"  <= "_unrest_usubst (x +L y) P"
+
+expr_constructor subst_app (1) \<comment> \<open> Only the second parameter (1) should be treated as a lifted expression. \<close> 
+expr_constructor subst_id
+expr_constructor subst_nil
+expr_constructor subst_default
+expr_constructor subst_upd
+expr_constructor subst_lookup
+
+ML_file \<open>Expr_Util.ML\<close>
+
+subsection \<open> Substitution Laws \<close>
+
+named_theorems usubst and usubst_eval
+
+lemma subst_id_apply [usubst]: "[\<leadsto>] \<dagger> P = P"
+  by (expr_auto)
+
+lemma subst_unrest [usubst]:
+  "\<lbrakk> vwb_lens x; $x \<sharp> v \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<dagger> v = \<sigma> \<dagger> v"
+  by expr_auto
+
+lemma subst_lookup_id [usubst]: "\<langle>[\<leadsto>]\<rangle>s x = [var x]e"
+  by expr_simp
+
+lemma subst_lookup_aext [usubst]: "\<langle>a\<up>\<rangle>s x = [getns_alpha a x]e"
+  by expr_auto
+
+lemma subst_id_var: "[\<leadsto>] = ($v)e"
+  by expr_simp
+
+lemma subst_upd_id_lam [usubst]: "subst_upd ($v)e x v = subst_upd [\<leadsto>] x v"
+  by expr_simp
+
+lemma subst_id [simp]: "[\<leadsto>] \<circ>s \<sigma> = \<sigma>" "\<sigma> \<circ>s [\<leadsto>] = \<sigma>"
+  by expr_auto+
+
+lemma subst_default_id [simp]: "\<lblot>\<leadsto>\<rblot> \<circ>s \<sigma> = \<lblot>\<leadsto>\<rblot>"
+  by (simp add: expr_defs comp_def)
+
+lemma subst_lookup_one_lens [usubst]: "\<langle>\<sigma>\<rangle>s 1L = \<sigma>"
+  by expr_simp
+
+text \<open> The following law can break expressions abstraction, so it is not by default a 
+  "usubst" law. \<close>
+
+lemma subst_apply_SEXP: "subst_app \<sigma> [e]e = [subst_app \<sigma> e]e"
+  by expr_simp
+
+lemma subst_apply_twice [usubst]: 
+  "\<rho> \<dagger> (\<sigma> \<dagger> e) = (\<sigma> \<circ>s \<rho>) \<dagger> e"
+  by expr_simp
+
+lemma subst_apply_twice_SEXP [usubst]: 
+  "\<rho> \<dagger> [\<sigma> \<dagger> e]e = (\<sigma> \<circ>s \<rho>) \<dagger> [e]e"
+  by expr_simp
+
+(* FIXME: Figure out how to make laws like this parse and simplify *)
+
+term "(f (\<sigma> \<dagger> e))e"
+
+term "(\<forall> x. x + $y > $z)e"
+
+term "(\<forall> k. P\<lbrakk>\<guillemotleft>k\<guillemotright>/x\<rbrakk>)e"
+
+lemma subst_get [usubst]: "\<sigma> \<dagger> getx = \<langle>\<sigma>\<rangle>s x"
+  by (simp add: expr_defs)
+
+lemma subst_var [usubst]: "\<sigma> \<dagger> ($x)e = \<langle>\<sigma>\<rangle>s x"
+  by (simp add: expr_defs)
+
+text \<open> We can't use this as simplification unfortunately as the expression structure is too
+  ambiguous to support automatic rewriting. \<close>
+
+lemma subst_uop: "\<sigma> \<dagger> (\<guillemotleft>f\<guillemotright> e)e = (\<guillemotleft>f\<guillemotright> (\<sigma> \<dagger> e))e"
+  by (simp add: expr_defs)
+
+lemma subst_bop: "\<sigma> \<dagger> (\<guillemotleft>f\<guillemotright> e1 e2)e = (\<guillemotleft>f\<guillemotright> (\<sigma> \<dagger> e1) (\<sigma> \<dagger> e2))e"
+  by (simp add: expr_defs)
+
+lemma subst_lit [usubst]: "\<sigma> \<dagger> (\<guillemotleft>v\<guillemotright>)e = (\<guillemotleft>v\<guillemotright>)e"
+  by (expr_simp)
+
+lemmas subst_basic_ops [usubst] =
+  subst_bop[where f=conj] 
+  subst_bop[where f=disj] 
+  subst_bop[where f=implies] 
+  subst_uop[where f=Not]
+  subst_bop[where f=HOL.eq] 
+  subst_bop[where f=less]
+  subst_bop[where f=less_eq]
+  subst_bop[where f=Set.member] 
+  subst_bop[where f=inf]
+  subst_bop[where f=sup]
+  subst_bop[where f=Pair]
+
+text \<open> A substitution update naturally yields the given expression. \<close>
+    
+lemma subst_lookup_upd [usubst]:
+  assumes "weak_lens x"
+  shows "\<langle>\<sigma>(x \<leadsto> v)\<rangle>s x = (v)e"
+  using assms by (simp add: expr_defs)
+
+lemma subst_lookup_upd_diff [usubst]:
+  assumes "x \<bowtie> y"
+  shows "\<langle>\<sigma>(y \<leadsto> v)\<rangle>s x = \<langle>\<sigma>\<rangle>s x"
+  using assms by (simp add: expr_defs)
+
+lemma subst_lookup_pair [usubst]: 
+  "\<langle>\<sigma>\<rangle>s (x +L y) = ((\<langle>\<sigma>\<rangle>s x, \<langle>\<sigma>\<rangle>s y))e"
+  by (expr_simp)
+
+text \<open> Substitution update is idempotent. \<close>
+
+lemma usubst_upd_idem [usubst]:
+  assumes "mwb_lens x"
+  shows "\<sigma>(x \<leadsto> u, x \<leadsto> v) = \<sigma>(x \<leadsto> v)"
+  using assms by (simp add: expr_defs)
+
+lemma usubst_upd_idem_sub [usubst]:
+  assumes "x \<subseteq>L y" "mwb_lens y"
+  shows "\<sigma>(x \<leadsto> u, y \<leadsto> v) = \<sigma>(y \<leadsto> v)"
+  using assms by (simp add: expr_defs assms fun_eq_iff sublens_put_put)
+
+text \<open> Substitution updates commute when the lenses are independent. \<close>
+    
+lemma subst_upd_comm:
+  assumes "x \<bowtie> y"
+  shows "\<sigma>(x \<leadsto> u, y \<leadsto> v) = \<sigma>(y \<leadsto> v, x \<leadsto> u)"
+  using assms unfolding subst_upd_def
+  by (auto simp add: subst_upd_def assms comp_def lens_indep_comm)
+
+lemma subst_upd_comm2:
+  assumes "z \<bowtie> y"
+  shows "\<sigma>(x \<leadsto> u, y \<leadsto> v, z \<leadsto> s) = \<sigma>(x \<leadsto> u, z \<leadsto> s, y \<leadsto> v)"
+  using assms unfolding subst_upd_def
+  by (auto simp add: subst_upd_def assms comp_def lens_indep_comm)
+
+lemma subst_upd_var_id [usubst]:
+  "vwb_lens x \<Longrightarrow> [x \<leadsto> $x] = [\<leadsto>]"
+  by (simp add: subst_upd_def subst_id_def id_lens_def SEXP_def)
+
+lemma subst_upd_pair [usubst]:
+  "\<sigma>((x, y) \<leadsto> (e, f)) = \<sigma>(y \<leadsto> f, x \<leadsto> e)"
+  by (simp add: subst_upd_def lens_defs SEXP_def fun_eq_iff)
+
+lemma subst_upd_comp [usubst]:
+  "\<rho>(x \<leadsto> v) \<circ>s \<sigma> = (\<rho> \<circ>s \<sigma>)(x \<leadsto> \<sigma> \<dagger> v)"
+  by (simp add: expr_defs fun_eq_iff)
+
+lemma swap_subst_inj: "\<lbrakk> vwb_lens x; vwb_lens y; x \<bowtie> y \<rbrakk> \<Longrightarrow> inj [(x, y) \<leadsto> ($y, $x)]"
+  by (simp add: expr_defs lens_defs inj_on_def)
+     (metis lens_indep.lens_put_irr2 lens_indep_get vwb_lens.source_determination vwb_lens_def wb_lens_weak weak_lens.put_get)
+
+subsection \<open> Proof rules \<close>
+
+text \<open> In proof, a lens can always be substituted for an arbitrary but fixed value. \<close>
+
+lemma taut_substI:
+  assumes "vwb_lens x" "\<And> v. `P\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk>`"
+  shows "`P`"
+  using assms by (expr_simp, metis vwb_lens.put_eq)
+
+lemma eq_substI:
+  assumes "vwb_lens x" "\<And> v. P\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk> = Q\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk>"
+  shows "P = Q"
+  using assms by (expr_simp, metis vwb_lens.put_eq)
+
+lemma bool_eq_substI:
+  assumes "vwb_lens x" "P\<lbrakk>True/x\<rbrakk> = Q\<lbrakk>True/x\<rbrakk>" "P\<lbrakk>False/x\<rbrakk> = Q\<lbrakk>False/x\<rbrakk>"
+  shows "P = Q"
+  by (metis (full_types) assms eq_substI)
+
+lemma less_eq_substI:
+  assumes "vwb_lens x" "\<And> v. P\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk> \<le> Q\<lbrakk>\<guillemotleft>v\<guillemotright>/x\<rbrakk>"
+  shows "P \<le> Q"
+  using assms by (expr_simp, metis le_funE le_funI vwb_lens_def wb_lens.source_stability)
+
+subsection \<open> Ordering substitutions \<close>
+
+text \<open> A simplification procedure to reorder substitutions maplets lexicographically by variable syntax \<close>
+
+simproc_setup subst_order ("subst_upd (subst_upd \<sigma> x u) y v") =
+  \<open> (fn _ => fn ctx => fn ct => 
+        case (Thm.term_of ct) of
+          Const (@{const_name subst_upd}, _) $ (Const (@{const_name subst_upd}, _) $ s $ x $ u) $ y $ v
+            => if (XML.content_of (YXML.parse_body (Syntax.string_of_term ctx x)) > XML.content_of (YXML.parse_body (Syntax.string_of_term ctx y)))
+               then SOME (mk_meta_eq @{thm subst_upd_comm})
+               else NONE  |
+          _ => NONE) 
+  \<close>
+
+subsection \<open> Substitution Unrestriction Laws \<close>
+
+lemma unrest_subst_lens [expr_simps]: "mwb_lens x \<Longrightarrow> ($x \<sharp>s \<sigma>) = (\<forall>s v. \<sigma> (putx s v) = putx (\<sigma> s) v)"
+  by (simp add: unrest_usubst_def, metis lens_override_def mwb_lens_weak weak_lens.create_get)
+
+lemma unrest_subst_empty [unrest]: "x \<sharp>s [\<leadsto>]"
+  by (expr_simp)
+
+lemma unrest_subst_upd [unrest]: "\<lbrakk> vwb_lens x; x \<bowtie> y; $x \<sharp> (e)e; $x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> $x \<sharp>s \<sigma>(y \<leadsto> e)"
+  by (expr_auto add: lens_indep_comm)
+
+lemma unrest_subst_upd_compl [unrest]: "\<lbrakk> vwb_lens x; y \<subseteq>L x; -$x \<sharp> (e)e; -$x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> -$x \<sharp>s \<sigma>(y \<leadsto> e)"
+  by (expr_auto, simp add: lens_override_def scene_override_commute)
+
+lemma unrest_subst_apply [unrest]:
+  "\<lbrakk> $x \<sharp> P; $x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> $x \<sharp> (\<sigma> \<dagger> P)"
+  by (expr_auto)
+
+lemma unrest_sset [unrest]:
+  "x \<bowtie> y \<Longrightarrow> $x \<sharp>s sset[$y, v]"
+  by (expr_auto, meson lens_indep_impl_scene_indep scene_override_commute_indep)
+
+subsection \<open> Conditional Substitution Laws \<close>
+
+lemma subst_cond_upd_1 [usubst]:
+  "\<sigma>(x \<leadsto> u) \<triangleleft> b \<triangleright> \<rho>(x \<leadsto> v) = (\<sigma> \<triangleleft> b \<triangleright> \<rho>)(x \<leadsto> (u \<triangleleft> b \<triangleright> v))"
+  by expr_auto
+
+lemma subst_cond_upd_2 [usubst]:
+  "\<lbrakk> vwb_lens x; $x \<sharp>s \<rho> \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> u) \<triangleleft> b \<triangleright> \<rho> = (\<sigma> \<triangleleft> b \<triangleright> \<rho>)(x \<leadsto> (u \<triangleleft> b \<triangleright> ($x)e))"
+  by (expr_auto, metis lens_override_def lens_override_idem)
+
+lemma subst_cond_upd_3 [usubst]:
+  "\<lbrakk> vwb_lens x; $x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> \<sigma> \<triangleleft> b \<triangleright> \<rho>(x \<leadsto> v) = (\<sigma> \<triangleleft> b \<triangleright> \<rho>)(x \<leadsto> (($x)e \<triangleleft> b \<triangleright> v))"
+  by (expr_auto, metis lens_override_def lens_override_idem)
+
+lemma expr_if_bool_var_left: "vwb_lens x \<Longrightarrow> P\<lbrakk>True/x\<rbrakk> \<triangleleft> $x \<triangleright> Q = P \<triangleleft> $x \<triangleright> Q"
+  by (expr_simp, metis (full_types) lens_override_def lens_override_idem)
+
+lemma expr_if_bool_var_right: "vwb_lens x \<Longrightarrow> P \<triangleleft> $x \<triangleright> Q\<lbrakk>False/x\<rbrakk> = P \<triangleleft> $x \<triangleright> Q"
+  by (expr_simp, metis (full_types) lens_override_def lens_override_idem)
+
+lemma subst_expr_if [usubst]: "\<sigma> \<dagger> (P \<triangleleft> B \<triangleright> Q) = (\<sigma> \<dagger> P) \<triangleleft> (\<sigma> \<dagger> B) \<triangleright> (\<sigma> \<dagger> Q)"
+  by expr_simp
+
+subsection \<open> Substitution Restriction Laws \<close>
+
+lemma subst_restrict_id [usubst]: "idem_scene a \<Longrightarrow> [\<leadsto>] \<restriction>s a = [\<leadsto>]"
+  by expr_simp
+
+lemma subst_restrict_out [usubst]: "\<lbrakk> vwb_lens x; vwb_lens a; x \<bowtie> a \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<restriction>s $a = \<sigma> \<restriction>s $a" 
+  by (expr_simp add: lens_indep.lens_put_irr2)
+
+lemma subst_restrict_in [usubst]: "\<lbrakk> vwb_lens x; vwb_lens y; x \<subseteq>L y \<rbrakk> \<Longrightarrow> \<sigma>(x \<leadsto> e) \<restriction>s $y = (\<sigma> \<restriction>s $y)(x \<leadsto> e)" 
+  by (expr_auto)
+
+lemma subst_restrict_twice [simp]: "\<sigma> \<restriction>s a \<restriction>s a = \<sigma> \<restriction>s a"
+  by expr_simp
+
+subsection \<open> Evaluation \<close>
+
+lemma subst_SEXP [usubst_eval]: "\<sigma> \<dagger> [\<lambda> s. e s]e = [\<lambda> s. e (\<sigma> s)]e"
+  by (simp add: SEXP_def subst_app_def fun_eq_iff)
+
+lemma get_subst_id [usubst_eval]: "getx ([\<leadsto>] s) = getx s"
+  by (simp add: subst_id_def)
+
+lemma get_subst_upd_same [usubst_eval]: "weak_lens x \<Longrightarrow> getx ((\<sigma>(x \<leadsto> e)) s) = e s"
+  by (simp add: subst_upd_def SEXP_def)
+
+lemma get_subst_upd_indep [usubst_eval]: 
+  "x \<bowtie> y \<Longrightarrow> getx ((\<sigma>(y \<leadsto> e)) s) = getx (\<sigma> s)"
+  by (simp add: subst_upd_def)
+
+lemma unrest_ssubst: "(a \<sharp> P) \<longleftrightarrow> (\<forall> s'. sset a s' \<dagger> P = (P)e)"
+  by (auto simp add: expr_defs fun_eq_iff)
+
+lemma unrest_ssubst_expr: "(a \<sharp> (P)e) = (\<forall>s'. sset[a, s'] \<dagger> (P)e = (P)e)"
+  by (simp add: unrest_ssubst)
+
+lemma get_subst_sset_out [usubst_eval]: "\<lbrakk> vwb_lens x; var_alpha x \<bowtie>S a \<rbrakk> \<Longrightarrow> getx (sset a s' s) = getx s"
+  by (simp add: expr_defs var_alpha_def get_scene_override_indep)
+
+lemma get_subst_sset_in [usubst_eval]: "\<lbrakk> vwb_lens x; var_alpha x \<le> a \<rbrakk> \<Longrightarrow> getx (sset a s' s) = getx s'"
+  by (simp add: get_scene_override_le sset_def var_alpha_def)
+
+lemma get_subst_ext [usubst_eval]: "getx (subst_ext a s) = getns_alpha a x s"
+  by (expr_simp)
+
+lemma unrest_sset_lens [unrest]: "\<lbrakk> mwb_lens x; mwb_lens y; x \<bowtie> y \<rbrakk> \<Longrightarrow> $x \<sharp>s sset[$y, s]"
+  by (simp add: sset_def unrest_subst_lens lens_indep_comm lens_override_def)
+
+lemma get_subst_restrict_out [usubst_eval]: "\<lbrakk> vwb_lens a; x \<bowtie> a \<rbrakk> \<Longrightarrow> getx ((\<sigma> \<restriction>s $a) s) = getx s"
+  by (expr_simp)
+
+lemma get_subst_restrict_in [usubst_eval]: "\<lbrakk> vwb_lens a; x \<subseteq>L a \<rbrakk> \<Longrightarrow> getx ((\<sigma> \<restriction>s $a) s) = getx (\<sigma> s)"
+  by (expr_simp, force)
+
+text \<open> If a variable is unrestricted in a substitution then it's application has no effect. \<close>
+
+lemma subst_apply_unrest:
+  "\<lbrakk> vwb_lens x; $x \<sharp>s \<sigma> \<rbrakk> \<Longrightarrow> \<langle>\<sigma>\<rangle>s x = var x"
+proof -
+  assume 1: "vwb_lens x" and "$x \<sharp>s \<sigma>"
+  hence "\<forall>s s'. \<sigma> (s \<oplus>L s' on x) = \<sigma> s \<oplus>L s' on x"
+    by (simp add: unrest_usubst_def)
+  thus "\<langle>\<sigma>\<rangle>s x = var x"
+    by (metis 1 lens_override_def lens_override_idem mwb_lens_weak subst_lookup_def vwb_lens_mwb weak_lens.put_get)
+qed
+
+text \<open> A tactic for proving unrestrictions by evaluating a special kind of substitution. \<close>
+
+method unrest uses add = (simp add: add unrest unrest_ssubst_expr var_alpha_combine usubst usubst_eval)
+
+text \<open> A tactic for evaluating substitutions. \<close>
+
+method subst_eval uses add = (simp add: add usubst_eval usubst unrest)
+
+text \<open> We can exercise finer grained control over substitutions with the following method. \<close>
+
+declare vwb_lens_mwb [lens]
+declare mwb_lens_weak [lens]
+
+method subst_eval' = (simp only: lens usubst_eval usubst unrest SEXP_apply)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Unrestriction.thy
--- /dev/null
+++ thys/Shallow_Expressions/Unrestriction.thy
@@ -0,0 +1,153 @@
+section \<open> Unrestriction \<close>
+
+theory Unrestriction
+  imports Expressions
+begin
+
+text \<open> Unrestriction means that an expression does not depend on the value of the state space
+  described by the given scene (i.e. set of variables) for its valuation. It is a semantic
+  characterisation of fresh variables. \<close>
+
+consts unrest :: "'s scene \<Rightarrow> 'p \<Rightarrow> bool"
+
+definition unrest_expr :: "'s scene \<Rightarrow> ('b, 's) expr \<Rightarrow> bool" where
+[expr_defs]: "unrest_expr a e = (\<forall> s s'. e (s \<oplus>S s' on a) = e s)"
+
+adhoc_overloading unrest \<rightleftharpoons> unrest_expr
+
+syntax
+  "_unrest" :: "salpha \<Rightarrow> logic \<Rightarrow> logic \<Rightarrow> logic" (infix "\<sharp>" 20)
+
+translations
+  "_unrest x p" == "CONST unrest x p"                                           
+
+named_theorems unrest
+
+lemma unrest_empty [unrest]: "\<emptyset> \<sharp> P"
+  by (simp add: expr_defs lens_defs)
+
+lemma unrest_var_union [unrest]:
+  "\<lbrakk> A \<sharp> P; B \<sharp> P \<rbrakk> \<Longrightarrow> A \<union> B \<sharp> P"
+  by (simp add: expr_defs lens_defs)
+     (metis scene_override_union scene_override_unit scene_union_incompat)
+
+lemma unrest_neg_union:
+  assumes "A ##S B" "- A \<sharp> P" "- B \<sharp> P"
+  shows "(- (A \<union> B)) \<sharp> P"
+  using assms by (simp add: unrest_expr_def scene_override_commute scene_override_union)
+
+text \<open> The following two laws greatly simplify proof when reasoning about unrestricted lens,
+  and so we add them to the expression simplification set. \<close>
+
+lemma unrest_lens [expr_simps]:
+  "mwb_lens x \<Longrightarrow> ($x \<sharp> e) = (\<forall> s v. e (putx s v) = e s)"
+  by (simp add: unrest_expr_def var_alpha_def comp_mwb_lens lens_override_def)
+     (metis mwb_lens.put_put)
+
+lemma unrest_compl_lens [expr_simps]:
+  "mwb_lens x \<Longrightarrow> (- $x \<sharp> e) = (\<forall>s s'. e (putx s' (getx s)) = e s)"
+  by (simp add: unrest_expr_def var_alpha_def comp_mwb_lens lens_override_def scene_override_commute)
+
+lemma unrest_subscene: "\<lbrakk> idem_scene a; a \<sharp> e; b \<subseteq>S a \<rbrakk> \<Longrightarrow> b \<sharp> e"
+  by (metis subscene_eliminate unrest_expr_def)
+
+lemma unrest_lens_comp [unrest]: "\<lbrakk> mwb_lens x; mwb_lens y; $x \<sharp> e \<rbrakk> \<Longrightarrow> $x:y \<sharp> e"
+  by (simp add: unrest_lens, simp add: lens_comp_def ns_alpha_def)
+
+lemma unrest_expr [unrest]: "x \<sharp> e \<Longrightarrow> x \<sharp> (e)e"
+  by (simp add: expr_defs)
+
+lemma unrest_lit [unrest]: "x \<sharp> (\<guillemotleft>v\<guillemotright>)e"
+  by (simp add: expr_defs)
+
+lemma unrest_var [unrest]: 
+  "\<lbrakk> vwb_lens x; idem_scene a; var_alpha x \<bowtie>S a \<rbrakk> \<Longrightarrow> a \<sharp> ($x)e"
+  by (auto simp add: unrest_expr_def scene_indep_override var_alpha_def)
+     (metis lens_override_def lens_override_idem mwb_lens_weak vwb_lens_mwb weak_lens_def)
+
+lemma unrest_var_single [unrest]:
+  "\<lbrakk> mwb_lens x; x \<bowtie> y \<rbrakk> \<Longrightarrow> $x \<sharp> ($y)e"
+  by (simp add: expr_defs lens_indep.lens_put_irr2 lens_indep_sym lens_override_def var_alpha_def)
+
+lemma unrest_sublens:
+  assumes "mwb_lens x" "$x \<sharp> P" "y \<subseteq>L x"
+  shows "$y \<sharp> P"
+  by (metis assms sublens_pres_mwb sublens_put_put unrest_lens)
+
+text \<open> If two lenses are equivalent, and thus they characterise the same state-space regions,
+  then clearly unrestrictions over them are equivalent. \<close>
+    
+lemma unrest_equiv:
+  assumes "mwb_lens y" "x \<approx>L y" "$x \<sharp> P"
+  shows "$y \<sharp> P"
+  using assms lens_equiv_def sublens_pres_mwb unrest_sublens by blast
+
+text \<open> If we can show that an expression is unrestricted on a bijective lens, then is unrestricted
+  on the entire state-space. \<close>
+
+lemma bij_lens_unrest_all:
+  assumes "bij_lens x" "$x \<sharp> P"
+  shows "\<Sigma> \<sharp> P"
+  by (metis assms bij_lens_vwb lens_scene_top_iff_bij_lens univ_alpha_def var_alpha_def vwb_lens_iff_mwb_UNIV_src)
+
+lemma bij_lens_unrest_all_eq:
+  assumes "bij_lens x"
+  shows "(\<Sigma> \<sharp> P) \<longleftrightarrow> ($x \<sharp> P)"
+  by (metis assms bij_lens_vwb lens_scene_top_iff_bij_lens univ_alpha_def var_alpha_def vwb_lens_mwb)
+
+text \<open> If an expression is unrestricted by all variables, then it is unrestricted by any variable \<close>
+
+lemma unrest_all_var:
+  assumes "\<Sigma> \<sharp> e"
+  shows "$x \<sharp> e"
+  by (metis assms scene_top_greatest top_idem_scene univ_alpha_def unrest_subscene)
+
+lemma unrest_pair [unrest]:
+  assumes "mwb_lens x" "mwb_lens y" "$x \<sharp> P" "$y \<sharp> P"
+  shows "$(x, y) \<sharp> P"
+  using assms
+  by expr_simp (simp add: lens_override_def lens_scene.rep_eq scene_override.rep_eq)
+
+lemma unrest_pair_split:
+  assumes "x \<bowtie> y" "vwb_lens x" "vwb_lens y"
+  shows "($(x, y) \<sharp> P) = (($x \<sharp> P) \<and> ($y \<sharp> P))"
+  using assms
+  by (metis lens_equiv_scene lens_indep_sym lens_plus_comm lens_plus_right_sublens plus_vwb_lens sublens_refl unrest_pair unrest_sublens var_alpha_def var_pair_def vwb_lens_def)
+
+lemma unrest_get [unrest]: "\<lbrakk> mwb_lens x; x \<bowtie> y \<rbrakk> \<Longrightarrow> $x \<sharp> gety"
+  by (expr_simp, simp add: lens_indep.lens_put_irr2)
+
+lemma unrest_conj [unrest]:
+  "\<lbrakk> x \<sharp> P; x \<sharp> Q \<rbrakk> \<Longrightarrow> x \<sharp> (P \<and> Q)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_not [unrest]:
+  "\<lbrakk> x \<sharp> P \<rbrakk> \<Longrightarrow> x \<sharp> (\<not> P)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_disj [unrest]:
+  "\<lbrakk> x \<sharp> P; x \<sharp> Q \<rbrakk> \<Longrightarrow> x \<sharp> (P \<or> Q)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_implies [unrest]:
+  "\<lbrakk> x \<sharp> P; x \<sharp> Q \<rbrakk> \<Longrightarrow> x \<sharp> (P \<longrightarrow> Q)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_expr_if [unrest]:
+  assumes "a \<sharp> P" "a \<sharp> Q" "a \<sharp> (e)e"
+  shows "a \<sharp> (P \<triangleleft> e \<triangleright> Q)"
+  using assms by expr_simp
+
+lemma unrest_uop:
+  "\<lbrakk> x \<sharp> e \<rbrakk> \<Longrightarrow> x \<sharp> (\<guillemotleft>f\<guillemotright> e)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_bop:
+  "\<lbrakk> x \<sharp> e1; x \<sharp> e2 \<rbrakk> \<Longrightarrow> x \<sharp> (\<guillemotleft>f\<guillemotright> e1 e2)e"
+  by (auto simp add: expr_defs)
+
+lemma unrest_trop:
+  "\<lbrakk> x \<sharp> e1; x \<sharp> e2; x \<sharp> e3 \<rbrakk> \<Longrightarrow> x \<sharp> (\<guillemotleft>f\<guillemotright> e1 e2 e3)e"
+  by (auto simp add: expr_defs)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/Variables.thy
--- /dev/null
+++ thys/Shallow_Expressions/Variables.thy
@@ -0,0 +1,329 @@
+section \<open> Variables as Lenses \<close>
+
+theory Variables
+  imports "Optics.Optics"
+begin
+
+text \<open> Here, we implement foundational constructors and syntax translations for state variables,
+  using lens manipulations, for use in shallow expressions. \<close>
+
+subsection \<open> Constructors \<close>
+
+text \<open> The following bundle allows us to override the colon operator, which we use for variable
+  namespaces. \<close>
+
+bundle Expression_Syntax
+begin
+
+no_notation   
+  Set.member  (\<open>'(:')\<close>) and
+  Set.member  (\<open>(\<open>notation=\<open>infix :\<close>\<close>_/ : _)\<close> [51, 51] 50)
+
+end
+
+unbundle Expression_Syntax
+
+declare fst_vwb_lens [simp] and snd_vwb_lens [simp]
+
+text \<open> Lenses~\cite{Foster2020-IsabelleUTP} can also be used to effectively define sets of variables. Here we define the
+  the universal alphabet ($\Sigma$) to be the bijective lens @{term "1L"}. This characterises
+  the whole of the source type, and thus is effectively the set of all alphabet variables. \<close>
+
+definition univ_var :: "('\<alpha> \<Longrightarrow> '\<alpha>)" ("v") where
+[lens_defs]: "univ_var = 1L"
+
+lemma univ_var_vwb [simp]: "vwb_lens univ_var"
+  by (simp add: univ_var_def)
+
+definition univ_alpha :: "'s scene" where
+[lens_defs]: "univ_alpha = \<top>S"
+
+definition emp_alpha :: "'s scene" where
+[lens_defs]: "emp_alpha = \<bottom>S"
+
+definition var_alpha :: "('a \<Longrightarrow> 's) \<Rightarrow> 's scene" where
+[lens_defs]: "var_alpha x = lens_scene x"
+
+definition ns_alpha :: "('b \<Longrightarrow> 'c) \<Rightarrow> ('a \<Longrightarrow> 'b) \<Rightarrow> 'a \<Longrightarrow> 'c" where
+[lens_defs]: "ns_alpha a x = x ;L a"
+
+definition var_fst :: "('a \<times> 'b \<Longrightarrow> 's) \<Rightarrow> ('a \<Longrightarrow> 's)" where
+[lens_defs]: "var_fst x = fstL ;L x" 
+
+definition var_snd :: "('a \<times> 'b \<Longrightarrow> 's) \<Rightarrow> ('b \<Longrightarrow> 's)" where
+[lens_defs]: "var_snd x = sndL ;L x" 
+
+definition var_pair :: "('a \<Longrightarrow> 's) \<Rightarrow> ('b \<Longrightarrow> 's) \<Rightarrow> ('a \<times> 'b \<Longrightarrow> 's)" where
+[lens_defs]: "var_pair x y = x +L y"
+
+abbreviation var_member :: "('a \<Longrightarrow> 's) \<Rightarrow> 's scene \<Rightarrow> bool" (infix "\<in>v" 50) where
+"x \<in>v A \<equiv> var_alpha x \<le> A"
+
+lemma ns_alpha_weak [simp]: "\<lbrakk> weak_lens a; weak_lens x \<rbrakk> \<Longrightarrow> weak_lens (ns_alpha a x)"
+  by (simp add: ns_alpha_def comp_weak_lens)
+
+lemma ns_alpha_mwb [simp]: "\<lbrakk> mwb_lens a; mwb_lens x \<rbrakk> \<Longrightarrow> mwb_lens (ns_alpha a x)"
+  by (simp add: ns_alpha_def comp_mwb_lens)
+
+lemma ns_alpha_vwb [simp]: "\<lbrakk> vwb_lens a; vwb_lens x \<rbrakk> \<Longrightarrow> vwb_lens (ns_alpha a x)"
+  by (simp add: ns_alpha_def comp_vwb_lens)
+
+lemma ns_alpha_indep_1 [simp]: "a \<bowtie> b \<Longrightarrow> ns_alpha a x \<bowtie> ns_alpha b y"
+  by (simp add: lens_indep_left_ext lens_indep_right_ext ns_alpha_def)
+
+lemma ns_alpha_indep_2 [simp]: "a \<bowtie> y \<Longrightarrow> ns_alpha a x \<bowtie> y"
+  by (simp add: lens_indep_left_ext ns_alpha_def)
+
+lemma ns_alpha_indep_3 [simp]: "x \<bowtie> b \<Longrightarrow> x \<bowtie> ns_alpha b y"
+  by (simp add: lens_indep_sym)
+
+lemma ns_alpha_indep_4 [simp]: "\<lbrakk> mwb_lens a; x \<bowtie> y \<rbrakk> \<Longrightarrow> ns_alpha a x \<bowtie> ns_alpha a y"
+  by (simp add: ns_alpha_def)
+
+lemma var_fst_mwb [simp]: "mwb_lens x \<Longrightarrow> mwb_lens (var_fst x)"
+  by (simp add: var_fst_def comp_mwb_lens)
+
+lemma var_snd_mwb [simp]: "mwb_lens x \<Longrightarrow> mwb_lens (var_snd x)"
+  by (simp add: var_snd_def comp_mwb_lens)
+
+lemma var_fst_vwb [simp]: "vwb_lens x \<Longrightarrow> vwb_lens (var_fst x)"
+  by (simp add: var_fst_def comp_vwb_lens)
+
+lemma var_snd_vwb [simp]: "vwb_lens x \<Longrightarrow> vwb_lens (var_snd x)"
+  by (simp add: var_snd_def comp_vwb_lens)
+
+lemma var_fst_indep_1 [simp]: "x \<bowtie> y \<Longrightarrow> var_fst x \<bowtie> y"
+  by (simp add: var_fst_def lens_indep_left_ext)
+
+lemma var_fst_indep_2 [simp]: "x \<bowtie> y \<Longrightarrow> x \<bowtie> var_fst y"
+  by (simp add: var_fst_def lens_indep_right_ext)
+
+lemma var_snd_indep_1 [simp]: "x \<bowtie> y \<Longrightarrow> var_snd x \<bowtie> y"
+  by (simp add: var_snd_def lens_indep_left_ext)
+
+lemma var_snd_indep_2 [simp]: "x \<bowtie> y \<Longrightarrow> x \<bowtie> var_snd y"
+  by (simp add: var_snd_def lens_indep_right_ext)
+
+lemma mwb_var_pair [simp]: "\<lbrakk> mwb_lens x; mwb_lens y; x \<bowtie> y \<rbrakk> \<Longrightarrow> mwb_lens (var_pair x y)"
+  by (simp add: var_pair_def plus_mwb_lens)
+
+lemma vwb_var_pair [simp]: "\<lbrakk> vwb_lens x; vwb_lens y; x \<bowtie> y \<rbrakk> \<Longrightarrow> vwb_lens (var_pair x y)"
+  by (simp add: var_pair_def)
+
+lemma var_pair_pres_indep [simp]:
+  "\<lbrakk> x \<bowtie> y; x \<bowtie> z \<rbrakk> \<Longrightarrow> x \<bowtie> var_pair y z"
+  "\<lbrakk> x \<bowtie> y; x \<bowtie> z \<rbrakk> \<Longrightarrow> var_pair y z \<bowtie> x"
+  by (simp_all add: var_pair_def lens_indep_sym)
+
+definition res_alpha :: "('a \<Longrightarrow> 'b) \<Rightarrow> ('c \<Longrightarrow> 'b) \<Rightarrow> 'a \<Longrightarrow> 'c" where
+[lens_defs]: "res_alpha x a = x /L a"
+
+lemma idem_scene_var [simp]:
+  "vwb_lens x \<Longrightarrow> idem_scene (var_alpha x)"
+  by (simp add: lens_defs)
+
+lemma var_alpha_combine: "\<lbrakk> vwb_lens x; vwb_lens y; x \<bowtie> y \<rbrakk> \<Longrightarrow> var_alpha x \<squnion>S var_alpha y = var_alpha (x +L y)"
+  by (simp add: lens_plus_scene var_alpha_def)
+
+lemma var_alpha_indep [simp]: 
+  assumes "vwb_lens x" "vwb_lens y"
+  shows "var_alpha x \<bowtie>S var_alpha y \<longleftrightarrow> x \<bowtie> y"
+  by (simp add: assms(1) assms(2) lens_indep_scene var_alpha_def)
+
+lemma pre_var_indep_prod [simp]: "x \<bowtie> a \<Longrightarrow> ns_alpha fstL x \<bowtie> a \<times>L b"
+  using lens_indep.lens_put_irr2
+  by (unfold_locales, force simp add: lens_defs prod.case_eq_if lens_indep_comm)+
+
+lemma post_var_indep_prod [simp]: "x \<bowtie> b \<Longrightarrow> ns_alpha sndL x \<bowtie> a \<times>L b"
+  using lens_indep.lens_put_irr2
+  by (unfold_locales, force simp add: lens_defs prod.case_eq_if lens_indep_comm)+
+
+lemma lens_indep_impl_scene_indep_var [simp]:
+  "(X \<bowtie> Y) \<Longrightarrow> var_alpha X \<bowtie>S var_alpha Y"
+  by (simp add: var_alpha_def)
+
+declare lens_scene_override [simp]
+declare uminus_scene_twice [simp]
+
+lemma var_alpha_override [simp]: 
+  "mwb_lens X \<Longrightarrow> s1 \<oplus>S s2 on var_alpha X = s1 \<oplus>L s2 on X"
+  by (simp add: var_alpha_def)
+
+(* Some extra Scene laws; should be moved to Optics at some point *)
+
+lemma var_alpha_indep_compl [simp]: 
+  assumes "vwb_lens x" "vwb_lens y"
+  shows "var_alpha x \<bowtie>S - var_alpha y \<longleftrightarrow> x \<subseteq>L y"
+  by (simp add: assms scene_le_iff_indep_inv sublens_iff_subscene var_alpha_def)
+
+lemma var_alpha_subset [simp]:
+  assumes "vwb_lens x" "vwb_lens y"
+  shows "var_alpha x \<le> var_alpha y \<longleftrightarrow> x \<subseteq>L y"
+  by (simp add: assms(1) assms(2) sublens_iff_subscene var_alpha_def)
+
+subsection \<open> Syntax Translations \<close>
+
+text \<open> In order to support nice syntax for variables, we here set up some translations. The first
+  step is to introduce a collection of non-terminals. \<close>
+  
+nonterminal svid and svids and salpha and sframe_enum and sframe
+
+text \<open> These non-terminals correspond to the following syntactic entities. Non-terminal 
+  @{typ "svid"} is an atomic variable identifier, and @{typ "svids"} is a list of identifier. 
+  @{typ "salpha"} is an alphabet or set of variables. @{typ "sframe"} is a frame. Such sets can 
+  be constructed only through lens composition due to typing restrictions. Next we introduce some 
+  syntax constructors. \<close>
+   
+syntax \<comment> \<open> Identifiers \<close>
+  "_svid"         :: "id_position \<Rightarrow> svid" ("_" [999] 999)
+  "_svlongid"     :: "longid_position \<Rightarrow> svid" ("_" [999] 999)
+  ""              :: "svid \<Rightarrow> svids" ("_")
+  "_svid_list"    :: "svid \<Rightarrow> svids \<Rightarrow> svids" ("_,/ _")
+  "_svid_alpha"   :: "svid" ("v")
+  "_svid_index"   :: "id_position \<Rightarrow> logic \<Rightarrow> svid" ("_'(_')" [999] 999)
+  "_svid_tuple"   :: "svids \<Rightarrow> svid" ("'(_')")
+  "_svid_dot"     :: "svid \<Rightarrow> svid \<Rightarrow> svid" ("_:_" [999,998] 998)
+  "_svid_res"     :: "svid \<Rightarrow> svid \<Rightarrow> svid" ("_\<restriction>_" [999,998] 998)
+  "_svid_pre"     :: "svid \<Rightarrow> svid" ("_<" [997] 997)
+  "_svid_post"    :: "svid \<Rightarrow> svid" ("_>" [997] 997)
+  "_svid_fst"     :: "svid \<Rightarrow> svid" ("_.1" [997] 997)
+  "_svid_snd"     :: "svid \<Rightarrow> svid" ("_.2" [997] 997)
+  "_mk_svid_list" :: "svids \<Rightarrow> logic" \<comment> \<open> Helper function for summing a list of identifiers \<close>
+  "_of_svid_list"   :: "logic \<Rightarrow> svids" \<comment> \<open> Reverse of the above \<close>
+  "_svid_view"    :: "logic \<Rightarrow> svid" ("\<V>[_]") \<comment> \<open> View of a symmetric lens \<close>
+  "_svid_coview"  :: "logic \<Rightarrow> svid" ("\<C>[_]") \<comment> \<open> Coview of a symmetric lens \<close>
+  "_svid_prod"    :: "svid \<Rightarrow> svid \<Rightarrow> svid" (infixr "\<times>" 85)
+  "_svid_pow2"    :: "svid \<Rightarrow> svid" ("_2" [999] 999)
+
+text \<open> A variable can be decorated with an ampersand, to indicate it is a predicate variable, with 
+  a dollar to indicate its an unprimed relational variable, or a dollar and ``acute'' symbol to 
+  indicate its a primed relational variable. Isabelle's parser is extensible so additional
+  decorations can be and are added later. \<close>
+
+syntax \<comment> \<open> Variable sets \<close>
+  "_salphaid"    :: "id_position \<Rightarrow> salpha" ("_" [990] 990)
+  "_salphavar"   :: "svid \<Rightarrow> salpha" ("$_" [990] 990)
+  "_salphaparen" :: "salpha \<Rightarrow> salpha" ("'(_')")
+  "_salphaunion" :: "salpha \<Rightarrow> salpha \<Rightarrow> salpha" (infixr "\<union>" 75)
+  "_salphainter" :: "salpha \<Rightarrow> salpha \<Rightarrow> salpha" (infixr "\<inter>" 75)
+  "_salphaminus" :: "salpha \<Rightarrow> salpha \<Rightarrow> salpha" (infixl "-" 65)
+  "_salphacompl" :: "salpha \<Rightarrow> salpha" ("- _" [81] 80)
+  "_salpha_all"  :: "salpha" ("\<Sigma>")
+  "_salpha_none" :: "salpha" ("\<emptyset>")
+(*  "_salpha_pre"  :: "salpha \<Rightarrow> salpha" ("_<" [989] 989)
+  "_salpha_post" :: "salpha \<Rightarrow> salpha" ("_>" [989] 989) *)
+  "_salphaset"   :: "svids \<Rightarrow> salpha" ("{_}")
+  "_sframeid"    :: "id \<Rightarrow> sframe" ("_")
+  "_sframeset"   :: "svids \<Rightarrow> sframe_enum" ("\<lbrace>_\<rbrace>")
+  "_sframeunion" :: "sframe \<Rightarrow> sframe \<Rightarrow> sframe" (infixr "\<union>" 75)
+  "_sframeinter" :: "sframe \<Rightarrow> sframe \<Rightarrow> sframe" (infixr "\<inter>" 75)
+  "_sframeminus" :: "sframe \<Rightarrow> sframe \<Rightarrow> sframe" (infixl "-" 65)
+  "_sframecompl" :: "sframe \<Rightarrow> sframe" ("- _" [81] 80)
+  "_sframe_all"  :: "sframe" ("\<Sigma>")
+  "_sframe_none" :: "sframe" ("\<emptyset>")
+  "_sframe_pre"  :: "sframe \<Rightarrow> sframe" ("_<" [989] 989)
+  "_sframe_post" :: "sframe \<Rightarrow> sframe" ("_>" [989] 989)
+  "_sframe_enum" :: "sframe_enum \<Rightarrow> sframe" ("_")
+  "_sframe_alpha" :: "sframe_enum \<Rightarrow> salpha" ("_")
+  "_salphamk"    :: "logic \<Rightarrow> salpha"
+  "_mk_alpha_list" :: "svids \<Rightarrow> logic"
+  "_mk_frame_list" :: "svids \<Rightarrow> logic"
+
+text \<open> The terminals of an alphabet are either HOL identifiers or UTP variable identifiers. 
+  We support two ways of constructing alphabets; by composition of smaller alphabets using
+  a semi-colon or by a set-style construction $\{a,b,c\}$ with a list of UTP variables. \<close>
+
+syntax \<comment> \<open> Quotations \<close>
+  "_svid_set"    :: "svids \<Rightarrow> logic" ("{_}v")
+  "_svid_empty"  :: "logic" ("{}v")
+  "_svar"        :: "svid \<Rightarrow> logic" ("'(_')v")
+  
+text \<open> For various reasons, the syntax constructors above all yield specific grammar categories and
+  will not parser at the HOL top level (basically this is to do with us wanting to reuse the syntax
+  for expressions). As a result we provide some quotation constructors above. 
+
+  Next we need to construct the syntax translations rules. Finally, we set up the translations rules. \<close>
+
+translations
+  \<comment> \<open> Identifiers \<close>
+  "_svid x" \<rightharpoonup> "x"
+  "_svlongid x" \<rightharpoonup> "x"
+  "_svid_alpha" \<rightleftharpoons> "CONST univ_var"
+  "_svid_tuple xs" \<rightharpoonup> "_mk_svid_list xs"
+  "_svid_dot x y" \<rightleftharpoons> "CONST ns_alpha x y"
+  "_svid_index x i" \<rightharpoonup> "x i"
+  "_svid_res x y" \<rightleftharpoons> "x /L y" 
+  "_svid_pre x" \<rightleftharpoons> "_svid_dot fstL x"
+  "_svid_post x" \<rightleftharpoons> "_svid_dot sndL x"
+  "_svid_fst x" \<rightleftharpoons> "CONST var_fst x"
+  "_svid_snd x" \<rightleftharpoons> "CONST var_snd x"
+  "_svid_prod x y" \<rightleftharpoons> "x \<times>L y"
+  "_svid_pow2 x" \<rightharpoonup> "x \<times>L x"
+  "_mk_svid_list (_svid_list x xs)" \<rightharpoonup> "CONST var_pair x (_mk_svid_list xs)"
+  "_mk_svid_list x" \<rightharpoonup> "x"
+  "_mk_alpha_list (_svid_list x xs)" \<rightharpoonup> "CONST var_alpha x \<squnion>S _mk_alpha_list xs"
+  "_mk_alpha_list x" \<rightharpoonup> "CONST var_alpha x"
+
+  "_svid_view a" => "\<V>a"
+  "_svid_coview a" => "\<C>a"
+
+  "_svid_list (_svid_tuple (_of_svid_list (CONST var_pair x y))) (_of_svid_list z)" \<leftharpoondown> "_of_svid_list (CONST var_pair (CONST var_pair x y) z)"
+  "_svid_list x (_of_svid_list y)"    \<leftharpoondown> "_of_svid_list (CONST var_pair x y)"
+  "x"                                 \<leftharpoondown> "_of_svid_list x"
+
+  "_svid_tuple (_svid_list x y)" \<leftharpoondown> "CONST var_pair x y"
+  "_svid_list x ys" \<leftharpoondown> "_svid_list x (_svid_tuple ys)"
+
+  \<comment> \<open> Alphabets \<close>
+  "_salphaparen a" \<rightharpoonup> "a"
+  "_salphaid x" \<rightharpoonup> "x"
+  "_salphaunion x y" \<rightharpoonup> "x \<squnion>S y"
+  "_salphainter x y" \<rightharpoonup> "x \<sqinter>S y"
+  "_salphaminus x y" \<rightharpoonup> "x - y"
+  "_salphacompl x"  \<rightharpoonup> "- x"
+(*  "_salpha_pre A" \<rightharpoonup> "A ;S fstL"
+  "_salpha_post A" \<rightharpoonup> "A ;S sndL" *)
+(*  "_salphaprod a b" \<rightleftharpoons> "a \<times>L b" *)
+  "_salphavar x" \<rightleftharpoons> "CONST var_alpha x"
+  "_salphaset A" \<rightharpoonup> "_mk_alpha_list A"  
+  "_sframeid A" \<rightharpoonup> "A"
+  "(_svid_list x (_salphamk y))" \<leftharpoondown> "_salphamk (x +L y)" 
+  "x" \<leftharpoondown> "_salphamk x"
+  "_salpha_all" \<rightleftharpoons> "CONST univ_alpha"
+  "_salpha_none" \<rightleftharpoons> "CONST emp_alpha"
+
+  \<comment> \<open> Quotations \<close>
+  "_svid_set A" \<rightharpoonup> "_mk_alpha_list A"
+  "_svid_empty" \<rightharpoonup> "0L"
+  "_svar x" \<rightharpoonup> "x"
+
+text \<open> The translation rules mainly convert syntax into lens constructions, using a mixture
+  of lens operators and the bespoke variable definitions. Notably, a colon variable identifier
+  qualification becomes a lens composition, and variable sets are constructed using len sum. 
+  The translation rules are carefully crafted to ensure both parsing and pretty printing. 
+
+  Finally we create the following useful utility translation function that allows us to construct a 
+  UTP variable (lens) type given a return and alphabet type. \<close>
+
+syntax
+  "_uvar_ty"      :: "type \<Rightarrow> type \<Rightarrow> type"
+
+parse_translation \<open>
+let
+  fun uvar_ty_tr [ty] = Syntax.const @{type_syntax lens} $ ty $ Syntax.const @{type_syntax dummy}
+    | uvar_ty_tr ts = raise TERM ("uvar_ty_tr", ts);
+in [(@{syntax_const "_uvar_ty"}, K uvar_ty_tr)] end
+\<close>
+
+subsection \<open> Simplifications \<close>
+
+lemma get_pre [simp]: "get(x<)v (s1, s2) = getx s1"
+  by (simp add: lens_defs)
+
+lemma get_post [simp]: "get(x>)v (s1, s2) = getx s2"
+  by (simp add: lens_defs)
+
+lemma get_prod_decomp: "getx s = (getvar_fst x s, getvar_snd x s)"
+  by (simp add: lens_defs)
+
+end
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/document/document.sty
--- /dev/null
+++ thys/Shallow_Expressions/document/document.sty
@@ -0,0 +1,19 @@
+% Support the outer command "paragraph {* *}".
+
+%%%%%%%%%%%
+% Colours %
+%%%%%%%%%%%
+
+\newcommand{\red}[1]{{\color{Red}#1}}
+\newcommand{\blue}[1]{{\color{Blue}#1}}
+\newcommand{\green}[1]{{\color{Green}#1}}
+\newcommand{\grey}[1]{{\color{Gray}#1}}
+\newcommand{\purple}[1]{{\color{Purple}#1}}
+
+%%%%%%%%%%%%%%%%%%%%%%%
+% Document Management %
+%%%%%%%%%%%%%%%%%%%%%%%
+
+\newcommand{\todo}[1]{\textbf{TODO}:~\red{#1}}
+\newcommand{\note}[1]{\textbf{NOTE}:~\green{#1}}
+\newcommand{\fixme}[1]{\textbf{FIXME}:~\blue{#1}}
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/document/root.bib
--- /dev/null
+++ thys/Shallow_Expressions/document/root.bib
@@ -0,0 +1,271 @@
+@InProceedings{Armstrong2012,
+  Title                    = {Automated Reasoning in Higher-Order Regular Algebra},
+  Author                   = {Armstrong, A. and Struth, G.},
+  Booktitle                = {RAMiCS 2012},
+  Year                     = {2012},
+  Month                    = {September},
+  Publisher                = {Springer},
+  Series                   = {LNCS},
+  Volume                   = {7560}
+}
+
+@InProceedings{Armstrong2013,
+  Title                    = {Program Analysis and Verification Based on {Kleene Algebra} in {Isabelle/HOL}},
+  Author                   = {Armstrong, A. and Struth, G. and Weber, T.},
+  Booktitle                = {ITP},
+  Year                     = 2013,
+  Publisher                = {Springer},
+  Series                   = {LNCS},
+  Volume                   = 7998
+}
+
+@Article{Armstrong2015,
+  author="Armstrong, A. and Gomes, V. and Struth, G.",
+  title="Building program construction and verification tools from algebraic principles",
+  journal="Formal Aspects of Computing",
+  year="2015",
+  volume="28",
+  number="2",
+  pages="265--293"
+}
+
+@InProceedings{Blanchette2011,
+  Title                    = {Automatic Proof and Disproof in {Isabelle/HOL}},
+  Author                   = {Blanchette, J. C. and Bulwahn, L. and Nipkow, T.},
+  Booktitle                = {FroCoS},
+  Year                     = {2011},
+  Pages                    = {12--27},
+  Publisher                = {Springer},
+  Series                   = {LNCS},
+  Volume                   = {6989}
+}
+
+@INCOLLECTION{Cavalcanti&06,
+  KEY           = "Cavalcanti\&06",
+  AUTHOR        = {Cavalcanti, A. and Woodcock, J.},
+  TITLE         = {A Tutorial Introduction to {CSP} in Unifying Theories of
+                   Programming},
+  BOOKTITLE     = {Refinement Techniques in Software Engineering},
+  SERIES        = {LNCS},
+  PUBLISHER     = {Springer},
+  ISBN          = {978-3-540-46253-8},
+  PAGES         = {220--268},
+  VOLUME        = {3167},
+  YEAR          = {2006},
+  COMMENT       = "BIB PGL"}
+
+@InProceedings{Dongol2015,
+  author = {Dongol, B. and Gomes, V. B. F. and Struth, G.},
+  title  = {A Program Construction and Verification Tool for Separation Logic},
+  booktitle = {Mathematics of Program Construction (MPC)},
+  series = {LNCS},
+  publisher = {Springer},
+  volume = {9129},
+  pages = {137--158},
+  year = {2015}
+}
+
+@InProceedings{Dongol19,
+  author =       {Dongol, B. and Hayes, I. and Meinicke, L. and Struth, G.},
+  title =        {Cylindric {Kleene} Lattices for Program Construction},
+  booktitle = {MPC},
+  month =     {October},
+  year =      2019,
+  series =    {LNCS},
+  volume =    {11825},
+  pages =     {192--225},
+  publisher = {Springer}}
+
+@InProceedings{Feliachi2010,
+  author =       {Feliachi, A. and Gaudel, M.-C. and Wolff, B.},
+  title =        {Unifying Theories in {Isabelle/HOL}},
+  booktitle    = {UTP 2010},
+  pages    =     {188--206},
+  year =         2010,
+  volume    =    6445,
+  series =       {LNCS},
+  publisher = {Springer}}
+
+@InProceedings{Feliachi2012,
+  author =       {Feliachi, A. and Gaudel, M.-C. and Wolff, B.},
+  title =        {{Isabelle/Circus}: a Process Specification and
+                  Verification Environment},
+  booktitle    = {VSTTE 2012},
+  pages    =     {243--260},
+  year =         2012,
+  volume    =    7152,
+  series =       {LNCS},
+  publisher = {Springer}}
+
+@InProceedings{Foster14c,
+  author =       {Foster, S. and Zeyda, F. and Woodcock, J.},
+  title =        {{Isabelle/UTP}: A Mechanised Theory Engineering Framework},
+  booktitle = {UTP},
+  year =      2014,
+  series =    {LNCS 8963},
+  pages =     {21--41},
+  publisher = {Springer}}
+
+@Article{Foster15,
+  author =       {Foster, S. and Struth, G.},
+  title =        {On the fine-structure of {Regular Algebra}},
+  journal =      {J. Automated Reasoning},
+  year =         2015,
+  volume =    42,
+  number =    2}
+
+@InProceedings{Foster16a,
+  author =       {Foster, S. and Zeyda, F. and Woodcock, J.},
+  title =        {Unifying heterogeneous state-spaces with lenses},
+  booktitle =    {ICTAC},
+  year =         2016,
+  series =       {LNCS 9965},
+  publisher =    {Springer}}
+
+@InProceedings{Foster16b,
+  author =       {Foster, S. and Thiele, B. and Cavalcanti, A. and Woodcock, J.},
+  title =        {Towards a {UTP} semantics for {Modelica}},
+  booktitle =    {UTP},
+  year =         2016,
+  series =       {LNCS 10134},
+  publisher =    {Springer}}
+
+@InProceedings{Foster16c,
+  author =       {Zeyda, F. and Foster, S. and Freitas, L.},
+  title =        {An axiomatic value model for {Isabelle/UTP}},
+  booktitle =    {UTP},
+  year =         2016,
+  series =       {LNCS 10134},
+  publisher =    {Springer}}
+
+@Article{Foster17b,
+  author =       {Foster, S. and Cavalcanti, A. and Woodcock, J. and Zeyda, F.},
+  title =        {Unifying theories of time with generalised reactive processes},
+  journal =      {Accepted for Information Processing Letters},
+  year =         2017,
+  month =        {Dec},
+  note =         {Preprint: \url{https://arxiv.org/abs/1712.10213}}}
+
+@InProceedings{Foster2020-LocalVars,
+  author =       {Foster, S. and Baxter, J.},
+  title =        {Automated Algebraic Reasoning for Collections and Local Variables with Lenses},
+  booktitle = {{RAMiCS}},
+  year =      2020,
+  volume =    12062,
+  series =    {LNCS},
+  publisher = {Springer}}
+
+@Article{Foster2020-IsabelleUTP,
+  author =       {Foster, S. and Baxter, J. and Cavalcanti, A. and Woodcock, J. and Zeyda, F.},
+  title =        {Unifying Semantic Foundations for Automated Verification Tools in {Isabelle/UTP}},
+  journal =      {Science of Computer Programming},
+  volume =       {197},
+  month =        {October},
+  year =         {2020}}
+
+@Book{Hoare&98,
+  Title                    = {Unifying {Theories} of {Programming}},
+  Author                   = {Hoare, T. and He, J.},
+  Publisher                = {Prentice-Hall},
+  Year                     = {1998},
+  ISBN                     = {ISBN-10: 0134587618},
+  Comment                  = {NA PGL},
+  Key                      = {Hoare\&98}
+}
+
+@InProceedings{Huffman13,
+  Title                    = {Lifting and Transfer: A Modular Design for Quotients in {Isabelle/HOL}},
+  Author                   = {Huffman, B. and Kun\v{c}ar, O.},
+  Booktitle                = {CPP},
+  Year                     = {2013},
+  Pages                    = {131--146},
+  Publisher                = {Springer},
+  Series                   = {LNCS},
+  Volume                   = {8307}
+}
+
+@InProceedings{Klein2012,
+  Title = {Mechanised Separation Algebra},
+  author = {Klein, G. and Kolanski, R. and Boyton, A.},
+  booktitle = {Interactive Theorem Proving (ITP)},
+  year = {2012},
+  publisher = {Springer},
+  series = {LNCS},
+  volume = {7406},
+  pages = {332--337}
+} 
+
+@article{kozen1997kleene,
+  title={Kleene algebra with tests},
+  author={Kozen, D.},
+  journal={ACM Transactions on Programming Languages and Systems (TOPLAS)},
+  volume={19},
+  number={3},
+  pages={427--443},
+  year={1997},
+  publisher={ACM}
+}
+
+@PHDTHESIS{Oliveira2005-PHD,
+AUTHOR = {M. V. M. Oliveira},
+TITLE = {{Formal Derivation of State-Rich Reactive Programs using {\sf\textsl{Circus}}}},
+SCHOOL = {{Department of Computer Science - University of York}},
+YEAR = {2006},
+NOTE = {YCST-2006-02},
+ADDRESS = {UK},
+}
+
+@BOOK{Morgan90a,
+  KEY           = "Morgan90",
+  AUTHOR        = "Carroll Morgan",
+  TITLE         = "Programming from Specifications",
+  PUBLISHER     = "Prentice-Hall",
+  ADDRESS       = "London, UK",
+  YEAR          = "1990",
+  SIZE          = "255",
+  ANNOTE        = "",
+  COMMENT       = "NA PGL (borrowed from ID's library"}
+
+@article{Munive2024IsaVODES,
+    author = {Huerta y Munive, J. J. and Foster, S. and Gleirscher, M. and Struth, G. and Laursen, C. P. and Hickman, T.},
+    title = {{IsaVODEs: Interactive Verification of Cyber-Physical Systems at Scale}},
+    journal = {Journal of Automated Reasoning},
+    year = {2024}
+}
+
+@INPROCEEDINGS{Oliveira07,
+  author =       {Oliveira, M. and Cavalcanti, A. and Woodcock, J.},
+  title =        {{Unifying theories in ProofPower-Z}},
+  booktitle =    {UTP 2006},
+  pages =        {123--140},
+  year =         {2007},
+  volume =       {4010},
+  series    =    {LNCS},
+  publisher    = {Springer}
+}
+
+@article{Optics-AFP,
+  author  = {Simon Foster and Christian Pardillo-Laursen and Frank Zeyda},
+  title   = {Optics},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2017},
+  note    = {\url{https://isa-afp.org/entries/Optics.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}
+
+@BOOK{zrm,
+  AUTHOR       = {J.~M.~Spivey},
+  TITLE        = {{The Z Notation: A Reference Manual}},
+  YEAR         = {1998},
+  PUBLISHER    = {Prentice Hall},
+  annote       = {Standard Z notation and mathematical toolkit}
+}
+
+@article{Foster2024ITrees,
+    author = {Foster, S. and Hur, C.-K. and Woodcock, J.},
+    title = {Unifying Model Execution and Deductive Verification with {Interaction Trees} in {Isabelle/HOL}},
+    journal = {ACM Trans. on Software Engineering Methodology (TOSEM)},
+    year = {2024}
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Shallow_Expressions/document/root.tex
--- /dev/null
+++ thys/Shallow_Expressions/document/root.tex
@@ -0,0 +1,108 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+\usepackage{stmaryrd}
+  %for \<Sqinter>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{Shallow Expressions}
+\author{Simon Foster \\[.5ex] University of York, UK \\[2ex] \texttt{\small simon.foster@york.ac.uk}}
+\maketitle
+
+\begin{abstract}
+\noindent Most verification techniques use expressions, for example when assigning to variables or 
+forming assertions over the state. Deep embeddings provide such expressions using a datatype, which 
+allows queries over the syntax, such as calculating the free variables, and performing substitutions. 
+Shallow embeddings, in contrast, model expressions as query functions over the state type, and are 
+more amenable to automating verification tasks. The goal of this library is provide an intuitive 
+implementation of shallow expressions, which nevertheless provides many of the benefits of a deep 
+embedding. We harness the Optics library to provide an algebraic semantics for state variables,
+and use syntax translations to provide an intuitive lifted expression syntax. Furthermore, we
+provide a variety of meta-logic-style queries on expressions, such as dependencies
+on a state variable, and substitution of a variable for an expression. We also provide proof 
+methods, based on the simplifier, to automate the associated proof tasks.
+\end{abstract}
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+\section{Introduction}
+
+This session provides a library for expressions in shallow embeddings, based on the Optics
+package~\cite{Optics-AFP}. It provides the following key features:
+
+\begin{enumerate}
+  \item Parse and print translations for converting between intuitive expression syntax,
+    and state functions using lenses to model variables. The translation uses the type
+    system to determine whether each free variable in an expression is (1) a lens (i.e. a state 
+    variable); (2) another expression; (3) a literal, and gives the correct interpretation for each.
+    The lifting mechanism is manifested through the bracket notation, $(\_)_e$, but can usually be hidden behind syntax.
+  \item Syntax for complex state variable constructions, using the lens operators~\cite{Optics-AFP}, such as simultaneous 
+    assignment, hierarchical state, and initial/final state variables.
+  \item The ``unrestriction'' predicate~\cite{Oliveira07,Foster2020-IsabelleUTP}, $x \mathop{\sharp} e$, 
+    which characterises whether an expression $e$ depends on a
+    particular variable $x$ or not, based on the lens laws. It can often be used as a replacement for 
+    syntactically checking for free variables in a deep embedding, as needed in several verification
+    techniques.
+  \item Semantic substitution of variables for expressions, $e[v/x]$, with support for evaluation from 
+    the simplifier. A notation is provided for expressing substitution objects as a sequence of simultaneous 
+    variable assignments, $[x_1 \leadsto e_1, x_2 \leadsto e_2, \cdots]$.
+  \item Collection lenses, $x[i]$, which can be used to model updating a component of a larger structure,
+    for example mutating an element of an array by its index.
+  \item Supporting transformations and constructors for expressions, such as state extension, 
+    state restriction, and quantifiers.
+\end{enumerate}
+
+The majority of these concepts have been adapted from Isabelle/UTP~\cite{Foster2020-IsabelleUTP}, but 
+have been generalised for use in other Isabelle-based verification tools. Several proof methods are 
+provided, such as for discharging unrestriction conditions (\texttt{unrest}) and evaluating substitutions (\texttt{subst\_eval}).
+
+The Shallow Expressions library has been applied in the IsaVODEs tool~\cite{Munive2024IsaVODES}, for 
+verifying hybrid systems, and Isabelle/ITrees~\cite{Foster2024ITrees}, for verification of process-algebraic languages. 
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/FLT_Sufficient_Conditions.thy
--- /dev/null
+++ thys/Sophie_Germain/FLT_Sufficient_Conditions.thy
@@ -0,0 +1,184 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory FLT_Sufficient_Conditions
+  imports SG_Preliminaries Fermat3_4.Fermat4
+begin
+  (*>*)
+
+
+
+section \<open>Sufficient Conditions for FLT\<close>
+
+text \<open>Recall that FLT stands for ``Fermat's Last Theorem''.
+      FLT states that there is no nontrivial integer solutions to the equation
+      term\<open>(x :: int) ^ n + y ^ n = z ^ n\<close> for any natural number term\<open>(2 :: int) < n\<close>.
+      as soon as the natural number \<open>n\<close> is greater than term\<open>2 :: nat\<close>.
+      More formally: term\<open>(2 :: nat) < n \<Longrightarrow> \<nexists>x y z :: int. x ^ n + y ^ n = z ^ n\<close>.
+      We give here some sufficient conditions.\<close>
+
+subsection \<open>Coprimality\<close>
+
+text \<open>We first notice that it is sufficient to prove FLT for integers
+      \<open>x\<close>, \<open>y\<close> and \<open>z\<close> that are (setwise) const\<open>coprime\<close>.\<close>
+
+lemma (in semiring_Gcd) FLT_setwise_coprime_reduction :
+  assumes \<open>x ^ n + y ^ n = z ^ n\<close> \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close>
+  defines \<open>d \<equiv> Gcd {x, y, z}\<close>
+  shows \<open>(x div d) ^ n + (y div d) ^ n = (z div d) ^ n\<close> \<open>x div d \<noteq> 0\<close>
+    \<open>y div d \<noteq> 0\<close> \<open>z div d \<noteq> 0\<close> \<open>Gcd {x div d, y div d, z div d} = 1\<close>
+proof -
+  have \<open>d dvd x\<close> \<open>d dvd y\<close> \<open>d dvd z\<close> by (unfold d_def, rule Gcd_dvd; simp)+
+  thus \<open>x div d \<noteq> 0\<close> \<open>y div d \<noteq> 0\<close> \<open>z div d \<noteq> 0\<close>
+    by (simp_all add: \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close> dvd_div_eq_0_iff)
+
+  have \<open>{x div d, y div d, z div d} = (\<lambda>n. n div d) ` {x, y, z}\<close> by blast
+  thus \<open>Gcd {x div d, y div d, z div d} = 1\<close>
+    by (metis GCD_div_Gcd_is_one \<open>x div d \<noteq> 0\<close> d_def div_by_0)
+
+  from  \<open>x ^ n + y ^ n = z ^ n\<close> show \<open>(x div d) ^ n + (y div d) ^ n = (z div d) ^ n\<close>
+    by (metis \<open>d dvd x\<close> \<open>d dvd y\<close> \<open>d dvd z\<close> div_add div_power dvd_power_same)
+qed
+
+corollary (in semiring_Gcd) FLT_for_coprime_is_sufficient : 
+  \<open>\<nexists>x y z. x \<noteq> 0 \<and> y \<noteq> 0 \<and> z \<noteq> 0 \<and> Gcd {x, y, z} = 1 \<and> x ^ n + y ^ n = z ^ n \<Longrightarrow>
+   \<nexists>x y z. x \<noteq> 0 \<and> y \<noteq> 0 \<and> z \<noteq> 0 \<and> x ^ n + y ^ n = z ^ n\<close>
+  by (metis (no_types) FLT_setwise_coprime_reduction)
+
+\<comment> \<open>These very generic lemmas are of course working for integers.\<close>
+lemma \<open>OFCLASS(int, semiring_Gcd_class)\<close> by intro_classes
+
+
+
+text \<open>This version involving congruences will be useful later.\<close>
+
+lemma FLT_setwise_coprime_reduction_mod_version :
+  fixes x y z :: int
+  assumes \<open>x ^ n + y ^ n = z ^ n\<close> \<open>[x \<noteq> 0] (mod m)\<close> \<open>[y \<noteq> 0] (mod m)\<close> \<open>[z \<noteq> 0] (mod m)\<close>
+  defines \<open>d \<equiv> Gcd {x, y, z}\<close>
+  shows \<open>(x div d) ^ n + (y div d) ^ n = (z div d) ^ n\<close> \<open>[x div d \<noteq> 0] (mod m)\<close>
+    \<open>[y div d \<noteq> 0] (mod m)\<close> \<open>[z div d \<noteq> 0] (mod m)\<close> \<open>Gcd {x div d, y div d, z div d} = 1\<close>
+proof -
+  have \<open>d dvd x\<close> \<open>d dvd y\<close> \<open>d dvd z\<close> by (unfold d_def, rule Gcd_dvd; simp)+
+  show \<open>[x div d \<noteq> 0] (mod m)\<close>
+    by (metis \<open>d dvd x\<close> assms(2) cong_0_iff dvd_mult dvd_mult_div_cancel)
+  show \<open>[y div d \<noteq> 0] (mod m)\<close>
+    by (metis \<open>d dvd y\<close> assms(3) cong_0_iff dvd_mult dvd_mult_div_cancel)
+  show \<open>[z div d \<noteq> 0] (mod m)\<close>
+    by (metis \<open>d dvd z\<close> assms(4) cong_0_iff dvd_mult dvd_mult_div_cancel)
+
+  have \<open>{x div d, y div d, z div d} = (\<lambda>n. n div d) ` {x, y, z}\<close> by blast
+  thus \<open>Gcd {x div d, y div d, z div d} = 1\<close>
+    by (metis GCD_div_Gcd_is_one \<open>[x div d \<noteq> 0] (mod m)\<close> cong_refl d_def div_by_0)
+
+  from  \<open>x ^ n + y ^ n = z ^ n\<close> show \<open>(x div d) ^ n + (y div d) ^ n = (z div d) ^ n\<close>
+    by (metis \<open>d dvd x\<close> \<open>d dvd y\<close> \<open>d dvd z\<close> div_plus_div_distrib_dvd_right div_power dvd_power_same)
+qed
+
+
+
+text \<open>Actually, it is sufficient to prove FLT for integers
+      \<open>x\<close>, \<open>y\<close> and \<open>z\<close> that are pairwise const\<open>coprime\<close>\<close>
+
+lemma (in semiring_Gcd) FLT_setwise_coprime_imp_pairwise_coprime :
+  \<open>coprime x y\<close> if \<open>n \<noteq> 0\<close> \<open>x ^ n + y ^ n = z ^ n\<close> \<open>Gcd {x, y, z} = 1\<close>
+proof (rule ccontr)
+  assume \<open>\<not> coprime x y\<close>
+  with is_unit_gcd obtain d where \<open>\<not> is_unit d\<close> \<open>d dvd x\<close> \<open>d dvd y\<close> by blast
+  from \<open>d dvd x\<close> \<open>d dvd y\<close> have \<open>d ^ n dvd x ^ n\<close> \<open>d ^ n dvd y ^ n\<close>
+    by (simp_all add: dvd_power_same)
+  moreover from calculation \<open>x ^ n + y ^ n = z ^ n\<close> have \<open>d ^ n dvd z ^ n\<close>
+    by (metis dvd_add_right_iff)
+  moreover from \<open>Gcd {x, y, z} = 1\<close> have \<open>Gcd {x ^ n, y ^ n, z ^ n} = 1\<close>
+    by (simp add: gcd_exp_weak)
+  ultimately have \<open>is_unit (d ^ n)\<close> by (metis Gcd_2 Gcd_insert gcd_greatest)
+  with \<open>\<not> is_unit d\<close> show False by (metis is_unit_power_iff \<open>n \<noteq> 0\<close>)
+qed
+
+
+
+subsection \<open>Odd prime Exponents\<close>
+
+text \<open>From session\<open>Fermat3_4\<close>, FLT is already proven for term\<open>n = (4 :: nat)\<close>.
+      Using this, we can prove that it is sufficient to
+      prove FLT for const\<open>odd\<close> const\<open>prime\<close> exponents.\<close>
+
+lemma (in semiring_1_no_zero_divisors) FLT_exponent_reduction :
+  assumes \<open>x ^ n + y ^ n = z ^ n\<close> \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close> \<open>p dvd n\<close>
+  shows \<open>(x ^ (n div p)) ^ p + (y ^ (n div p)) ^ p = (z ^ (n div p)) ^ p\<close>
+    \<open>x ^ (n div p) \<noteq> 0\<close> \<open>y ^ (n div p) \<noteq> 0\<close> \<open>z ^ (n div p) \<noteq> 0\<close>
+proof -
+  from power_not_zero[OF \<open>x \<noteq> 0\<close>] show \<open>x ^ (n div p) \<noteq> 0\<close> .
+  from power_not_zero[OF \<open>y \<noteq> 0\<close>] show \<open>y ^ (n div p) \<noteq> 0\<close> .
+  from power_not_zero[OF \<open>z \<noteq> 0\<close>] show \<open>z ^ (n div p) \<noteq> 0\<close> .
+
+  from \<open>p dvd n\<close> have * : \<open>n = (n div p) * p\<close> by simp
+  from \<open>x ^ n + y ^ n = z ^ n\<close>
+  show \<open>(x ^ (n div p)) ^ p + (y ^ (n div p)) ^ p = (z ^ (n div p)) ^ p\<close>
+    by (subst (asm) (1 2 3) "*") (simp add: power_mult)
+qed
+
+lemma \<open>OFCLASS(int, semiring_1_no_zero_divisors_class)\<close> by intro_classes
+
+
+
+lemma odd_prime_or_four_factorE :
+  fixes n :: nat assumes \<open>2 < n\<close>
+  obtains p where \<open>p dvd n\<close> \<open>odd p\<close> \<open>prime p\<close> | \<open>4 dvd n\<close>
+proof -
+  assume hyp1 : \<open>p dvd n \<Longrightarrow> odd p \<Longrightarrow> prime p \<Longrightarrow> thesis\<close> for p
+  assume hyp2 : \<open>4 dvd n \<Longrightarrow> thesis\<close>
+
+  show thesis
+  proof (cases \<open>\<exists>p. p dvd n \<and> odd p \<and> prime p\<close>)
+    from hyp1 show \<open>\<exists>p. p dvd n \<and> odd p \<and> prime p \<Longrightarrow> thesis\<close> by blast
+  next
+    assume \<open>\<nexists>p. p dvd n \<and> odd p \<and> prime p\<close>
+    hence \<open>p \<in> prime_factors n \<Longrightarrow> p = 2\<close> for p
+      by (metis in_prime_factors_iff primes_dvd_imp_eq two_is_prime_nat)
+    then obtain k where \<open>prime_factorization n = replicate_mset k 2\<close>
+      by (metis set_mset_subset_singletonD singletonI subsetI)
+    hence \<open>n = 2 ^ k\<close> by (subst prod_mset_prime_factorization_nat[symmetric])
+        (use assms in simp_all)
+    with \<open>2 < n\<close> have \<open>1 < k\<close> by (metis nat_power_less_imp_less pos2 power_one_right)
+    with \<open>n = 2 ^ k\<close> have \<open>4 dvd n\<close>
+      by (metis Suc_leI dvd_power_iff_le numeral_Bit0_eq_double
+          power.simps(2) power_one_right verit_comp_simplify1(2))
+    with hyp2 show thesis by blast
+  qed
+qed
+
+
+
+text \<open>Finally, proving FLT for odd prime exponents is sufficient.\<close>
+
+corollary FLT_for_odd_prime_exponents_is_sufficient : 
+  \<open>\<nexists>x y z :: int. x \<noteq> 0 \<and> y \<noteq> 0 \<and> z \<noteq> 0 \<and> x ^ n + y ^ n = z ^ n\<close> if \<open>2 < n\<close>
+  and odd_prime_FLT :
+  \<open>\<And>p. odd p \<Longrightarrow> prime p \<Longrightarrow>
+        \<nexists>x y z :: int.  x \<noteq> 0 \<and> y \<noteq> 0 \<and> z \<noteq> 0 \<and> x ^ p + y ^ p = z ^ p\<close>
+proof (rule ccontr)
+  assume \<open>\<not> (\<nexists>x y z :: int. x \<noteq> 0 \<and> y \<noteq> 0 \<and> z \<noteq> 0 \<and> x ^ n + y ^ n = z ^ n)\<close>
+  then obtain x y z :: int
+    where \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close> \<open>x ^ n + y ^ n = z ^ n\<close> by blast
+  from odd_prime_or_four_factorE \<open>2 < n\<close>
+  consider p where \<open>p dvd n\<close> \<open>odd p\<close> \<open>prime p\<close> | \<open>4 dvd n\<close> by blast
+  thus False
+  proof cases
+    fix p assume \<open>p dvd n\<close> \<open>odd p\<close> \<open>prime p\<close>
+    from FLT_exponent_reduction [OF \<open>x ^ n + y ^ n = z ^ n\<close> \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close> \<open>p dvd n\<close>]
+      odd_prime_FLT[OF \<open>odd p\<close> \<open>prime p\<close>]
+    show False by blast
+  next
+    assume \<open>4 dvd n\<close>
+    from fermat_mult4[OF \<open>x ^ n + y ^ n = z ^ n\<close> \<open>4 dvd n\<close>] \<open>x \<noteq> 0\<close> \<open>y \<noteq> 0\<close> \<open>z \<noteq> 0\<close>
+    show False by (metis mult_eq_0_iff)
+  qed
+qed
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/ROOT
--- /dev/null
+++ thys/Sophie_Germain/ROOT
@@ -0,0 +1,16 @@
+chapter AFP
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+session Sophie_Germain = Fermat3_4 +
+  options [timeout = 600]
+  sessions Wlog
+  theories
+    SG_Introduction
+    FLT_Sufficient_Conditions
+    SG_Preliminaries
+    SG_Theorem
+    SG_Generalization
+  document_files
+    "root.tex"
+    "root.bib"
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/SG_Generalization.thy
--- /dev/null
+++ thys/Sophie_Germain/SG_Generalization.thy
@@ -0,0 +1,1126 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory SG_Generalization
+  imports SG_Theorem
+begin 
+  (*>*)
+
+
+section \<open>Sophie Germain's Theorem: generalized Version\<close>
+
+text \<open>The proof we give here is from cite\<open>kiefer2012theoreme\<close>.\<close>
+
+subsection \<open>Auxiliary Primes\<close>
+
+abbreviation non_consecutivity_condition :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close> (\<open>NC\<close>)
+  where \<open>NC p q \<equiv> \<nexists>x y :: int. [x \<noteq> 0] (mod q) \<and> [y \<noteq> 0] (mod q) \<and> [x ^ p = 1 + y ^ p] (mod q)\<close>
+
+lemma non_consecutivity_condition_bis :
+  \<open>NC p q \<longleftrightarrow> (\<nexists>x y a b. [a :: int \<noteq> 0] (mod q) \<and> [a ^ p = x] (mod q) \<and>
+                         [b :: int \<noteq> 0] (mod q) \<and> [b ^ p = y] (mod q) \<and> [x = 1 + y] (mod q))\<close>
+  by (simp add: cong_def) (metis mod_add_right_eq)
+
+
+abbreviation not_pth_power :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close> (\<open>PPP\<close>)
+  where \<open>PPP p q \<equiv> \<nexists>x :: int. [p = x ^ p] (mod q)\<close>
+
+
+
+definition auxiliary_prime :: \<open>nat \<Rightarrow> nat \<Rightarrow> bool\<close> (\<open>aux'_prime\<close>)
+  where \<open>aux_prime p q \<equiv> prime p \<and> prime q \<and> NC p q \<and> PPP p q\<close>
+
+lemma auxiliary_primeI :
+  \<open>\<lbrakk>prime p; prime q; NC p q; PPP p q\<rbrakk> \<Longrightarrow> aux_prime p q\<close>
+  unfolding auxiliary_prime_def by auto
+
+lemma auxiliary_primeD :
+  \<open>prime p\<close> \<open>prime q\<close> \<open>NC p q\<close> \<open>PPP p q\<close> if \<open>aux_prime p q\<close>
+  using that by (auto simp add: auxiliary_prime_def)
+
+
+text \<open>We do not give code equation yet, let us first work around these notions.\<close>
+
+
+
+lemma gen_mult_group_mod_prime_as_ord : \<open>ord p g = p - 1\<close>
+  if \<open>prime p\<close> \<open>{1 .. p - 1} = {g ^ k mod p |k. k \<in> UNIV}\<close>
+proof -
+  from that(2) have \<open>g mod p \<in> {1 .. p - 1}\<close>
+    by (simp add: set_eq_iff) (metis power_one_right)
+  hence \<open>[g \<noteq> 0] (mod p)\<close> by (simp add: cong_def)
+
+  with cong_0_iff prime_imp_coprime \<open>prime p\<close>
+  have \<open>ord p g = (LEAST d. 0 < d \<and> [g ^ d = 1] (mod p))\<close>
+    unfolding ord_def by auto
+  also have \<open>\<dots> = p - 1\<close>
+  proof (rule ccontr)
+    assume \<open>(LEAST d. 0 < d \<and> [g ^ d = 1] (mod p)) \<noteq> p - 1\<close>
+    with fermat_theorem \<open>prime p\<close> \<open>[g \<noteq> 0] (mod p)\<close>
+    obtain k where \<open>0 < k\<close> \<open>k < p - 1\<close> \<open>[g ^ k = 1] (mod p)\<close>
+      by (metis calculation cong_0_iff coprime_ord linorder_neqE_nat
+          lucas_coprime_lemma ord_minimal prime_gt_1_nat zero_less_diff)
+    { fix l m
+      have \<open>g ^ (m + (l * k)) mod p = (g ^ m mod p * ((g ^ k) ^ l mod p)) mod p\<close>
+        by (simp add: mod_mult_eq mult.commute power_add power_mult)
+      also from \<open>[g ^ k = 1] (mod p)\<close> have \<open>((g ^ k) ^ l mod p) = 1\<close>
+        by (metis cong_def cong_pow mod_if power_one prime_nat_iff \<open>prime p\<close>)
+      finally have \<open>g ^ (m + (l * k)) mod p = g ^ m mod p\<close> by simp
+    } note $ = this
+
+    have \<open>UNIV = (\<Union>l. {m + (l * k) |m. m \<in> {0..k - 1}})\<close>
+      by auto (metis Suc_pred \<open>0 < k\<close> add.commute div_mod_decomp mod_Suc_le_divisor)
+    hence \<open>{g ^ k mod p |k. k \<in> UNIV} =
+             (\<Union>l. {g ^ (m + (l * k)) mod p |m. m \<in> {0..k - 1}})\<close>
+      by (simp add: set_eq_iff) metis
+    also have \<open>\<dots> = {g ^ m mod p |m. m \<in> {0..k - 1}}\<close> by (simp add: "$") 
+    finally have \<open>card {g ^ k mod p |k. k \<in> UNIV} = card \<dots>\<close> by simp
+    also have \<open>{g ^ m mod p |m. m \<in> {0..k - 1}} =
+                 (\<lambda>m. g ^ m mod p) ` {0..k - 1}\<close> by auto
+    also from card_image_le have \<open>card \<dots> \<le> card {0..k - 1}\<close> by blast
+    also have \<open>\<dots> = k\<close> by (simp add: \<open>0 < k\<close>)
+    finally show False
+      by (metis that(2) \<open>k < p - 1\<close> card_atLeastAtMost diff_Suc_1 linorder_not_less)
+  qed
+  finally show \<open>ord p g = p - 1\<close> .
+qed
+
+
+lemma exists_nth_power_mod_prime_iff :
+  fixes p n assumes \<open>prime p\<close>
+  defines d_def : \<open>d \<equiv> gcd n (p - 1)\<close>
+  shows \<open>(\<exists>x :: int. [a = x ^ n] (mod p)) \<longleftrightarrow>
+         (n \<noteq> 0 \<and> [a = 0] (mod p)) \<or> [a ^ ((p - 1) div d) = 1] (mod p)\<close>
+proof (cases \<open>n = 0\<close>)
+  show \<open>n = 0 \<Longrightarrow> ?thesis\<close>
+    by (simp add: d_def)
+      (metis \<open>prime p\<close> Suc_0_not_prime_nat Suc_pred div_self power_one_right prime_gt_0_nat)
+next
+  show ?thesis if \<open>n \<noteq> 0\<close>
+  proof (cases \<open>[a = 0] (mod p)\<close>)
+    show \<open>[a = 0] (mod p) \<Longrightarrow> ?thesis\<close>
+      by (auto simp add: cong_def power_0_left \<open>n \<noteq> 0\<close> intro!: exI[of _ 0])
+  next
+    have \<open>0 < int p\<close> by (simp add: prime_gt_0_nat \<open>prime p\<close>)
+    from \<open>prime p\<close> residue_prime_mult_group_has_gen gen_mult_group_mod_prime_as_ord
+    obtain g where * : \<open>{1 .. p - 1} = {g ^ k mod p |k. k \<in> UNIV}\<close> and \<open>ord p g = p - 1\<close> by blast
+    have \<open>[g \<noteq> 0] (mod p)\<close>
+      by (metis \<open>ord p g = p - 1\<close> \<open>prime p\<close> nat_less_le ord_0_right_nat
+          ord_cong prime_nat_iff zero_less_diff)
+
+    show \<open>?thesis\<close> if \<open>[a \<noteq> 0] (mod p)\<close>
+    proof (rule iffI)
+      assume \<open>\<exists>x. [a = x ^ n] (mod p)\<close>
+      then obtain x where \<open>[a = x ^ n] (mod p)\<close> ..
+      from cong_less_unique_int[OF \<open>0 < int p\<close>, of x]
+      obtain y :: nat where \<open>0 \<le> y\<close> \<open>y < p\<close> \<open>[x = y] (mod p)\<close>
+        by (metis int_nat_eq less_eq_nat.simps(1) nat_less_iff)
+      from \<open>[a \<noteq> 0] (mod p)\<close> \<open>[a = x ^ n] (mod p)\<close> have \<open>[x \<noteq> 0] (mod p)\<close>
+        by (metis cong_pow cong_sym cong_trans power_0_left \<open>n \<noteq> 0\<close>)
+      with \<open>[x = y] (mod p)\<close> have \<open>y \<noteq> 0\<close> by (metis of_nat_0)
+      with \<open>0 \<le> y\<close> \<open>y < p\<close> have \<open>y \<in> {1 .. p - 1}\<close> by simp
+      with "*" \<open>[x = y] (mod p)\<close> zmod_int obtain k where \<open>[x = g ^ k] (mod p)\<close> by auto
+
+      with \<open>[a = x ^ n] (mod p)\<close> have \<open>[a = g ^ (k * n)] (mod p)\<close>
+        by (metis (no_types, lifting) cong_pow cong_trans of_nat_power power_mult)
+      hence \<open>[a ^ ((p - 1) div d) = (g ^ (k * n)) ^ ((p - 1) div d)] (mod p)\<close>
+        by (simp add: cong_pow)
+      moreover have \<open>[(g ^ (k * n)) ^ ((p - 1) div d) = g ^ (k * n * (p - 1) div d)] (mod p)\<close>
+        by (metis (no_types) d_def cong_refl div_mult_swap gcd_dvd2 power_mult)
+      moreover have \<open>[g ^ (k * n * (p - 1) div d) = (g ^ (k * n div d)) ^ (p - 1)] (mod p)\<close>
+        by (metis (no_types) d_def cong_def dvd_div_mult dvd_mult gcd_dvd1 power_mult)
+      moreover have \<open>[(g ^ (k * n div d)) ^ (p - 1) = 1] (mod p)\<close>
+        by (rule fermat_theorem[OF \<open>prime p\<close>])
+          (metis \<open>[g \<noteq> 0] (mod p)\<close> cong_0_iff prime_dvd_power_nat \<open>prime p\<close>)
+      ultimately have \<open>[a ^ ((p - 1) div d) = 1] (mod p)\<close>
+        by (metis (no_types, opaque_lifting) cong_def cong_int_iff of_nat_1)
+      thus \<open> n \<noteq> 0 \<and> [a = 0] (mod p) \<or> [a ^ ((p - 1) div d) = 1] (mod p)\<close> ..
+    next
+      assume \<open> n \<noteq> 0 \<and> [a = 0] (mod p) \<or> [a ^ ((p - 1) div d) = 1] (mod p)\<close>
+      with \<open>[a \<noteq> 0] (mod p)\<close> have \<open>[a ^ ((p - 1) div d) = 1] (mod p)\<close> by blast
+
+      from cong_less_unique_int[OF \<open>0 < int p\<close>, of a]
+      obtain b :: nat where \<open>0 \<le> b\<close> \<open>b < p\<close> \<open>[a = b] (mod p)\<close>
+        by (metis int_nat_eq less_eq_nat.simps(1) nat_less_iff)
+      from \<open>[a \<noteq> 0] (mod p)\<close> \<open>[a = b] (mod p)\<close> have \<open>b \<noteq> 0\<close> by (metis of_nat_0)
+      with \<open>0 \<le> b\<close> \<open>b < p\<close> have \<open>b \<in> {1 .. p - 1}\<close> by simp
+      with "*" have \<open>b \<in> {g ^ k mod p |k. k \<in> UNIV}\<close> by blast
+      with \<open>[a = b] (mod p)\<close> zmod_int obtain k where \<open>[a = g ^ k] (mod p)\<close> by auto
+      from this[THEN cong_pow, of \<open>(p - 1) div d\<close>] \<open>[a ^ ((p - 1) div d) = 1] (mod p)\<close>
+      have \<open>[(g ^ k) ^ ((p - 1) div d) = 1] (mod p)\<close>
+        by (simp flip: cong_int_iff) (metis (no_types) cong_def)
+      hence \<open>[g ^ (k * (p - 1) div d) = 1] (mod p)\<close>
+        by (metis (no_types) d_def div_mult_swap gcd_dvd2 power_mult)
+      hence \<open>p - 1 dvd k * (p - 1) div d\<close>
+        by (simp add: ord_divides' \<open>ord p g = p - 1\<close>)
+      hence \<open>d dvd k\<close>
+        by (metis \<open>prime p\<close> d_def div_mult_swap dvd_div_eq_0_iff dvd_mult_div_cancel
+            dvd_times_right_cancel_iff gcd_dvd2 less_numeral_extra(3) prime_gt_1_nat zero_less_diff)
+      then obtain l where \<open>k = l * d\<close> by (metis dvd_div_mult_self)
+      moreover from bezout_nat[OF \<open>n \<noteq> 0\<close>]
+      obtain u v where \<open>u * n = v * (p - 1) + d\<close> by (metis d_def mult.commute)
+      ultimately have \<open>l * u * n = l * v * (p - 1) + k\<close>
+        by (metis distrib_left mult.assoc)
+      hence \<open>(g ^ (l * u)) ^ n = (g ^ (l * v)) ^ (p - 1) * g ^ k\<close>
+        by (metis power_add power_mult)
+      hence \<open>[(g ^ (l * u)) ^ n = (g ^ (l * v)) ^ (p - 1) * g ^ k] (mod p)\<close> by simp
+      moreover have \<open>[(g ^ (l * v)) ^ (p - 1) = 1] (mod p)\<close>
+        by (metis \<open>ord p g = p - 1\<close> dvd_triv_right ord_divides power_mult)
+      ultimately have \<open>[(g ^ (l * u)) ^ n = g ^ k] (mod p)\<close>
+        by (metis cong_scalar_right cong_trans mult_1)
+      with \<open>[a = g ^ k] (mod p)\<close> have \<open>[a = (g ^ (l * u)) ^ n] (mod p)\<close>
+        by (meson cong_int_iff cong_sym cong_trans)
+      thus \<open>\<exists>x. [a = x ^ n] (mod p)\<close> by auto
+    qed
+  qed
+qed
+
+
+corollary not_pth_power_iff :
+  \<open>PPP p q \<longleftrightarrow> [p \<noteq> 0] (mod q) \<and> [p ^ ((q - 1) div gcd p (q - 1)) \<noteq> 1] (mod q)\<close>
+  if \<open>prime p\<close> \<open>prime q\<close>
+  by (subst exists_nth_power_mod_prime_iff[OF \<open>prime q\<close>, of p p])
+    (metis cong_int_iff not_prime_0 of_nat_0 of_nat_1 of_nat_power \<open>prime p\<close>)
+
+corollary not_pth_power_iff_mod :
+  \<open>PPP p q \<longleftrightarrow> \<not> q dvd p \<and> p ^ ((q - 1) div gcd p (q - 1)) mod q \<noteq> 1\<close>
+  if \<open>prime p\<close> and \<open>prime q\<close>
+  by (subst not_pth_power_iff[OF \<open>prime p\<close> \<open>prime q\<close>])
+    (simp add: cong_def mod_eq_0_iff_dvd prime_gt_Suc_0_nat)
+
+
+lemma non_consecutivity_condition_iff_enum_mod :
+  \<comment> \<open>This version is oriented towards code generation.\<close>
+  \<open>NC p q \<longleftrightarrow>
+   (\<forall>x \<in> {1..q - 1}. let x_p_mod = x ^ p mod q
+                      in \<forall>y \<in> {1..q - 1}. x_p_mod \<noteq> (1 + y ^ p mod q) mod q)\<close>
+  if \<open>q \<noteq> 0\<close>
+proof (unfold Let_def, intro iffI conjI ballI)
+  fix x y assume \<open>NC p q\<close> \<open>x \<in> {1..q - 1}\<close> \<open> y \<in> {1..q - 1}\<close>
+  show \<open>x ^ p mod q \<noteq> (1 + y ^ p mod q) mod q\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> x ^ p mod q \<noteq> (1 + y ^ p mod q) mod q\<close>
+    hence \<open>[x ^ p = 1 + y ^ p] (mod q)\<close>
+      by (simp add: cong_def) presburger
+    with \<open>NC p q\<close> have \<open>[x = 0] (mod q) \<or> [y = 0] (mod q)\<close>
+      by (metis (mono_tags, opaque_lifting) cong_int_iff int_ops(1)
+          of_nat_Suc of_nat_power_eq_of_nat_cancel_iff plus_1_eq_Suc)
+    with cong_less_modulus_unique_nat
+    have \<open>x \<notin> {1..q - 1} \<or> y \<notin> {1..q - 1}\<close> by force
+    with \<open>x \<in> {1..q - 1}\<close> \<open> y \<in> {1..q - 1}\<close> show False by blast
+  qed
+next
+  show \<open>NC p q\<close> if * : \<open>\<forall>x\<in>{1..q - 1}. \<forall>y\<in>{1..q - 1}. x ^ p mod q \<noteq> (1 + y ^ p mod q) mod q\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> NC p q\<close>
+    then obtain x y :: int where \<open>[x \<noteq> 0] (mod q)\<close> \<open>[y \<noteq> 0] (mod q)\<close>
+      \<open>[x ^ p = 1 + y ^ p] (mod q)\<close> by blast
+
+    from \<open>[x \<noteq> 0] (mod q)\<close> \<open>q \<noteq> 0\<close> have \<open>x mod q \<in> {1..q - 1}\<close>
+      by (simp add: cong_0_iff)
+        (metis linorder_not_le mod_by_1 mod_eq_0_iff_dvd
+          mod_pos_pos_trivial of_nat_0_less_iff pos_mod_sign)
+    then obtain x' :: nat where \<open>x' \<in> {1..q - 1}\<close> \<open>x' = x mod q\<close>
+      by (cases \<open>x mod q\<close>) simp_all
+
+    from \<open>[y \<noteq> 0] (mod q)\<close> \<open>q \<noteq> 0\<close> have \<open>y mod q \<in> {1..q - 1}\<close>
+      by (simp add: cong_0_iff)
+        (metis linorder_not_le mod_by_1 mod_eq_0_iff_dvd
+          mod_pos_pos_trivial of_nat_0_less_iff pos_mod_sign)
+    then obtain y' :: nat where \<open>y' \<in> {1..q - 1}\<close> \<open>y' = y mod q\<close>
+      by (cases \<open>y mod q\<close>) simp_all
+
+    from \<open>[x ^ p = 1 + y ^ p] (mod q)\<close>
+    have \<open>(x mod q) ^ p mod q = (1 + (y mod q) ^ p mod q) mod q\<close>
+      by (simp add: cong_def mod_add_right_eq power_mod)
+    hence \<open>x' ^ p mod q = (1 + y' ^ p mod q) mod q\<close>
+      by (metis \<open>x' = x mod int q\<close> \<open>y' = y mod int q\<close> nat_mod_as_int
+          of_nat_Suc of_nat_power plus_1_eq_Suc zmod_int)
+    with "*" \<open>x' \<in> {1..q - 1}\<close> \<open>y' \<in> {1..q - 1}\<close> show False by blast
+  qed
+qed
+
+
+lemma auxiliary_prime_iff_enum_mod [code] :
+  \<comment> \<open>We will have a more optimized version later.\<close>
+  \<open>aux_prime p q \<longleftrightarrow>
+   prime p \<and> prime q \<and>
+   \<not> q dvd p \<and> p ^ ((q - 1) div gcd p (q - 1)) mod q \<noteq> 1 \<and>
+   (\<forall>x \<in> {1..q - 1}. let x_p_mod = x ^ p mod q
+                      in \<forall>y \<in> {1..q - 1}. x_p_mod \<noteq> (1 + y ^ p mod q) mod q)\<close>
+proof (cases \<open>prime p\<close>; cases \<open>prime q\<close>)
+  assume \<open>prime p\<close> and \<open>prime q\<close>
+  from \<open>prime q\<close> have \<open>q \<noteq> 0\<close> by auto
+  show ?thesis
+    unfolding auxiliary_prime_def not_pth_power_iff_mod[OF \<open>prime p\<close> \<open>prime q\<close>]
+      non_consecutivity_condition_iff_enum_mod[OF \<open>q \<noteq> 0\<close>] by blast
+next
+  show \<open>\<not> prime q \<Longrightarrow> ?thesis\<close>
+    and \<open>\<not> prime q \<Longrightarrow> ?thesis\<close>
+    and \<open>\<not> prime p \<Longrightarrow> ?thesis\<close>
+    by (simp_all add: auxiliary_prime_def)
+qed
+
+
+text \<open>We can for example compute pairs of auxiliary primes less than term\<open>110 :: nat\<close>.\<close>
+
+value \<open>[(p, q). p \<leftarrow> [1..110], q \<leftarrow> [1..110], aux_prime (nat p) (nat q)]\<close>
+
+
+lemma auxiliary_primeI' :
+  \<open>\<lbrakk>prime p; prime q; \<not> q dvd p; p ^ ((q - 1) div gcd p (q - 1)) mod q \<noteq> 1;
+    \<And>x y. x \<in> {1..q - 1} \<Longrightarrow> y \<in> {1..q - 1} \<Longrightarrow> [x ^ p \<noteq> 1 + y ^ p] (mod q)\<rbrakk>
+   \<Longrightarrow> aux_prime p q\<close>
+  by (simp add: auxiliary_prime_iff_enum_mod cong_def mod_Suc_eq)
+
+
+
+lemma two_is_not_auxiliary_prime : \<open>\<not> aux_prime p 2\<close>
+  by (simp add: auxiliary_prime_iff_enum_mod) presburger
+
+
+lemma auxiliary_prime_of_2 : \<open>aux_prime 2 q \<longleftrightarrow> q = 3 \<or> q = 5\<close>
+proof (rule iffI)
+  show \<open>q = 3 \<or> q = 5 \<Longrightarrow> aux_prime 2 q\<close>
+  proof (elim disjE)
+    show \<open>q = 3 \<Longrightarrow> aux_prime 2 q\<close>
+    proof (intro auxiliary_primeI')
+      show \<open>prime (2 :: nat)\<close> and \<open>q = 3 \<Longrightarrow> prime q\<close> by simp_all
+    next
+      fix x y assume \<open>q = 3\<close> \<open>x \<in> {1..q - 1}\<close> \<open>y \<in> {1..q - 1}\<close>
+      hence \<open>x = 1 \<and> y = 1 \<or> x = 1 \<and> y = 2 \<or> x = 2 \<and> y = 1 \<or> x = 2 \<and> y = 2\<close>
+        by simp (metis le_Suc_eq le_antisym numeral_2_eq_2)
+      with \<open>q = 3\<close> show \<open>[x2 \<noteq> 1 + y2] (mod q)\<close>
+        by (fastforce simp add: cong_def)
+    next
+      show \<open>q = 3 \<Longrightarrow> \<not> q dvd 2\<close> by simp
+    next
+      show \<open>q = 3 \<Longrightarrow> 2 ^ ((q - 1) div gcd 2 (q - 1)) mod q \<noteq> 1\<close> by simp
+    qed
+  next
+    show \<open>q = 5 \<Longrightarrow> aux_prime 2 q\<close>
+    proof (intro auxiliary_primeI')
+      show \<open>prime (2 :: nat)\<close> and \<open>q = 5 \<Longrightarrow> prime q\<close> by simp_all
+    next
+      fix x y assume \<open>q = 5\<close> \<open>x \<in> {1..q - 1}\<close> \<open>y \<in> {1..q - 1}\<close>
+      hence \<open>(x = 1 \<or> x = 2 \<or> x = 3 \<or> x = 4) \<and> (y = 1 \<or> y = 2 \<or> y = 3 \<or> y = 4)\<close>
+        by (simp add: numeral_eq_Suc) linarith
+      with \<open>q = 5\<close> show \<open>[x2 \<noteq> 1 + y2] (mod q)\<close>
+        by (fastforce simp add: cong_def)
+    next
+      show \<open>q = 5 \<Longrightarrow> \<not> q dvd 2\<close> by simp
+    next
+      show \<open>q = 5 \<Longrightarrow> 2 ^ ((q - 1) div gcd 2 (q - 1)) mod q \<noteq> 1\<close> by simp
+    qed
+  qed
+next
+  assume aux_q : \<open>aux_prime 2 q\<close>
+  with two_is_not_auxiliary_prime have \<open>q \<noteq> 2\<close> by blast
+  show \<open>q = 3 \<or> q = 5\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> (q = 3 \<or> q = 5)\<close>
+    with \<open>q \<noteq> 2\<close> auxiliary_primeD(1-2)[OF aux_q] prime_prime_factor
+      le_neq_implies_less prime_ge_2_nat
+    have \<open>prime q\<close> \<open>odd q\<close> \<open>2 < q\<close> by metis+
+
+    hence \<open>5 \<le> q\<close>
+      by (metis Suc_1 \<open>\<not> (q = 3 \<or> q = 5)\<close> add.commute add_Suc_right eval_nat_numeral(3)
+          even_numeral not_less_eq_eq numeral_Bit0 order_antisym_conv plus_1_eq_Suc prime_ge_2_nat)
+    with \<open>prime q\<close> have \<open>gcd 4 q = (1 :: int)\<close>
+      by (metis coprime_imp_gcd_eq_1 eval_nat_numeral(3) gcd.commute less_Suc_eq
+          of_nat_1 order_less_le_trans prime_nat_iff'' zero_less_numeral)
+    with cong_solve_dvd_int obtain inv_4 :: int
+      where inv_4: \<open>[4 * inv_4 = 1] (mod q)\<close>
+      by (metis dvd_refl gcd_int_int_eq of_nat_numeral)
+    define x where \<open>x \<equiv> 1 + inv_4\<close>
+    define y where \<open>y \<equiv> 1 - inv_4\<close>
+
+    from inv_4 have \<open>[x2 = 1 + y2] (mod q)\<close>
+      by (simp add: x_def y_def power2_eq_square cong_iff_dvd_diff ring_class.ring_distribs)
+    moreover obtain x' y' :: nat where \<open>[x' = x] (mod q)\<close> \<open>[y' = y] (mod q)\<close>
+      by (metis \<open>prime q\<close> cong_less_unique_int cong_sym int_eq_iff of_nat_0_less_iff prime_gt_0_nat)
+    ultimately have \<open>[x'2 = 1 + y'2] (mod q)\<close>
+      by (simp flip: cong_int_iff)
+        (meson cong_add cong_pow cong_refl cong_sym_eq cong_trans)
+    moreover have \<open>[x' \<noteq> 0] (mod q)\<close>
+    proof (rule ccontr)
+      assume \<open>\<not> [x' \<noteq> 0] (mod q)\<close>
+      with \<open>[x' = x] (mod q)\<close> have \<open>[x = 0] (mod q)\<close>
+        by (metis cong_0_iff cong_dvd_iff int_dvd_int_iff)
+      hence \<open>[4 * x = 0] (mod q)\<close>
+        by (metis cong_scalar_left mult_zero_right)
+      with cong_add[OF cong_refl[of 4] inv_4] have \<open>q dvd 5\<close>
+        by (simp add: x_def) (metis cong_0_iff cong_dvd_iff int_ops(3) of_nat_dvd_iff)
+      with \<open>prime q\<close> have \<open>q = 5\<close> by (auto intro: primes_dvd_imp_eq)
+      with \<open>\<not> (q = 3 \<or> q = 5)\<close> show False by blast
+    qed
+    moreover have \<open>[y' \<noteq> 0] (mod q)\<close>
+    proof (rule ccontr)
+      assume \<open>\<not> [y' \<noteq> 0] (mod q)\<close>
+      with \<open>[y' = y] (mod q)\<close> have \<open>[y = 0] (mod q)\<close>
+        by (metis cong_0_iff cong_dvd_iff int_dvd_int_iff)
+      hence \<open>[4 * y = 0] (mod q)\<close>
+        by (metis cong_scalar_left mult_zero_right)
+      with cong_diff[OF cong_refl[of 4] inv_4] have \<open>q dvd 3\<close>
+        by (simp add: y_def) (metis cong_0_iff cong_dvd_iff int_ops(3) of_nat_dvd_iff)
+      with \<open>prime q\<close> have \<open>q = 3\<close> by (auto intro: primes_dvd_imp_eq)
+      with \<open>\<not> (q = 3 \<or> q = 5)\<close> show False by blast
+    qed
+    ultimately have \<open>[(int x')2 = 1 + (int y')2] (mod q) \<and>
+      [int x' \<noteq> 0] (mod q) \<and> [int y' \<noteq> 0] (mod q)\<close>
+      by (metis cong_int_iff of_nat_0 of_nat_Suc of_nat_power plus_1_eq_Suc)
+    with auxiliary_primeD(3) aux_q show False by blast
+  qed
+qed
+
+
+
+
+text \<open>An auxiliary prime \<open>q\<close> of \<open>p\<close> is generally of the form term\<open>q = 2 * n * p + 1\<close>.\<close>
+
+lemma auxiliary_prime_pattern_aux :
+  \<open>\<exists>x y. [x \<noteq> 0] (mod q) \<and> [y \<noteq> 0] (mod q) \<and> [x ^ p = 1 + y ^ p] (mod q)\<close>
+  if \<open>p \<noteq> 0\<close> \<open>prime q\<close> \<open>coprime p (q - 1)\<close> \<open>odd q\<close>
+proof -
+  from bezout_nat \<open>coprime p (q - 1)\<close> \<open>p \<noteq> 0\<close>
+  obtain u v where bez : \<open>u * p = 1 + v * (q - 1)\<close>
+    by (metis add.commute coprime_imp_gcd_eq_1 mult.commute)
+  have \<open>[x \<noteq> 0] (mod q) \<Longrightarrow> [(x ^ v) ^ (q - 1) = 1] (mod q)\<close> for x
+    by (meson cong_0_iff fermat_theorem prime_dvd_power \<open>prime q\<close>)
+  hence * : \<open>[(x ^ u) ^ p = x] (mod q)\<close> for x
+    by (fold power_mult, unfold bez, unfold power_add power_mult)
+      (metis cong_0_iff cong_def cong_scalar_left \<open>prime q\<close>
+        mult.right_neutral power_one_right prime_dvd_mult_iff)
+  obtain x0 y0 where \<open>[x0 \<noteq> 0] (mod q)\<close> \<open>[y0 \<noteq> 0] (mod q)\<close> \<open>[x0 = 1 + y0] (mod q)\<close>
+    by (metis \<open>odd q\<close> \<open>prime q\<close> cong_0_iff cong_refl not_prime_unit
+        one_add_one prime_prime_factor two_is_prime_nat)
+  from "*" this(3) have \<open>[(x0 ^ u) ^ p = 1 + (y0 ^ u) ^ p] (mod q)\<close>
+    by (metis cong_add_lcancel_nat cong_def)
+  moreover from \<open>[x0 \<noteq> 0] (mod q)\<close> \<open>[y0 \<noteq> 0] (mod q)\<close>
+  have \<open>[x0 ^ u \<noteq> 0] (mod q)\<close> \<open>[y0 ^ u \<noteq> 0] (mod q)\<close>
+    by (meson cong_0_iff prime_dvd_power \<open>prime q\<close>)+
+  ultimately show ?thesis by blast
+qed
+
+
+theorem auxiliary_prime_pattern :
+  \<open>p = 2 \<and> (q = 3 \<or> q = 5) \<or> odd p \<and> (\<exists>n\<ge>1. q = 2 * n * p + 1)\<close> if aux_p : \<open>aux_prime p q\<close>
+proof -
+  from auxiliary_prime_of_2 consider \<open>p = 2\<close> \<open>q = 3 \<or> q = 5\<close> | \<open>odd p\<close> \<open>q \<noteq> 2\<close>
+    by (metis aux_p auxiliary_prime_def prime_prime_factor two_is_not_auxiliary_prime)
+  thus ?thesis
+  proof cases
+    show \<open>p = 2 \<Longrightarrow> q = 3 \<or> q = 5 \<Longrightarrow> ?thesis\<close> by blast
+  next
+    assume \<open>odd p\<close> \<open>q \<noteq> 2\<close>
+    have \<open>2 < q\<close> \<open>odd q\<close>
+      by (use \<open>q \<noteq> 2\<close> auxiliary_prime_def le_neq_implies_less prime_ge_2_nat that in presburger)
+        (metis \<open>q \<noteq> 2\<close> auxiliary_prime_def prime_prime_factor that two_is_prime_nat)
+    from aux_p have \<open>prime p\<close> and \<open>prime q\<close> by (simp_all add: auxiliary_primeD(1-2))
+    from euler_criterion[OF \<open>prime q\<close> \<open>2 < q\<close>]
+    have * : \<open>[Legendre p q = p ^ ((q - 1) div 2)] (mod q)\<close> by simp
+    have \<open>\<not> coprime p (q - 1)\<close>
+      by (metis auxiliary_prime_def cong_0_iff coprime_iff_gcd_eq_1
+          div_by_1 fermat_theorem not_pth_power_iff aux_p)
+    with \<open>prime p\<close> prime_imp_coprime have \<open>p dvd q - 1\<close> by blast
+    then obtain k where \<open>q = k * p + 1\<close>
+      by (metis add.commute \<open>prime q\<close> dvd_div_mult_self
+          le_add_diff_inverse less_or_eq_imp_le prime_gt_1_nat)
+    with \<open>odd q\<close> \<open>odd p\<close> have \<open>even k\<close> by simp
+    then obtain n where \<open>k = 2 * n\<close> by fast
+    with \<open>q = k * p + 1\<close> have \<open>q = 2 * n * p + 1\<close> by blast
+    with \<open>2 < q\<close> have \<open>1 \<le> n\<close> 
+      by (metis One_nat_def add.commute less_one linorder_not_less
+          mult_is_0 one_le_numeral plus_1_eq_Suc)
+    with \<open>odd p\<close> \<open>q = 2 * n * p + 1\<close> show ?thesis by blast
+  qed
+qed
+
+
+
+lemma auxiliary_prime_imp_less : \<open>aux_prime p q \<Longrightarrow> p < q\<close>
+  by (auto dest: auxiliary_prime_pattern simp add: less_Suc_eq)
+
+lemma auxiliary_primeE :
+  assumes \<open>aux_prime p q\<close>
+  obtains \<open>p = 2\<close> \<open>q = 3\<close> | \<open>p = 2\<close> \<open>q = 5\<close>
+  | n where \<open>odd p\<close> \<open>1 \<le> n\<close> \<open>q = 2 * n * p + 1\<close>
+    \<open>NC p (2 * n * p + 1)\<close> \<open>PPP p (2 * n * p + 1)\<close>
+  by (metis assms auxiliary_prime_def auxiliary_prime_pattern)
+
+
+
+text \<open>With this, we can quickly eliminate numbers that cannot be auxiliary.\<close>
+
+declare auxiliary_prime_iff_enum_mod [code del]
+
+lemma auxiliary_prime_iff_enum_mod_optimized [code] :
+  \<open>aux_prime p q \<longleftrightarrow>
+   p = 2 \<and> (q = 3 \<or> q = 5) \<or>
+   p < q \<and>
+   2 * p dvd q - 1 \<and>
+   prime p \<and> prime q \<and>
+   \<not> q dvd p \<and> p ^ ((q - 1) div gcd p (q - 1)) mod q \<noteq> 1 \<and>
+   (\<forall>x \<in> {1..q - 1}. let x_p_mod = x ^ p mod q
+                      in \<forall>y \<in> {1..q - 1}. x_p_mod \<noteq> (1 + y ^ p mod q) mod q)\<close>
+  by (fold auxiliary_prime_iff_enum_mod)
+    (metis add_diff_cancel_right' auxiliary_prime_imp_less auxiliary_prime_of_2
+      auxiliary_prime_pattern dvd_refl even_mult_iff mult_dvd_mono)
+
+value \<open>[(p, q). p \<leftarrow> [1..1000], q \<leftarrow> [1..110], aux_prime (nat p) (nat q)]\<close>
+
+
+
+subsection \<open>Sophie Germain Primes are auxiliary\<close>
+
+text \<open>When \<open>p\<close> is an const\<open>odd\<close> const\<open>prime\<close> and term\<open>2 * p + 1 :: nat\<close> is also
+      a const\<open>prime\<close> (what we call a Sophie Germain const\<open>prime\<close>),
+      term\<open>2 * p + 1 :: nat\<close> is automatically an const\<open>auxiliary_prime\<close>.\<close>
+
+lemma SophGer_prime_imp_auxiliary_prime :
+  fixes p assumes \<open>SG p\<close> defines q_def: \<open>q \<equiv> 2 * p + 1\<close>
+  shows \<open>aux_prime p q\<close>
+proof (rule auxiliary_primeI)
+  from SophGer_primeD(2-3)[OF \<open>SG p\<close>]
+  show \<open>prime p\<close> and \<open>prime q\<close> by (unfold q_def)
+next
+  from SophGer_primeD[OF \<open>SG p\<close>, folded q_def]
+  have \<open>odd p\<close> \<open>prime p\<close> \<open>prime q\<close> \<open>prime (int q)\<close> \<open>p \<noteq> 0\<close> by simp_all
+  show \<open>NC p q\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> NC p q\<close>
+    then obtain x y :: int where \<open>[x \<noteq> 0] (mod q)\<close> \<open>[y \<noteq> 0] (mod q)\<close>
+      \<open>[x ^ p = 1 + y ^ p] (mod q)\<close> by blast
+    from SG_simps.p_th_power_mod_q \<open>[x \<noteq> 0] (mod q)\<close> \<open>SG p\<close> cong_0_iff q_def
+    consider \<open>[x ^ p = 1] (mod q)\<close> | \<open>[x ^ p = - 1] (mod q)\<close> by blast
+    thus False
+    proof cases
+      assume \<open>[x ^ p = 1] (mod q)\<close>
+      with \<open>[x ^ p = 1 + y ^ p] (mod q)\<close> have \<open>[y ^ p = 0] (mod q)\<close>
+        by (metis add.right_neutral cong_add_lcancel cong_sym cong_trans)
+      with \<open>[y \<noteq> 0] (mod q)\<close> show False
+        by (meson \<open>prime (int q)\<close> cong_0_iff prime_dvd_power_int)
+    next
+      assume \<open>[x ^ p = - 1] (mod q)\<close>
+      with \<open>[x ^ p = 1 + y ^ p] (mod q)\<close>
+      have \<open>[- 1 = 1 + y ^ p] (mod q)\<close> by (metis cong_def)
+      moreover have \<open>- (1::int) = 1 + - 2\<close> by force
+      ultimately have \<open>[y ^ p = - 2] (mod q)\<close>
+        by (metis cong_add_lcancel cong_sym)
+      with \<open>odd p\<close> cong_minus_minus_iff have \<open>[(- y) ^ p = 2] (mod q)\<close> by force
+      with cong_sym have \<open>\<exists>x :: int. [2 = x ^ p] (mod q)\<close> by blast
+      with \<open>p \<noteq> 0\<close> exists_nth_power_mod_prime_iff[OF \<open>prime q\<close>]
+      have \<open>([2 = 0] (mod q) \<or> [4 = 1] (mod q))\<close>
+        by (simp add: q_def flip: cong_int_iff)
+      hence \<open>p \<le> 1\<close>
+      proof (elim disjE)
+        from \<open>p \<noteq> 0\<close> show \<open>[2 = 0] (mod q) \<Longrightarrow> p \<le> 1\<close>
+          by (auto simp add: q_def cong_def)
+      next
+        from linorder_not_less show \<open>[4 = 1] (mod q) \<Longrightarrow> p \<le> 1\<close>
+          by (fastforce simp add: q_def cong_def)
+      qed
+      with \<open>SG p\<close> show False
+        by (metis \<open>prime p\<close> linorder_not_less prime_nat_iff)
+    qed
+  qed
+next
+  from \<open>SG p\<close>[THEN SophGer_primeD(3), folded q_def]
+  have \<open>prime q\<close> \<open>prime (int q)\<close> by simp_all
+  from SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]
+  have \<open>\<not> q dvd x \<Longrightarrow> [x ^ p = 1] (mod q) \<or> [x ^ p = - 1] (mod q)\<close> for x :: int
+    unfolding q_def .
+  moreover have \<open>[p \<noteq> (1 :: int)] (mod q)\<close> \<open>[p \<noteq> - 1] (mod q)\<close>
+    using SG_simps.notcong_p(1, 3)[OF \<open>SG p\<close>] cong_sym unfolding q_def by blast+
+  ultimately have \<open>\<not> q dvd x \<Longrightarrow> [p \<noteq> x ^ p] (mod q)\<close> for x :: int
+    using cong_trans by blast
+  moreover have \<open>q dvd x \<Longrightarrow> [p \<noteq> x ^ p] (mod q)\<close> for x :: int
+    by (metis SG_simps.pos Suc_eq_plus1 \<open>SG p\<close> cong_dvd_iff dvd_power dvd_trans
+        int_dvd_int_iff less_add_Suc1 mult.commute mult_2_right nat_dvd_not_less q_def)
+  ultimately show \<open>PPP p q\<close> by blast
+qed
+
+
+
+subsection \<open>Main Theorems\<close>
+
+theorem Sophie_Germain_auxiliary_prime :
+  \<open>q dvd x \<or> q dvd y \<or> q dvd z\<close>
+  if \<open>x ^ p + y ^ p = z ^ p\<close> and \<open>aux_prime p q\<close> for x y z :: int
+proof (rule ccontr)
+  assume not_dvd : \<open>\<not> (q dvd x \<or> q dvd y \<or> q dvd z)\<close>
+  from auxiliary_primeD[OF \<open>aux_prime p q\<close>]
+  have \<open>prime p\<close> \<open>prime q\<close> \<open>NC p q\<close> by simp_all
+
+  have \<open>coprime x q\<close>
+    by (meson coprime_commute not_dvd prime_imp_coprime prime_nat_int_transfer \<open>prime q\<close>)
+  with bezout_int[of x q] obtain inv_x v :: int where \<open>inv_x * x + v * q = 1\<close> by auto
+  hence inv_x : \<open>[x * inv_x = 1] (mod q)\<close> by (metis cong_iff_lin mult.commute)
+
+  from \<open>x ^ p + y ^ p = z ^ p\<close> have \<open>z ^ p = x ^ p + y ^ p\<close> ..
+  hence \<open>(inv_x * z) ^ p = (inv_x * x) ^ p + (inv_x * y) ^ p\<close>
+    by (metis distrib_left power_mult_distrib)
+  with inv_x have \<open>[(inv_x * z) ^ p = 1 + (inv_x * y) ^ p] (mod q)\<close>
+    by (metis cong_add_rcancel cong_pow mult.commute power_one)
+  moreover from inv_x \<open>prime q\<close> have \<open>[(inv_x * z) ^ p \<noteq> 0] (mod q)\<close>
+    by (metis cong_dvd_iff dvd_0_right not_dvd not_prime_unit
+        prime_dvd_mult_eq_int prime_dvd_power prime_nat_int_transfer)
+  moreover from inv_x \<open>prime q\<close> have \<open>[(inv_x * y) ^ p \<noteq> 0] (mod q)\<close>
+    by (metis cong_dvd_iff dvd_0_right not_dvd not_prime_unit
+        prime_dvd_mult_eq_int prime_dvd_power prime_nat_int_transfer)
+  moreover obtain y' z' :: int where \<open>[y' = inv_x * y] (mod q)\<close> \<open>[z' = inv_x * z] (mod q)\<close>
+    by (metis \<open>prime q\<close> cong_less_unique_int cong_sym int_eq_iff of_nat_0_less_iff prime_gt_0_nat)
+  ultimately show False
+    by (metis \<open>NC p q\<close> \<open>prime p\<close> \<open>prime q\<close> cong_0_iff
+        prime_dvd_power_iff prime_gt_0_nat prime_nat_int_transfer)
+qed
+
+
+
+theorem Sophie_Germain_generalization :
+  \<open>\<nexists>x y z :: int. x ^ p + y ^ p = z ^ p \<and>
+                  [x \<noteq> 0] (mod p2) \<and> [y \<noteq> 0] (mod p2) \<and> [z \<noteq> 0] (mod p2)\<close>
+  if odd_p : \<open>odd p\<close> and aux_prime : \<open>aux_prime p q\<close>
+proof (rule ccontr) \<comment> \<open>The proof is done by contradiction.\<close>
+  from \<open>aux_prime p q\<close> have prime_p : \<open>prime p\<close>
+    by (metis auxiliary_primeD(1))
+  hence not_p_0 : \<open>p \<noteq> 0\<close> and prime_int_p : \<open>prime (int p)\<close> by simp_all
+  from \<open>aux_prime p q\<close> have prime_q : \<open>prime q\<close>
+    by (metis auxiliary_primeD(2))
+  hence prime_int_q : \<open>prime (int q)\<close> by simp
+  from \<open>odd p\<close> \<open>prime p\<close> have p_ge_3 : \<open>3 \<le> p\<close>
+    by (simp add: numeral_eq_Suc)
+      (metis Suc_le_eq dvd_refl le_antisym not_less_eq_eq prime_gt_Suc_0_nat)
+
+  assume \<open>\<not> (\<nexists>x y z. x ^ p + y ^ p = z ^ p \<and> [x \<noteq> 0] (mod int (p2)) \<and>
+                     [y \<noteq> 0] (mod int (p2)) \<and> [z \<noteq> 0] (mod (int p)2))\<close>
+  then obtain x y z :: int
+    where fermat : \<open>x ^ p + y ^ p = z ^ p\<close>
+      and not_cong_0 : \<open>[x \<noteq> 0] (mod p2)\<close> \<open>[y \<noteq> 0] (mod p2)\<close> \<open>[z \<noteq> 0] (mod p2)\<close> by auto
+
+\<comment> \<open>We first assume without loss of generality that
+    term\<open>x\<close>, term\<open>y\<close> and term\<open>z\<close> are setwise const\<open>coprime\<close>.\<close>
+  let ?gcd = \<open>Gcd {x, y, z}\<close>
+  wlog coprime : \<open>?gcd = 1\<close> goal False generalizing x y z keeping fermat not_cong_0
+    using FLT_setwise_coprime_reduction_mod_version[OF fermat not_cong_0]
+      hypothesis by blast
+
+\<comment> \<open>Then we can deduce that term\<open>x\<close>, term\<open>y\<close> and term\<open>z\<close> are pairwise const\<open>coprime\<close>.\<close>
+  have coprime_x_y : \<open>coprime x y\<close>
+    by (fact FLT_setwise_coprime_imp_pairwise_coprime[OF \<open>p \<noteq> 0\<close> fermat coprime])
+  have coprime_y_z : \<open>coprime y z\<close>
+  proof (subst coprime_minus_right_iff[symmetric],
+      rule FLT_setwise_coprime_imp_pairwise_coprime[OF \<open>p \<noteq> 0\<close>])
+    from fermat \<open>odd p\<close> show \<open>y ^ p + (- z) ^ p = (- x) ^ p\<close> by simp
+  next
+    show \<open>Gcd {y, - z, - x} = 1\<close>
+      by (metis Gcd_insert coprime gcd_neg1_int insert_commute)
+  qed
+  have coprime_x_z : \<open>coprime x z\<close>
+  proof (subst coprime_minus_right_iff[symmetric],
+      rule FLT_setwise_coprime_imp_pairwise_coprime[OF \<open>p \<noteq> 0\<close>])
+    from fermat \<open>odd p\<close> show \<open>x ^ p + (- z) ^ p = (- y) ^ p\<close> by simp
+  next
+    show \<open>Gcd {x, - z, - y} = 1\<close> 
+      by (metis Gcd_insert coprime gcd_neg1_int insert_commute)
+  qed
+
+\<comment> \<open>From @{thm [show_question_marks = false] Sophie_Germain_auxiliary_prime}
+    we have that among term\<open>x\<close>, term\<open>y\<close> and term\<open>z\<close>,
+    one (and only one, see below) is a multiple of term\<open>q\<close>.\<close>
+  from Sophie_Germain_auxiliary_prime[OF fermat aux_prime]
+  have q_dvd_xyz : \<open>q dvd x \<or> q dvd y \<or> q dvd z\<close> .
+
+\<comment> \<open>Without loss of generality, we can assume that term\<open>x\<close> is a multiple of term\<open>q\<close>.\<close>
+  wlog q_dvd_z : \<open>q dvd z\<close> goal False generalizing x y z
+    keeping fermat not_cong_0 coprime_x_y coprime_y_z coprime_x_z
+  proof -
+    from negation q_dvd_xyz have \<open>q dvd x \<or> q dvd y\<close> by simp
+    thus False
+    proof (elim disjE)
+      show \<open>q dvd x \<Longrightarrow> False\<close>
+        by (erule hypothesis[of x \<open>- y\<close> z])
+          (use fermat not_cong_0 \<open>odd p\<close> in
+            \<open>simp_all add: cong_0_iff coprime_commute coprime_x_y coprime_x_z coprime_y_z\<close>)
+    next
+      show \<open>q dvd y \<Longrightarrow> False\<close>
+        by (erule hypothesis[of y \<open>- x\<close> z])
+          (use fermat not_cong_0 \<open>odd p\<close> in
+            \<open>simp_all add: cong_0_iff coprime_commute coprime_x_y coprime_x_z coprime_y_z\<close>)
+    qed
+  qed
+
+  define S where \<open>S \<equiv> \<lambda>y z :: int. \<Sum>k = 0..p - 1. (- z) ^ (p - 1 - k) * y ^ k\<close>
+
+\<comment> \<open>Now we prove that term\<open>x\<close>, term\<open>y\<close> or term\<open>x\<close> is dividable by term\<open>p\<close>.\<close>
+  have p_dvd_xyz : \<open>p dvd x \<or> p dvd y \<or> p dvd z\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> (p dvd x \<or> p dvd y \<or> p dvd z)\<close>
+    hence \<open>[x \<noteq> 0] (mod p)\<close> \<open>[y \<noteq> 0] (mod p)\<close> \<open>[z \<noteq> 0] (mod p)\<close> 
+      by (simp_all add: cong_0_iff)
+    from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- z\<close> x y]
+      coprime_x_y fermat \<open>[z \<noteq> 0] (mod p)\<close>
+    obtain a \<alpha> where \<open>x + y = a ^ p\<close> \<open>S x y = \<alpha> ^ p\<close>
+      by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+    from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- x\<close> z \<open>- y\<close>]
+      coprime_y_z fermat \<open>[x \<noteq> 0] (mod int p)\<close>
+    obtain b \<beta> where \<open>z - y = b ^ p\<close> \<open>S z (- y) = \<beta> ^ p\<close>
+      by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+    from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- y\<close> z \<open>- x\<close>]
+      coprime_x_z fermat \<open>[y \<noteq> 0] (mod p)\<close>
+    obtain c \<gamma> where \<open>z - x = c ^ p\<close> \<open>S z (- x) = \<gamma> ^ p\<close>
+      by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+
+    have \<open>a ^ p + b ^ p + c ^ p = x + y + (z - y) + (z - x)\<close>
+      by (simp add: \<open>x + y = a ^ p\<close> \<open>z - y = b ^ p\<close> \<open>z - x = c ^ p\<close>)
+    also have \<open>\<dots> = 2 * z\<close> by simp
+    also from \<open>q dvd z\<close> have \<open>[\<dots> = 0] (mod q)\<close> by (simp add: cong_0_iff)
+    finally have \<open>[a ^ p + b ^ p + c ^ p = 0] (mod q)\<close> .
+
+    have \<open>[a = 0] (mod q) \<or> [b = 0] (mod q) \<or> [c = 0] (mod q)\<close>
+    proof (rule ccontr)
+      assume \<open>\<not> ([a = 0] (mod q) \<or> [b = 0] (mod q) \<or> [c = 0] (mod q))\<close>
+      hence \<open>[a \<noteq> 0] (mod q)\<close> \<open>[b \<noteq> 0] (mod q)\<close> \<open>[c \<noteq> 0] (mod q)\<close> by simp_all
+      from \<open>[c \<noteq> 0] (mod q)\<close> have \<open>gcd c q = 1\<close>
+        by (meson aux_prime auxiliary_prime_def cong_0_iff coprime_iff_gcd_eq_1
+            residues_prime.p_coprime_right_int residues_prime_def)
+      from bezout_int[of c q, unfolded this]
+      obtain u v where \<open>u * c + v * int q = 1\<close> by blast
+      with \<open>[a \<noteq> 0] (mod q)\<close> have \<open>[u \<noteq> 0] (mod q)\<close>
+        by (metis cong_0_iff cong_dvd_iff cong_iff_lin dvd_mult mult.commute unit_imp_dvd)
+
+      from \<open>[a ^ p + b ^ p + c ^ p = 0] (mod q)\<close>
+      have \<open>[(u * a) ^ p + (u * b) ^ p + (u * c) ^ p = 0] (mod q)\<close>
+        by (simp add: power_mult_distrib)
+          (metis cong_scalar_left distrib_left mult.commute mult_zero_left)
+      also from \<open>u * c + v * int q = 1\<close> have \<open>u * c = 1 - v * int q\<close> by simp
+      finally have \<open>[(u * a) ^ p + (u * b) ^ p + (1 - v * int q) ^ p = 0] (mod q)\<close> .
+      moreover have \<open>[(1 - v * int q) ^ p = 1] (mod q)\<close>
+        by (metis add_uminus_conv_diff cong_0_iff cong_add_lcancel_0
+            cong_pow dvd_minus_iff dvd_triv_right power_one)
+      ultimately have \<open>[(u * a) ^ p + (u * b) ^ p + 1 = 0] (mod q)\<close>
+        by (meson cong_add_lcancel cong_sym cong_trans)
+      hence \<open>[1 + (u * b) ^ p = (- (u * a)) ^ p] (mod q)\<close>
+        by (simp add: \<open>odd p\<close> cong_iff_dvd_diff) presburger
+      hence \<open>[(- (u * a)) ^ p = 1 + (u * b) ^ p] (mod q)\<close> by (fact cong_sym)
+      moreover from \<open>[a \<noteq> 0] (mod q)\<close> \<open>[u \<noteq> 0] (mod q)\<close>
+        cong_prime_prod_zero_int[OF _ \<open>prime (int q)\<close>, of u a] cong_minus_minus_iff
+      have \<open>[- u * a \<noteq> 0] (mod q)\<close> by force
+      moreover from \<open>[b \<noteq> 0] (mod q)\<close> \<open>[u \<noteq> 0] (mod q)\<close>
+        cong_prime_prod_zero_int[OF _ \<open>prime (int q)\<close>, of u b]
+      have \<open>[u * b \<noteq> 0] (mod q)\<close> by blast
+      ultimately show False
+        using aux_prime auxiliary_primeD(3) by auto
+    qed
+    hence \<open>q dvd a\<close>
+    proof (elim disjE)
+      show \<open>[a = 0] (mod q) \<Longrightarrow> q dvd a\<close> by (simp add: cong_0_iff)
+    next
+      assume \<open>[b = 0] (mod q)\<close>
+      with \<open>z - y = b ^ p\<close> \<open>q dvd z\<close> \<open>prime p\<close> have \<open>q dvd y\<close>
+        by (metis cong_0_iff cong_dvd_iff cong_iff_dvd_diff
+            dvd_power dvd_trans prime_gt_0_nat)
+      with \<open>prime (int q)\<close> \<open>q dvd z\<close> \<open>coprime y z\<close> have False
+        by (metis coprime_def not_prime_unit)
+      thus \<open>q dvd a\<close> ..
+    next
+      assume \<open>[c = 0] (mod q)\<close>
+      with \<open>z - x = c ^ p\<close> \<open>q dvd z\<close> \<open>prime p\<close> have \<open>q dvd x\<close>
+        by (metis cong_0_iff cong_dvd_iff cong_iff_dvd_diff
+            dvd_power dvd_trans prime_gt_0_nat)
+      with \<open>prime (int q)\<close> \<open>q dvd z\<close> \<open>coprime x z\<close> have False
+        by (metis coprime_def not_prime_unit)
+      thus \<open>q dvd a\<close> ..
+    qed
+    with \<open>x + y = a ^ p\<close> \<open>p \<noteq> 0\<close> \<open>prime (int q)\<close> have \<open>[y = - x] (mod q)\<close>
+      by (metis add.commute add.inverse_inverse add_uminus_conv_diff
+          bot_nat_0.not_eq_extremum cong_iff_dvd_diff prime_dvd_power_int_iff)
+    hence \<open>[S x y = p * x ^ (p - 1)] (mod q)\<close>
+      unfolding S_def by (fact Sophie_Germain_lemma_computation_cong_simp[OF \<open>p \<noteq> 0\<close>])
+    moreover from \<open>z - x = c ^ p\<close> \<open>q dvd z\<close> have \<open>[x = (- c) ^ p] (mod q)\<close>
+      by (simp add: \<open>odd p\<close> cong_iff_dvd_diff)
+        (metis add_diff_cancel_left' diff_diff_eq2)
+    ultimately have \<open>[S x y = p * ((- c) ^ p) ^ (p - 1)] (mod q)\<close>
+      by (meson cong_pow cong_scalar_left cong_trans)
+    also have \<open>S x y = \<alpha> ^ p\<close> by (fact \<open>S x y = \<alpha> ^ p\<close>)
+    also have \<open>((- c) ^ p) ^ (p - 1) = ((- c) ^ (p - 1)) ^ p\<close>
+      by (metis mult.commute power_mult)
+    finally have \<open>[\<alpha> ^ p = p * ((- c) ^ (p - 1)) ^ p] (mod q)\<close> .
+
+    have \<open>gcd q ((- c) ^ (p - 1)) = 1\<close>
+      by (metis \<open>[x = (- c) ^ p] (mod int q)\<close> \<open>int q dvd z\<close> cong_dvd_iff
+          coprime_def coprime_imp_gcd_eq_1 coprime_x_z dvd_mult not_prime_unit
+          power_eq_if prime_imp_coprime_int \<open>p \<noteq> 0\<close> \<open>prime (int q)\<close>)
+    with bezout_int obtain u v
+      where \<open>u * int q + v * (- c) ^ (p - 1) = 1\<close> by metis
+    hence \<open>v * (- c) ^ (p - 1) = 1 - u * int q\<close> by simp
+    have \<open>[1 - u * int q = 1] (mod q)\<close> by (simp add: cong_iff_lin)
+    from \<open>[\<alpha> ^ p = p * ((- c) ^ (p - 1)) ^ p] (mod q)\<close>
+    have \<open>[(v * \<alpha>) ^ p = p * (v * (- c) ^ (p - 1)) ^ p] (mod q)\<close>
+      by (simp add: power_mult_distrib) (metis cong_scalar_left mult.left_commute)
+    with \<open>[1 - u * int q = 1] (mod int q)\<close> have \<open>[(v * \<alpha>) ^ p = p] (mod q)\<close>
+      by (unfold \<open>v * (- c) ^ (p - 1) = 1 - u * int q\<close>)
+        (metis cong_pow cong_scalar_left cong_trans mult.comm_neutral power_one)
+    with \<open>aux_prime p q\<close>[THEN auxiliary_primeD(4)] cong_sym show False by blast
+  qed
+
+\<comment> \<open>Without loss of generality, we can assume that it is term\<open>z\<close>.\<close>
+  wlog p_dvd_z : \<open>p dvd z\<close> goal False generalizing x y z S
+    keeping fermat not_cong_0 coprime_x_y coprime_y_z coprime_x_z S_def
+  proof -
+    from negation p_dvd_xyz have \<open>p dvd x \<or> p dvd y\<close> by simp
+    thus False
+    proof (elim disjE)
+      show \<open>p dvd x \<Longrightarrow> False\<close>
+        by (erule hypothesis[of x \<open>- y\<close> z])
+          (use fermat not_cong_0 \<open>odd p\<close> in
+            \<open>simp_all add: cong_0_iff coprime_commute coprime_x_y coprime_x_z coprime_y_z\<close>)
+    next
+      show \<open>p dvd y \<Longrightarrow> False\<close>
+        by (erule hypothesis[of y \<open>- x\<close> z])
+          (use fermat not_cong_0 \<open>odd p\<close> in
+            \<open>simp_all add: cong_0_iff coprime_commute coprime_x_y coprime_x_z coprime_y_z\<close>)
+    qed
+  qed
+
+\<comment> \<open>The rest of the proof consists in deducing that actually term\<open>p2\<close>
+    divides term\<open>z\<close>, which contradicts @{thm \<open>[z \<noteq> 0] (mod p2)\<close>}.\<close>
+  from \<open>p dvd z\<close> coprime_x_z coprime_y_z
+  have \<open>[x \<noteq> 0] (mod p)\<close> \<open>[y \<noteq> 0] (mod p)\<close>
+    by (simp_all add: cong_0_iff)
+      (meson not_coprimeI not_prime_unit \<open>prime (int p)\<close>)+
+
+  from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- x\<close> z \<open>- y\<close>]
+    coprime_y_z fermat \<open>[x \<noteq> 0] (mod int p)\<close>
+  obtain b \<beta> where \<open>z - y = b ^ p\<close> \<open>S z (- y) = \<beta> ^ p\<close>
+    by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+  from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- y\<close> z \<open>- x\<close>]
+    coprime_x_z fermat \<open>[y \<noteq> 0] (mod p)\<close>
+  obtain c \<gamma> where \<open>z - x = c ^ p\<close> \<open>S z (- x) = \<gamma> ^ p\<close>
+    by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+
+  from fermat have \<open>(- z) ^ p + x ^ p + y ^ p = 0\<close> by (simp add: \<open>odd p\<close>)
+  from Sophie_Germain_lemma_computation[OF \<open>odd p\<close>] fermat
+  have \<open>(x + y) * S x y = z ^ p\<close> by (simp add: S_def)
+  with \<open>[z \<noteq> 0] (mod p2)\<close> have \<open>x + y \<noteq> 0\<close> \<open>S x y \<noteq> 0\<close> by auto
+
+  define z' where \<open>z' \<equiv> z div p\<close>
+  from \<open>p dvd z\<close> \<open>[z \<noteq> 0] (mod p2)\<close> \<open>p \<noteq> 0\<close>
+  have \<open>z = z' * p\<close> \<open>[z' \<noteq> 0] (mod p)\<close>
+    by (simp_all add: cong_0_iff z'_def dvd_div_iff_mult power2_eq_square)
+
+  from Sophie_Germain_lemma_only_possible_prime_common_divisor[OF \<open>prime p\<close> _ \<open>coprime x y\<close>]
+  have \<open>\<exists>q\<in>#prime_factorization r. q \<noteq> p \<Longrightarrow> \<not> r dvd x + y \<or> \<not> r dvd S x y\<close> for r :: nat
+    unfolding S_def
+    by (meson dvd_trans in_prime_factors_iff int_dvd_int_iff
+        of_nat_eq_iff prime_nat_int_transfer)
+  from this[OF Ex_other_prime_factor[OF _ _ \<open>prime p\<close>]]
+  have \<open>r dvd x + y \<Longrightarrow> r dvd S x y \<Longrightarrow> r = 0 \<or> (\<exists>k. r = p ^ k)\<close> for r :: nat by auto
+  moreover have \<open>\<not> (p ^ k dvd x + y \<and> p ^ k dvd S x y)\<close> if \<open>1 < k\<close> for k
+  proof (rule ccontr)
+    assume \<open>\<not> (\<not> (p ^ k dvd x + y \<and> p ^ k dvd S x y))\<close>
+    moreover from \<open>1 < k\<close> have \<open>p2 dvd p ^ k\<close>
+      by (simp add: le_imp_power_dvd)
+    ultimately have \<open>p2 dvd x + y\<close> \<open>p2 dvd S x y\<close>
+      by (meson dvd_trans of_nat_dvd_iff)+
+    from \<open>p2 dvd x + y\<close> have \<open>[y = - x] (mod p2)\<close>
+      by (simp add: add.commute cong_iff_dvd_diff)
+    hence \<open>[S x y = p * x ^ (p - 1)] (mod p2)\<close>
+      unfolding S_def by (fact Sophie_Germain_lemma_computation_cong_simp[OF \<open>p \<noteq> 0\<close>])
+    moreover from \<open>[x \<noteq> 0] (mod p)\<close> \<open>z = z' * p\<close> \<open>[z \<noteq> 0] (mod p2)\<close> \<open>prime (int p)\<close>
+    have \<open>\<not> p2 dvd p * x ^ (p - 1)\<close>
+      by (metis cong_0_iff dvd_mult_cancel_left mult_zero_right
+          of_nat_power power2_eq_square prime_dvd_power_int)
+    ultimately show False using \<open>p2 dvd S x y\<close> cong_dvd_iff by blast
+  qed
+  ultimately have p_only_nontrivial_div :
+    \<open>r dvd x + y \<Longrightarrow> r dvd S x y \<Longrightarrow> r = 1 \<or> r = p\<close> for r :: nat
+    by (metis \<open>S x y \<noteq> 0\<close> dvd_0_left_iff less_one
+        linorder_neqE_nat of_nat_eq_0_iff power_0 power_one_right)
+      \<comment> \<open>term\<open>p\<close> is therefore the only possible nontrivial common divisor.\<close>
+
+  define mul_x_plus_y where \<open>mul_x_plus_y = multiplicity p (x + y)\<close>
+  define mul_S_x_y where \<open>mul_S_x_y = multiplicity p (S x y)\<close>
+  from \<open>(x + y) * S x y = z ^ p\<close>
+  have \<open>(x + y) * S x y = z' ^ p * p ^ p\<close>
+    by (simp add: \<open>z = z' * p\<close> power_mult_distrib)
+
+
+  have \<open>mul_x_plus_y + mul_S_x_y = multiplicity p (z ^ p)\<close>
+    unfolding mul_x_plus_y_def mul_S_x_y_def
+    by (metis \<open>(x + y) * S x y = z ^ p\<close> \<open>S x y \<noteq> 0\<close> \<open>x + y \<noteq> 0\<close> prime_def
+        prime_elem_multiplicity_mult_distrib \<open>prime (int p)\<close>)
+  also have \<open>z ^ p = z' ^ p * p ^ p\<close>
+    by (simp add: \<open>z = z' * p\<close> power_mult_distrib)
+  also have \<open>multiplicity p \<dots> = p\<close>
+    by (metis \<open>[z' \<noteq> 0] (mod int p)\<close> aux_prime auxiliary_prime_def cong_0_iff
+        mult_of_nat_commute multiplicity_decomposeI of_nat_eq_0_iff of_nat_power
+        prime_dvd_power_iff prime_gt_0_nat \<open>p \<noteq> 0\<close> \<open>prime (int p)\<close>)
+  finally have \<open>mul_x_plus_y + mul_S_x_y = p\<close> .
+
+  moreover have \<open>(2 \<le> mul_x_plus_y \<longrightarrow> mul_S_x_y    \<le> 1) \<and>
+                 (2 \<le> mul_S_x_y    \<longrightarrow> mul_x_plus_y \<le> 1)\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> ?thesis\<close>
+    hence \<open>2 \<le> mul_x_plus_y \<and> 2 \<le> mul_S_x_y\<close> by linarith
+    hence \<open>p2 dvd (x + y) \<and> p2 dvd (S x y)\<close>
+      by (simp add: mul_x_plus_y_def mul_S_x_y_def multiplicity_dvd')
+    thus False
+      by (metis cong_0_iff less_numeral_extra(3) one_eq_prime_power_iff
+          p_dvd_z p_only_nontrivial_div pos2 \<open>[z \<noteq> 0] (mod p2)\<close> \<open>prime p\<close>)
+  qed
+  ultimately consider \<open>mul_x_plus_y = p\<close> \<open>mul_S_x_y = 0\<close>
+    | \<open>mul_x_plus_y = 0\<close> \<open>mul_S_x_y = p\<close>
+    | \<open>mul_x_plus_y = 1\<close> \<open>mul_S_x_y = p - 1\<close>
+    | \<open>mul_x_plus_y = p - 1\<close> \<open>mul_S_x_y = 1\<close>
+    by (metis Nat.add_diff_assoc add_cancel_right_right add_diff_cancel_left'
+        diff_is_0_eq not_less_eq_eq one_add_one plus_1_eq_Suc)
+  thus False
+  proof cases
+
+    assume \<open>mul_x_plus_y = p\<close> \<open>mul_S_x_y = 0\<close>
+    from \<open>mul_x_plus_y = p\<close> have \<open>p dvd (x + y)\<close>
+      by (metis mul_x_plus_y_def not_dvd_imp_multiplicity_0 \<open>p \<noteq> 0\<close>)
+    hence \<open>[y = - x] (mod p)\<close>
+      by (simp add: add.commute cong_iff_dvd_diff)
+    hence \<open>[S x y = S x (- x)] (mod p)\<close>
+      unfolding S_def
+      by (meson cong_minus_minus_iff cong_pow cong_scalar_right cong_sum)
+    also have \<open>S x (- x) = (\<Sum>k = 0..p - 1. x ^ (p - 1))\<close>
+      unfolding S_def
+      by (rule sum.cong[OF refl], simp)
+        (metis One_nat_def diff_Suc_eq_diff_pred le_add_diff_inverse2 power_add)
+    also from \<open>p \<noteq> 0\<close> have \<open>\<dots> = p * x ^ (p - 1)\<close> by simp
+    finally have \<open>[S x y = p * x ^ (p - 1)] (mod p)\<close> .
+    with \<open>[x \<noteq> 0] (mod p)\<close> have \<open>p dvd S x y\<close> by (simp add: cong_dvd_iff)
+    with \<open>mul_S_x_y = 0\<close> show False
+      by (metis \<open>S x y \<noteq> 0\<close> mul_S_x_y_def not_one_le_zero not_prime_unit
+          power_dvd_iff_le_multiplicity power_one_right \<open>prime (int p)\<close>)
+  next
+
+    assume \<open>mul_x_plus_y = 0\<close> \<open>mul_S_x_y = p\<close>
+    from \<open>mul_S_x_y = p\<close> have \<open>p dvd S x y\<close>
+      by (metis mul_S_x_y_def not_dvd_imp_multiplicity_0 \<open>p \<noteq> 0\<close>)
+    from Sophie_Germain_lemma_computation[OF \<open>odd p\<close>]
+    have \<open>(x + y) * S x y = x ^ p + y ^ p\<close> unfolding S_def .
+    moreover from \<open>p dvd S x y\<close> have \<open>[(x + y) * S x y = 0] (mod p)\<close>
+      by (simp add: cong_0_iff)
+    moreover have \<open>[x ^ p + y ^ p = x + y] (mod p)\<close>
+    proof (rule cong_add)
+      have \<open>x ^ p = x * x ^ (p - 1)\<close>
+        by (simp add: power_eq_if \<open>p \<noteq> 0\<close>)
+      moreover from \<open>[x \<noteq> 0] (mod p)\<close> have \<open>[x ^ (p - 1)  = 1] (mod p)\<close>
+        using cong_0_iff fermat_theorem_int \<open>prime p\<close> by blast
+      ultimately show \<open>[x ^ p = x] (mod p)\<close>
+        by (metis cong_scalar_left mult.right_neutral)
+    next
+      have \<open>y ^ p = y * y ^ (p - 1)\<close>
+        by (simp add: power_eq_if \<open>p \<noteq> 0\<close>)
+      moreover from \<open>[y \<noteq> 0] (mod p)\<close> have \<open>[y ^ (p - 1)  = 1] (mod p)\<close>
+        using cong_0_iff fermat_theorem_int \<open>prime p\<close> by blast
+      ultimately show \<open>[y ^ p = y] (mod p)\<close>
+        by (metis cong_scalar_left mult.right_neutral)
+    qed
+    ultimately have \<open>p dvd x + y\<close>
+      by (simp add: cong_0_iff cong_dvd_iff)
+    with \<open>mul_x_plus_y = 0\<close> show False
+      by (metis \<open>x + y \<noteq> 0\<close> mul_x_plus_y_def multiplicity_eq_zero_iff
+          not_prime_unit \<open>prime (int p)\<close>)
+  next
+
+    define x_plus_y' where \<open>x_plus_y' \<equiv> (x + y) div p\<close>
+    define S_x_y' where \<open>S_x_y' \<equiv> (S x y) div p ^ (p - 1)\<close>
+    define s where \<open>s \<equiv> x + y\<close>
+    let ?f  = \<open>\<lambda>k. (p choose k) * (- x) ^ k * s ^ (p - k)\<close>
+    let ?f' = \<open>\<lambda>k. (p choose k) * (- x) ^ k * s ^ (p - 1 - k)\<close>
+
+    assume \<open>mul_x_plus_y = 1\<close> \<open>mul_S_x_y = p - 1\<close>
+    hence \<open>s = p * x_plus_y'\<close> \<open>S x y = p ^ (p - 1) * S_x_y'\<close>
+      by (simp_all add: s_def x_plus_y'_def S_x_y'_def)
+        (metis dvd_mult_div_cancel mul_x_plus_y_def multiplicity_dvd power_Suc0_right,
+          metis dvd_mult_div_cancel mul_S_x_y_def multiplicity_dvd)
+
+    from fermat have \<open>s * S x y = (s - x) ^ p + x ^ p\<close>
+      by (simp add: s_def \<open>(x + y) * S x y = z ^ p\<close>)
+    also have \<open>s - x = (- x + s)\<close> by simp
+    also have \<open>(- x + s) ^ p = (\<Sum>k\<le>p. (p choose k) * (- x) ^ k * s ^ (p - k))\<close>
+      by (fact binomial_ring)
+    also have \<open>\<dots> = (\<Sum>k\<in>{0..p - 1}. ?f k) + (\<Sum>k\<in>{p}. ?f k)\<close>
+      by (rule sum_Un_eq[symmetric])
+        (auto simp add: linorder_not_le prime_gt_0_nat \<open>prime p\<close>)
+    also have \<open>(\<Sum>k\<in>{0..p - 1}. ?f k) = (\<Sum>k\<in>{0}. ?f k) + (\<Sum>k\<in>{1..p - 1}. ?f k)\<close>
+      by (rule sum_Un_eq[symmetric]) auto
+    also have \<open>(\<Sum>k\<in>{1..p - 1}. ?f k) = (\<Sum>k\<in>{1..p - 2}. ?f k) + (\<Sum>k\<in>{p - 1}. ?f k)\<close>
+      by (rule sum_Un_eq[symmetric]) (use \<open>3 \<le> p\<close> in auto)
+    also have \<open>(\<Sum>k\<in>{0}. ?f k) = s * s ^ (p - 1)\<close> by (simp add: power_eq_if \<open>p \<noteq> 0\<close>)
+    also have \<open>(\<Sum>k\<in>{1..p - 2}. ?f k) = (\<Sum>k\<in>{1..p - 2}. s * ?f' k)\<close>
+      by (rule sum.cong, simp_all)
+        (metis Suc_diff_Suc diff_less less_2_cases_iff less_zeroE
+          linorder_neqE order.strict_iff_not power_Suc \<open>p \<noteq> 0\<close>)
+    also have \<open>\<dots> = s * (\<Sum>k\<in>{1..p - 2}. ?f' k)\<close>
+      by (simp add: mult.assoc sum_distrib_left)
+    also have \<open>(\<Sum>k\<in>{p - 1}. ?f k) = s * p * x ^ (p - 1)\<close>
+      by (simp del: One_nat_def, subst binomial_symmetric[symmetric])
+        (use \<open>3 \<le> p\<close> in \<open>auto simp add: \<open>odd p\<close>\<close>)
+    finally have \<open>s * S x y =
+                  s * (s ^ (p - 1) + (\<Sum>k\<in>{1..p - 2}. ?f' k) + p * x ^ (p - 1))\<close>
+      by (simp add: \<open>odd p\<close> distrib_left int_distrib(4))
+    hence S_x_y_developed : \<open>S x y = s ^ (p - 1) + (\<Sum>k\<in>{1..p - 2}. ?f' k) + p * x ^ (p - 1)\<close>
+      using \<open>x + y \<noteq> 0\<close> mult_cancel_left unfolding s_def by blast
+    have \<open>[p * x ^ (p - 1) = 0] (mod p2)\<close>
+    proof (rule cong_trans[OF _ cong_sym])
+      show \<open>[p * x ^ (p - 1) = 0 + 0 + p * x ^ (p - 1)] (mod p2)\<close> by simp
+    next
+      show \<open>[0 = 0 + 0 + p * x ^ (p - 1)] (mod p2)\<close>
+      proof (rule cong_trans)
+        have \<open>p ^ (p - 1) dvd S x y\<close>
+          by (simp add: \<open>S x y = p ^ (p - 1) * S_x_y'\<close>)
+        moreover from \<open>3 \<le> p\<close> have \<open>p ^ 2 dvd p ^ (p - 1)\<close>
+          by (auto intro: le_imp_power_dvd)
+        ultimately show \<open>[0 = S x y] (mod p2)\<close>
+          by (metis cong_0_iff cong_sym dvd_trans of_nat_dvd_iff)
+      next
+        show \<open>[S x y = 0 + 0 + p * x ^ (p - 1)] (mod p2)\<close>
+        proof (unfold S_x_y_developed, rule cong_add[OF _ cong_refl])
+          show \<open>[s ^ (p - 1) + (\<Sum>k = 1..p - 2. ?f' k) = 0 + 0] (mod p2)\<close>
+          proof (rule cong_add)
+            have \<open>p dvd s\<close> by (simp add: \<open>s = p * x_plus_y'\<close>)
+            hence \<open>p ^ (p - 1) dvd s ^ (p - 1)\<close> by (simp add: dvd_power_same)
+            moreover from \<open>3 \<le> p\<close> have \<open>p ^ 2 dvd p ^ (p - 1)\<close>
+              by (auto intro: le_imp_power_dvd)
+            ultimately show \<open>[s ^ (p - 1) = 0] (mod p2)\<close>
+              by (metis cong_0_iff dvd_trans of_nat_dvd_iff)
+          next
+            show \<open>[\<Sum>k = 1..p - 2. ?f' k = 0] (mod p2)\<close>
+            proof (rule cong_trans)
+              show \<open>[\<Sum>k = 1..p - 2. ?f' k = \<Sum>k\<in>{1..p - 2}. 0] (mod p2)\<close>
+              proof (rule cong_sum)
+                fix k assume \<open>k \<in> {1..p - 2}\<close>
+                from \<open>k \<in> {1..p - 2}\<close> have \<open>p dvd (p choose k)\<close>
+                  by (auto intro: dvd_choose_prime simp add: \<open>prime p\<close>)
+                moreover from \<open>k \<in> {1..p - 2}\<close> have \<open>p dvd s ^ (p - 1 - k)\<close>
+                  by (auto simp add: \<open>prime p\<close> \<open>s = p * x_plus_y'\<close> prime_dvd_power_int_iff)
+                ultimately show \<open>[?f' k = 0] (mod p2)\<close>
+                  by (simp add: cong_0_iff mult_dvd_mono power2_eq_square)
+              qed
+            next
+              show \<open>[\<Sum>k = 1..p - 2. 0 = 0] (mod int (p2))\<close> by simp
+            qed
+          qed
+        qed
+      qed
+    qed
+    hence \<open>p dvd x ^ (p - 1)\<close> by (simp add: cong_iff_dvd_diff power2_eq_square \<open>p \<noteq> 0\<close>)
+    with prime_dvd_power \<open>prime (int p)\<close> have \<open>p dvd x\<close> by blast
+    with \<open>[x \<noteq> 0] (mod p)\<close> show False by (simp add: cong_0_iff)
+  next
+
+    define x_plus_y' where \<open>x_plus_y' \<equiv> (x + y) div p ^ (p - 1)\<close>
+    define S_x_y' where \<open>S_x_y' \<equiv> (S x y) div p\<close>
+    assume \<open>mul_x_plus_y = p - 1\<close> \<open>mul_S_x_y = 1\<close>
+    with \<open>(x + y) * S x y = z ^ p\<close> \<open>mul_x_plus_y + mul_S_x_y = p\<close>
+    have \<open>x_plus_y' * S_x_y' = z' ^ p\<close>
+      unfolding x_plus_y'_def S_x_y'_def z'_def
+      by (metis (no_types, opaque_lifting)
+          div_mult_div_if_dvd div_power mul_S_x_y_def mul_x_plus_y_def
+          multiplicity_dvd of_nat_power p_dvd_z power_add power_one_right)
+    have \<open>coprime x_plus_y' S_x_y'\<close>
+    proof (rule ccontr)
+      assume \<open>\<not> coprime x_plus_y' S_x_y'\<close>
+      then obtain r where \<open>prime r\<close> \<open>r dvd x_plus_y'\<close> \<open>r dvd S_x_y'\<close>
+        by (metis coprime_absorb_left not_coprime_nonzeroE
+            prime_factor_int unit_imp_dvd zdvd1_eq)
+      from \<open>r dvd x_plus_y'\<close> have \<open>r dvd x + y\<close>
+        by (metis \<open>mul_x_plus_y = p - 1\<close> dvd_div_iff_mult dvd_mult_left
+            mul_x_plus_y_def multiplicity_dvd of_nat_eq_0_iff of_nat_power
+            power_not_zero \<open>p \<noteq> 0\<close> x_plus_y'_def)
+      moreover from \<open>r dvd S_x_y'\<close> have \<open>r dvd S x y\<close>
+        by (metis S_x_y'_def \<open>mul_S_x_y = 1\<close> dvd_mult_div_cancel dvd_trans
+            dvd_triv_right mul_S_x_y_def multiplicity_dvd power_one_right)
+      ultimately have \<open>r = p\<close>
+        by (metis \<open>prime r\<close> not_prime_1 p_only_nontrivial_div
+            pos_int_cases prime_gt_0_int prime_nat_int_transfer)
+      with \<open>[z' \<noteq> 0] (mod int p)\<close> \<open>r dvd S_x_y'\<close> \<open>prime p\<close> \<open>prime (int p)\<close>
+        \<open>x_plus_y' * S_x_y' = z' ^ p\<close> show False
+        by (metis  cong_0_iff dvd_trans dvd_triv_right prime_dvd_power_int_iff prime_gt_0_nat)
+    qed
+    from prod_is_some_powerE[OF \<open>coprime x_plus_y' S_x_y'\<close> \<open>x_plus_y' * S_x_y' = z' ^ p\<close>]
+    obtain a where \<open>normalize x_plus_y' = a ^ p\<close> by blast
+    moreover from prod_is_some_powerE
+      [OF coprime_commute[THEN iffD1, OF \<open>coprime x_plus_y' S_x_y'\<close>]]
+    obtain \<alpha> where \<open>normalize S_x_y' = \<alpha> ^ p\<close>
+      by (metis \<open>x_plus_y' * S_x_y' = z' ^ p\<close> mult.commute)
+    ultimately have \<open>x_plus_y' = (if 0 \<le> x_plus_y' then a ^ p else (- a) ^ p)\<close>
+      \<open>S_x_y' = (if 0 \<le> S_x_y' then \<alpha> ^ p else (- \<alpha>) ^ p)\<close>
+      by (metis \<open>odd p\<close> abs_of_nonneg abs_of_nonpos add.inverse_inverse
+          linorder_linear normalize_int_def power_minus_odd)+
+    then obtain \<alpha> a where \<open>x_plus_y' = a ^ p\<close> \<open>S_x_y' = \<alpha> ^ p\<close> by meson
+
+    from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- x\<close> z \<open>- y\<close>]
+      coprime_y_z fermat \<open>[x \<noteq> 0] (mod int p)\<close>
+    obtain b \<beta> where \<open>z - y = b ^ p\<close> \<open>S z (- y) = \<beta> ^ p\<close>
+      by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+    from Sophie_Germain_lemma[OF \<open>odd p\<close> \<open>prime p\<close>, of \<open>- y\<close> z \<open>- x\<close>]
+      coprime_x_z fermat \<open>[y \<noteq> 0] (mod p)\<close>
+    obtain c \<gamma> where \<open>z - x = c ^ p\<close> \<open>S z (- x) = \<gamma> ^ p\<close>
+      by (simp add: S_def \<open>odd p\<close> coprime_commute) (meson cong_0_iff dvd_minus_iff)
+
+    have \<open>\<not> p dvd c\<close>
+      by (metis \<open>[x \<noteq> 0] (mod int p)\<close> \<open>z - x = c ^ p\<close> cong_0_iff cong_dvd_iff
+          cong_iff_dvd_diff dvd_def dvd_trans p_dvd_z power_eq_if \<open>p \<noteq> 0\<close>)
+    have \<open>\<not> p dvd b\<close>
+      by (metis \<open>[y \<noteq> 0] (mod int p)\<close> \<open>z - y = b ^ p\<close> cong_0_iff cong_dvd_iff
+          cong_iff_dvd_diff dvd_def dvd_trans p_dvd_z power_eq_if \<open>p \<noteq> 0\<close>)
+
+    have \<open>p dvd 2 * z - x - y\<close>
+      by (metis \<open>mul_S_x_y = 1\<close> \<open>mul_x_plus_y + mul_S_x_y = p\<close> comm_monoid_add_class.add_0 diff_diff_eq dvd_diff int_ops(2)
+          mul_x_plus_y_def not_dvd_imp_multiplicity_0 one_dvd p_dvd_z prime_dvd_mult_iff wlog_keep.prime_int_p)
+    hence \<open>[2 * z - x - y = 0] (mod p)\<close> by (simp add: cong_0_iff)
+    also from \<open>z - x = c ^ p\<close> \<open>z - y = b ^ p\<close>
+    have \<open>2 * z - x - y = c ^ p + b ^ p\<close> by presburger
+    also have \<open>\<dots> = c ^ (p - 1) * c + b ^ (p - 1) * b\<close>
+      by (simp add: power_eq_if \<open>p \<noteq> 0\<close>)
+    finally have \<open>[c ^ (p - 1) * c + b ^ (p - 1) * b = 0] (mod p)\<close> .
+    moreover have \<open>[c ^ (p - 1) = 1] (mod p)\<close>
+      by (fact fermat_theorem_int[OF \<open>prime p\<close> \<open>\<not> p dvd c\<close>])
+    moreover have \<open>[b ^ (p - 1) = 1] (mod p)\<close>
+      by (fact fermat_theorem_int[OF \<open>prime p\<close> \<open>\<not> p dvd b\<close>])
+    ultimately have \<open>[1 * c + 1 * b = 0] (mod p)\<close>
+      by (meson cong_add cong_scalar_right cong_sym_eq cong_trans)
+    hence \<open>[c + b = 0] (mod p)\<close> by simp
+    hence \<open>[b = - c] (mod p)\<close> by (simp add: add.commute cong_iff_dvd_diff)
+    hence \<open>[S c b = p * c ^ (p - 1)] (mod p)\<close>
+      unfolding S_def
+      by (fact Sophie_Germain_lemma_computation_cong_simp[OF \<open>p \<noteq> 0\<close>])
+    hence \<open>p dvd S c b\<close> by (simp add: cong_dvd_iff)
+    moreover from \<open>[c + b = 0] (mod p)\<close> have \<open>p dvd c + b\<close> by (simp add: cong_dvd_iff)
+    moreover from Sophie_Germain_lemma_computation[OF \<open>odd p\<close>]
+    have \<open>c ^ p + b ^ p = (c + b) * S c b\<close> unfolding S_def ..
+    ultimately have \<open>p2 dvd c ^ p + b ^ p\<close>
+      by (simp add: mult_dvd_mono power2_eq_square)
+    moreover have \<open>p2 dvd x + y\<close>
+      by (metis \<open>mul_S_x_y = 1\<close> \<open>mul_x_plus_y + mul_S_x_y = p\<close> add_0
+          bot_nat_0.not_eq_extremum dvd_trans linorder_not_less
+          mul_x_plus_y_def multiplicity_dvd' nat_dvd_not_less odd_even_add
+          odd_p of_nat_power one_dvd prime_prime_factor \<open>prime p\<close>)
+    ultimately have \<open>p2 dvd 2 * z\<close>
+      by (metis \<open>2 * z - x - y = c ^ p + b ^ p\<close> diff_diff_eq zdvd_zdiffD)
+    moreover have \<open>coprime (p2) 2\<close> by (simp add: \<open>odd p\<close>)
+    ultimately have \<open>p2 dvd z\<close>
+      by (simp add: coprime_dvd_mult_left_iff mult.commute)
+    with \<open>[z \<noteq> 0] (mod p2)\<close> show False by (simp add: cong_0_iff)
+  qed
+qed
+
+
+
+text \<open>Since @{thm [show_question_marks = false] SophGer_prime_imp_auxiliary_prime},
+      this result is a generalization of
+      @{thm [show_question_marks = false] Sophie_Germain_theorem}.\<close>
+
+
+
+(*<*)
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/SG_Introduction.thy
--- /dev/null
+++ thys/Sophie_Germain/SG_Introduction.thy
@@ -0,0 +1,66 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory SG_Introduction
+  imports Main
+begin
+
+context fixes n p :: nat and x y z :: int
+begin
+
+(*>*)
+
+
+
+section \<open>Introduction\<close>
+
+text \<open>
+Fermat's Last Theorem (often abbreviated to FLT) states that for any integer 
+term\<open>2 < n\<close>, the equation term\<open>x ^ n + y ^ n = z ^ n\<close> has no nontrivial solution in the integers.
+Pierre de Fermat first conjectured this result in the 17th century,
+claiming to have a proof that did not fit in the margin of his notebook.
+However, it remained an open problem for centuries until Andrew Wiles and Richard Taylor provided
+a complete proof in 1995 using advanced techniques from algebraic geometry and modular forms.
+
+But in the meantime, many mathematicians have made partial progress on the problem.
+In particular, Sophie Germain's theorem states that term\<open>p\<close> is a prime
+such that term\<open>2 * p + 1\<close> is also a prime, then there are no integer
+solutions to the equation term\<open>x ^ p + y ^ p = z ^ p\<close> such that term\<open>p\<close>
+divides neither term\<open>x\<close>, term\<open>y\<close> nor term\<open>z\<close>.
+
+This result is not only included in the extended list of Freek's ``Top 100 theorems''
+footnote\<open>url\<open>http://www.cs.ru.nl/~freek/100/\<close>\<close>,
+but is also very familiar to students taking the French
+``agrégation'' mathematics competitive examination.
+Hoping that this submission might also be useful to them,
+we developed separately the classical version of the theorem
+as presented in cite\<open>Francinou_Gianella_Nicolas_2014\<close>
+and a generalization that one can find for example in cite\<open>kiefer2012theoreme\<close>.\<close>
+
+text \<open>
+\begin{figure}[h]
+\label{sessionGraph}
+\centering
+\includegraphics[scale=0.5]{session_graph} 
+\caption{Dependency graph of the session session\<open>Sophie_Germain\<close>}
+\end{figure}
+\newpage
+\<close>
+
+text \<open>
+The session displayed in \ref{sessionGraph} is organized as follows:
+  item verbatim\<open>FLT_Sufficient_Conditions\<close> provides sufficient conditions for proving FLT,
+  item verbatim\<open>SG_Premilinaries\<close> establish some useful lemmas and introduces the concept
+    of Sophie Germain prime,
+  item verbatim\<open>SG_Theorem\<close> proves Sophie Germain's theorem and
+  item verbatim\<open>SG_Generalization\<close> gives a generalization of it.
+\<close>
+
+
+
+(*<*)
+end
+
+end
+  (*>*)
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/SG_Preliminaries.thy
--- /dev/null
+++ thys/Sophie_Germain/SG_Preliminaries.thy
@@ -0,0 +1,221 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory SG_Preliminaries
+  imports "HOL-Number_Theory.Number_Theory" Wlog.Wlog
+begin
+  (*>*)
+
+
+
+section \<open>Preliminaries\<close>
+
+subsection \<open>Coprimality\<close>
+
+text \<open>We start with this useful elimination rule: when \<open>a\<close> and \<open>b\<close> are
+      not const\<open>coprime\<close> and are not both equal to term\<open>0 :: 'a :: factorial_semiring_gcd\<close>,
+      there exists some common const\<open>prime\<close> factor.\<close>
+
+lemma (in factorial_semiring_gcd) not_coprime_nonzeroE :
+  \<open>\<lbrakk>\<not> coprime a b; a \<noteq> 0 \<or> b \<noteq> 0; \<And>p. prime p \<Longrightarrow> p dvd a \<Longrightarrow> p dvd b \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis\<close>
+  by (metis gcd_eq_0_iff gcd_greatest_iff is_unit_gcd prime_divisor_exists)
+
+
+
+text \<open>Still referring to the notion of const\<open>coprime\<close> (but generalized to a set),
+      we prove that when term\<open>Gcd A \<noteq> 0\<close>, the elements of term\<open>{a div Gcd A |a. a \<in> A}\<close>
+      are setwise const\<open>coprime\<close>.\<close>
+
+lemma (in semiring_Gcd) GCD_div_Gcd_is_one :
+  \<open>(GCD a \<in> A. a div Gcd A) = 1\<close> if \<open>Gcd A \<noteq> 0\<close>
+proof (rule ccontr)
+  assume \<open>(GCD a\<in>A. a div Gcd A) \<noteq> 1\<close>
+  then obtain d where \<open>\<not> is_unit d\<close> \<open>\<forall>a\<in>A. d dvd (a div Gcd A)\<close>
+    by (metis (no_types, lifting) Gcd_dvd associated_eqI
+        image_eqI normalize_1 normalize_Gcd one_dvd)
+  from \<open>\<forall>a\<in>A. d dvd (a div Gcd A)\<close> have \<open>\<forall>a\<in>A. d * Gcd A dvd a\<close>
+    by (meson Gcd_dvd dvd_div_iff_mult \<open>Gcd A \<noteq> 0\<close>)
+  with Gcd_greatest have \<open>d * Gcd A dvd Gcd A\<close> by blast
+  with \<open>\<not> is_unit d\<close> show False by (metis div_self dvd_mult_imp_div that)
+qed
+
+
+
+subsection \<open>Power\<close>
+
+text \<open>Now we want to characterize the fact of admitting an \<open>n\<close>-th root
+      with a condition on the const\<open>multiplicity\<close> of each prime factor.\<close>
+
+lemma exists_nth_root_iff :
+  \<open>(\<exists>x. normalize y = x ^ n) \<longleftrightarrow> (\<forall>p\<in>prime_factors y. n dvd multiplicity p y)\<close>
+  if \<open>y \<noteq> 0\<close> for y :: \<open>'a :: factorial_semiring_multiplicative\<close>
+proof (rule iffI)
+  show \<open>\<exists>x. normalize y = x ^ n \<Longrightarrow> \<forall>p\<in>prime_factors y. n dvd multiplicity p y\<close>
+  proof (elim exE, rule ballI)
+    fix x p assume \<open>normalize y = x ^ n\<close> and \<open>p \<in> prime_factors y\<close>
+    hence \<open>p \<in> prime_factors x\<close>
+      by (metis prime_factorization_normalize empty_iff power_0 prime_factorization_1
+          prime_factors_power set_mset_empty zero_less_iff_neq_zero)
+    with \<open>normalize y = x ^ n\<close> show \<open>n dvd multiplicity p y\<close>
+      by (metis dvd_def in_prime_factors_iff multiplicity_normalize_right
+          normalization_semidom_class.prime_def prime_elem_multiplicity_power_distrib)
+  qed
+next
+  assume * : \<open>\<forall>p\<in>prime_factors y. n dvd multiplicity p y\<close>
+  define f where \<open>f p \<equiv> multiplicity p y div n\<close> for p
+  have \<open>normalize y = (\<Prod>p\<in>prime_factors y. p ^ multiplicity p y)\<close>
+    by (fact prod_prime_factors[OF \<open>y \<noteq> 0\<close>, symmetric])
+  also have \<open>\<dots> = (\<Prod>p\<in>prime_factors y. p ^ (f p * n))\<close>
+    by (rule prod.cong[OF refl]) (simp add: "*" f_def)
+  also have \<open>\<dots> = (\<Prod>p\<in>prime_factors y. p ^ f p) ^ n\<close>
+    by (simp add: power_mult prod_power_distrib)
+  finally show \<open>\<exists>x. normalize y = x ^ n\<close> ..
+qed
+
+
+
+text \<open>We use this result to obtain the following elimination rule.\<close>
+
+corollary prod_is_some_powerE :
+  fixes a b :: \<open>'a :: factorial_semiring_multiplicative\<close>
+  assumes \<open>coprime a b\<close> and \<open>a * b = x ^ n\<close>
+  obtains \<alpha> where \<open>normalize a = \<alpha> ^ n\<close>
+proof (cases \<open>a = 0\<close>)
+  from \<open>a * b = x ^ n\<close> show \<open>(\<And>\<alpha>. normalize a = \<alpha> ^ n \<Longrightarrow> thesis) \<Longrightarrow> a = 0 \<Longrightarrow> thesis\<close> by simp
+next
+  assume \<open>a \<noteq> 0\<close> and hyp : \<open>normalize a = \<alpha> ^ n \<Longrightarrow> thesis\<close> for \<alpha>
+  from \<open>a \<noteq> 0\<close> have \<open>\<exists>\<alpha>. normalize a = \<alpha> ^ n\<close>
+  proof (rule exists_nth_root_iff[THEN iffD2, rule_format])
+    fix p assume \<open>p \<in> prime_factors a\<close>
+    with \<open>a * b = x ^ n\<close> have \<open>p dvd x\<close>
+      by (metis dvd_mult2 in_prime_factors_iff prime_dvd_power)
+    hence \<open>p ^ n dvd x ^ n\<close> by (simp add: dvd_power_same)
+    with \<open>p \<in> prime_factors a\<close> \<open>a * b = x ^ n\<close> have \<open>n dvd multiplicity p (x ^ n)\<close>
+      by (metis dvd_triv_left gcd_nat.extremum in_prime_factors_iff multiplicity_unit_left
+          multiplicity_zero not_dvd_imp_multiplicity_0 power_0_left
+          prime_elem_multiplicity_power_distrib prime_imp_prime_elem)
+    also from \<open>p \<in> prime_factors a\<close> \<open>coprime a b\<close> \<open>a * b = x ^ n\<close>
+    have \<open>multiplicity p (x ^ n) = multiplicity p a\<close>
+      by (metis (no_types, opaque_lifting) add.right_neutral coprime_0_right_iff coprime_def
+          in_prime_factors_iff normalization_semidom_class.prime_def prime_factorization_empty_iff
+          prime_elem_multiplicity_eq_zero_iff prime_elem_multiplicity_mult_distrib)
+    finally show \<open>n dvd multiplicity p a\<close> .
+  qed
+  with hyp show thesis by blast
+qed
+
+
+
+subsection \<open>Sophie Germain Prime\<close>
+
+text \<open>Finally, we introduce Sophie Germain primes.\<close>
+
+definition SophGer_prime :: \<open>nat \<Rightarrow> bool\<close> (\<open>SG\<close>)
+  where \<open>SG p \<equiv> odd p \<and> prime p \<and> prime (2 * p + 1)\<close>
+
+lemma SophGer_primeI : \<open>odd p \<Longrightarrow> prime p \<Longrightarrow> prime (2 * p + 1) \<Longrightarrow> SG p\<close>
+  unfolding SophGer_prime_def by simp
+
+lemma SophGer_primeD : \<open>odd p\<close> \<open>prime p\<close> \<open>prime (2 * p + 1)\<close> if \<open>SG p\<close>
+  using that unfolding SophGer_prime_def by simp_all
+
+
+text \<open>We can easily compute Sophie Germain primes less than term\<open>2000 :: nat\<close>.\<close>
+
+value \<open>[p. p \<leftarrow> [0..2000], SG (nat p)]\<close>
+
+
+
+context fixes p assumes \<open>SG p\<close> begin
+
+local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "SG_simps")\<close>
+
+lemma nonzero : \<open>p \<noteq> 0\<close> using \<open>SG p\<close> by (simp add: odd_pos SophGer_primeD(1))
+
+lemma pos : \<open>0 < p\<close> using nonzero by blast
+
+lemma ge_3 : \<open>3 \<le> p\<close>
+  by (metis \<open>SG p\<close> SophGer_prime_def gcd_nat.order_iff_strict not_less_eq_eq
+      numeral_2_eq_2 numeral_3_eq_3 order_antisym_conv prime_ge_2_nat)
+
+lemma ge_7 : \<open>7 \<le> 2 * p + 1\<close> using ge_3 by auto
+
+
+lemma notcong_zero :
+  \<open>[- 3 \<noteq> 0 :: int] (mod 2 * p + 1)\<close> \<open>[- 1 \<noteq> 0 :: int] (mod 2 * p + 1)\<close>
+  \<open>[  1 \<noteq> 0 :: int] (mod 2 * p + 1)\<close> \<open>[  3 \<noteq> 0 :: int] (mod 2 * p + 1)\<close>
+  using SophGer_primeD(2)[OF \<open>SG p\<close>] 
+  by (simp_all add: cong_def zmod_zminus1_not_zero prime_nat_iff'')
+
+lemma notcong_p :
+  \<open>[- 1 \<noteq> p :: int] (mod 2 * p + 1)\<close>
+  \<open>[  0 \<noteq> p :: int] (mod 2 * p + 1)\<close>
+  \<open>[  1 \<noteq> p :: int] (mod 2 * p + 1)\<close>
+  using SophGer_primeD(2)[OF \<open>SG p\<close>]
+  by (auto simp add: pos cong_def zmod_zminus1_eq_if)
+
+
+lemma p_th_power_mod_q :
+  \<open>[m ^ p = 1] (mod 2 * p + 1) \<or> [m ^ p = - 1] (mod 2 * p + 1)\<close>
+  if \<open>\<not> 2 * p + 1 dvd m\<close> for m :: int
+proof -
+  wlog \<open>0 < m\<close> generalizing m keeping that
+    by (cases \<open>0 < - m\<close>)
+      (metis (no_types, opaque_lifting) \<open>\<not> 2 * p + 1 dvd m\<close> add.inverse_inverse
+        cong_minus_minus_iff dvd_minus_iff hypothesis uminus_power_if,
+        use \<open>\<not> 2 * p + 1 dvd m\<close> \<open>\<not> 0 < m\<close> in auto)
+
+  with \<open>\<not> 2 * p + 1 dvd m\<close> obtain n where \<open>m = int n\<close> \<open>\<not> 2 * p + 1 dvd n\<close>
+    by (metis int_dvd_int_iff pos_int_cases)
+  from \<open>0 < m\<close> have \<open>0 < m ^ p\<close> by simp
+
+  have \<open>[m ^ (2 * p) = n ^ (2 * p)] (mod 2 * p + 1)\<close> by (simp add: \<open>m = int n\<close>)
+  moreover have \<open>[n ^ (2 * p) = 1] (mod 2 * p + 1)\<close>
+    by (metis SophGer_prime_def \<open>SG p\<close> \<open>\<not> 2 * p + 1 dvd n\<close> add_implies_diff fermat_theorem)
+  ultimately have \<open>[m ^ (2 * p) = 1] (mod 2 * p + 1)\<close> by (metis cong_def int_ops(2) zmod_int)
+  also have \<open>m ^ (2 * p) = m ^ p * m ^ p\<close> by (simp add: mult_2 power_add)
+  finally have \<open>[m ^ p * m ^ p = 1] (mod 2 * p + 1)\<close> .
+  thus \<open>[m ^ p = 1] (mod 2 * p + 1) \<or> [m ^ p = - 1] (mod 2 * p + 1)\<close>
+    by (meson SophGer_primeD(3) \<open>0 < m ^ p\<close> \<open>SG p\<close> cong_square prime_nat_int_transfer)
+qed
+
+
+
+end
+
+
+
+subsection \<open>Fermat's little Theorem for Integers\<close>
+
+lemma fermat_theorem_int :
+  \<open>[a ^ (p - 1) = 1] (mod p)\<close> if \<open>prime p\<close> and \<open>\<not> p dvd a\<close>
+  for p :: nat and a :: int
+proof (cases a)
+  show \<open>a = int n \<Longrightarrow> [a ^ (p - 1) = 1] (mod p)\<close> for n
+    by (metis cong_int_iff fermat_theorem int_dvd_int_iff of_nat_1 of_nat_power that)
+next
+  fix n assume \<open>a = - int (Suc n)\<close>
+  from \<open>prime p\<close> have \<open>p = 2 \<or> odd p\<close>
+    by (metis prime_prime_factor two_is_prime_nat)
+  thus \<open>[a ^ (p - 1) = 1] (mod p)\<close>
+  proof (elim disjE)
+    assume \<open>p = 2\<close>
+    with \<open>\<not> p dvd a\<close> have \<open>[a = 1] (mod p)\<close> by (simp add: cong_iff_dvd_diff)
+    with \<open>p = 2\<close> show \<open>[a ^ (p - 1) = 1] (mod p)\<close> by simp
+  next
+    assume \<open>odd p\<close>
+    hence \<open>even (p - 1)\<close> by simp
+    hence \<open>a ^ (p - 1) = (int (Suc n)) ^ (p - 1)\<close>
+      by (metis \<open>a = - int (Suc n)\<close> uminus_power_if)
+    also have \<open>[\<dots> = 1] (mod p)\<close>
+      by (metis \<open>a = - int (Suc n)\<close> cong_int_iff dvd_minus_iff
+          fermat_theorem int_dvd_int_iff of_nat_1 of_nat_power that)
+    finally show \<open>[a ^ (p - 1) = 1] (mod p)\<close> .
+  qed
+qed
+  
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/SG_Theorem.thy
--- /dev/null
+++ thys/Sophie_Germain/SG_Theorem.thy
@@ -0,0 +1,315 @@
+
+(* Author: Benoît Ballenghien, Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF *)
+
+(*<*)
+theory SG_Theorem
+  imports FLT_Sufficient_Conditions
+begin
+  (*>*)
+
+
+section \<open>Sophie Germain's Theorem: classical Version\<close>
+
+text \<open>The proof we give here is from cite\<open>Francinou_Gianella_Nicolas_2014\<close>.\<close>
+
+subsection \<open>A Crucial Lemma\<close>
+
+lemma Sophie_Germain_lemma_computation : 
+  fixes x y :: int assumes \<open>odd p\<close>
+  defines \<open>S \<equiv> \<Sum>k = 0..p - 1. (- y) ^ (p - 1 - k) * x ^ k\<close>
+  shows \<open>(x + y) * S = x ^ p + y ^ p\<close>
+proof -
+  from \<open>odd p\<close> have \<open>0 < p\<close> by (simp add: odd_pos)
+
+  from int_distrib(1) have \<open>(x + y) * S = x * S - (- y) * S\<close> by auto
+  have \<open>x * S = (\<Sum>k = 0..p - 1. (- y) ^ (p - 1 - k) * x ^ (k + 1))\<close>
+    by (unfold S_def, subst sum_distrib_left) (rule sum.cong[OF refl], simp)
+  also have \<open>\<dots> = (\<Sum>k = 0..p - 1. (- y) ^ (p - (k + 1)) * x ^ (k + 1))\<close> by simp
+  also have \<open>\<dots> = x ^ p + (\<Sum>k = 1..p - 1. (- y) ^ (p - k) * x ^ k)\<close>
+    by (subst sum.shift_bounds_cl_nat_ivl[symmetric])
+      (simp, metis One_nat_def \<open>0 < p\<close> not_gr0 power_eq_if)
+  finally have S1 : \<open>x * S = x ^ p + (\<Sum>k = 1..p - 1. (- y) ^ (p - k) * x ^ k)\<close> .
+
+  have \<open>k \<in> {0..p - 1} \<Longrightarrow> (- y) ^ Suc (p - 1 - k) * x ^ k = (- y) ^ (p - k) * x ^ k\<close> for k
+    by (rule arg_cong[where f = \<open>\<lambda>n. (- y) ^ n * x ^ _\<close>])
+      (metis Suc_diff_le Suc_pred' \<open>0 < p\<close> atLeastAtMost_iff)
+  hence \<open>(- y) * S = (\<Sum>k = 0..p - 1. (- y) ^ (p - k) * x ^ k)\<close>
+    by (unfold S_def, subst sum_distrib_left, intro sum.cong[OF refl])
+      (subst mult.assoc[symmetric], subst power_Suc[symmetric], simp)
+  also have \<open>\<dots> = (- y) ^ (p - 0) * x ^ 0 + (\<Sum>k = 1..p - 1. (- y) ^ (p - k) * x ^ k)\<close>
+    by (unfold One_nat_def, subst sum.atLeast_Suc_atMost[symmetric]) simp_all
+  also have \<open>(- y) ^ (p - 0) * x ^ 0 = - (y ^ p)\<close>
+    by (simp add: \<open>odd p\<close>)
+  finally have S2 : \<open>- y * S = - (y ^ p) + (\<Sum>k = 1..p - 1. (- y) ^ (p - k) * x ^ k)\<close> .
+
+  show \<open>(x + y) * S = x ^ p + y ^ p\<close>
+    unfolding \<open>(x + y) * S = x * S - (- y) * S\<close> S1 S2 by simp
+qed
+
+lemma Sophie_Germain_lemma_computation_cong_simp :
+  fixes p :: nat and n x y :: int assumes \<open>p \<noteq> 0\<close> \<open>[y = - x] (mod n)\<close>
+  defines \<open>S \<equiv> \<lambda>x y. \<Sum>k = 0..p - 1. (- y) ^ (p - 1 - k) * x ^ k\<close>
+  shows \<open>[S x y = p * x ^ (p - 1)] (mod n)\<close>
+proof -
+  from \<open>[y = - x] (mod n)\<close> have \<open>[S x y = S x (- x)] (mod n)\<close>
+    unfolding S_def
+    by (meson cong_minus_minus_iff cong_pow cong_scalar_right cong_sum)
+  also have \<open>S x (- x) = (\<Sum>k = 0..p - 1. x ^ (p - 1))\<close>
+    unfolding S_def
+    by (rule sum.cong[OF refl], simp)
+      (metis One_nat_def diff_Suc_eq_diff_pred le_add_diff_inverse2 power_add)
+  also from \<open>p \<noteq> 0\<close> have \<open>\<dots> = p * x ^ (p - 1)\<close> by simp
+  finally show \<open>[S x y = p * x ^ (p - 1)] (mod n)\<close> .
+qed
+
+lemma Sophie_Germain_lemma_only_possible_prime_common_divisor :
+  fixes x y z :: int and p :: nat
+  defines S_def: \<open>S \<equiv> \<lambda>x y. \<Sum>k = 0..p - 1. (- y) ^ (p - 1 - k) * x ^ k\<close>
+  assumes \<open>prime p\<close> \<open>prime q\<close> \<open>coprime x y\<close> \<open>q dvd x + y\<close> \<open>q dvd S x y\<close>
+  shows \<open>q = p\<close>
+proof (rule ccontr)
+  from \<open>prime p\<close> have \<open>p \<noteq> 0\<close> by auto
+  assume \<open>q \<noteq> p\<close>
+  from \<open>q dvd x + y\<close> have \<open>[y = - x] (mod q)\<close>
+    by (metis add_minus_cancel cong_iff_dvd_diff uminus_add_conv_diff)
+  from Sophie_Germain_lemma_computation_cong_simp[OF \<open>p \<noteq> 0\<close> this]
+  have \<open>[S x y = p * x ^ (p - 1)] (mod q)\<close> unfolding S_def .
+  with \<open>q dvd S x y\<close> \<open>q \<noteq> p\<close> \<open>prime q\<close> \<open>prime p\<close> have \<open>q dvd x ^ (p - 1)\<close>
+    by (metis cong_dvd_iff prime_dvd_mult_iff prime_nat_int_transfer primes_dvd_imp_eq)
+  with \<open>prime q\<close> prime_dvd_power_int prime_nat_int_transfer have \<open>q dvd x\<close> by blast
+  with \<open>q dvd x + y\<close> \<open>[y = - x] (mod q)\<close> have \<open>q dvd y\<close> by (simp add: cong_dvd_iff)
+  with \<open>coprime x y\<close> \<open>q dvd x\<close> \<open>prime q\<close> show False
+    by (metis coprime_def not_prime_unit)
+qed
+
+lemma Sophie_Germain_lemma :
+  fixes x y z :: int
+  assumes \<open>odd p\<close> and \<open>prime p\<close> and fermat : \<open>x ^ p + y ^ p + z ^ p = 0\<close>
+    and \<open>[x \<noteq> 0] (mod p)\<close> and \<open>coprime y z\<close>
+  defines \<open>S \<equiv> \<Sum>k = 0..p - 1. (- z) ^ (p - 1 - k) * y ^ k\<close>
+  shows \<open>\<exists>a \<alpha>. y + z = a ^ p \<and> S = \<alpha> ^ p\<close>
+proof -
+  from Sophie_Germain_lemma_computation[OF \<open>odd p\<close>]
+  have \<open>(y + z) * S = y ^ p + z ^ p\<close> unfolding S_def .
+  also from fermat have \<open>\<dots> = (- x) ^ p\<close> by (simp add: \<open>odd p\<close>)
+  finally have \<open>(y + z) * S = \<dots>\<close> .
+  
+  have \<open>coprime (y + z) S\<close>
+  proof (rule ccontr)
+    assume \<open>\<not> coprime (y + z) S\<close>
+    then consider \<open>y + z = 0\<close> | \<open>S = 0\<close> | q :: nat where \<open>prime q\<close> \<open>q dvd y + z\<close> \<open>q dvd S\<close>
+      by (elim not_coprime_nonzeroE)
+        (use \<open>(y + z) * S = (- x) ^ p\<close> \<open>[x \<noteq> 0] (mod p)\<close> in force,
+          metis nat_0_le prime_int_nat_transfer)
+    hence \<open>p dvd (y + z) * S\<close>
+    proof cases
+      fix q :: nat assume \<open>prime q\<close> \<open>q dvd y + z\<close> \<open>q dvd S\<close>
+      from Sophie_Germain_lemma_only_possible_prime_common_divisor
+        [OF \<open>prime p\<close> _ \<open>coprime y z\<close> \<open>q dvd y + z\<close> \<open>q dvd S\<close>[unfolded S_def]]
+      show \<open>p dvd (y + z) * S\<close> using \<open>int q dvd S\<close> \<open>prime q\<close> by auto
+    qed simp_all
+    with \<open>(y + z) * S = (- x) ^ p\<close> \<open>[x \<noteq> 0] (mod p)\<close> show False
+      by (metis \<open>prime p\<close> cong_0_iff dvd_minus_iff prime_dvd_power_int prime_nat_int_transfer)
+  qed
+
+  from prod_is_some_powerE[OF coprime_commute[THEN iffD1, OF \<open>coprime (y + z) S\<close>]]
+  obtain \<alpha> where \<open>normalize S = \<alpha> ^ p\<close>
+    by (metis (no_types, lifting) \<open>(y + z) * S = (- x) ^ p\<close> mult.commute)
+  moreover from prod_is_some_powerE[OF \<open>coprime (y + z) S\<close> \<open>(y + z) * S = (- x) ^ p\<close>]
+  obtain a where \<open>normalize (y + z) = a ^ p\<close> by blast
+  ultimately have \<open>S = (if 0 \<le> S then \<alpha> ^ p else (- \<alpha>) ^ p)\<close>
+    \<open>y + z = (if 0 \<le> y + z then a ^ p else (- a) ^ p)\<close>
+    by (metis \<open>odd p\<close> abs_of_nonneg abs_of_nonpos
+        add.inverse_inverse linorder_linear normalize_int_def power_minus_odd)+
+  thus \<open>\<exists>a \<alpha>. y + z = a ^ p \<and> S = \<alpha> ^ p\<close> by meson
+qed
+
+
+
+subsection \<open>The Theorem\<close>
+
+theorem Sophie_Germain_theorem :
+  \<open>\<nexists>x y z :: int. x ^ p + y ^ p = z ^ p \<and> [x \<noteq> 0] (mod p) \<and>
+                  [y \<noteq> 0] (mod p) \<and> [z \<noteq> 0] (mod p)\<close> if SG : \<open>SG p\<close>
+proof (rule ccontr) \<comment> \<open>The proof is done by contradiction.\<close>
+  from SophGer_primeD(1)[OF \<open>SG p\<close>] have odd_p : \<open>odd p\<close> .
+  from SG_simps.pos[OF \<open>SG p\<close>] have pos_p : \<open>0 < p\<close> .
+
+  assume \<open>\<not> (\<nexists>x y z. x ^ p + y ^ p = z ^ p \<and> [x \<noteq> 0] (mod int p) \<and>
+                     [y \<noteq> 0] (mod int p) \<and> [z \<noteq> 0] (mod int p))\<close>
+  then obtain x y z :: int
+    where fermat : \<open>x ^ p + y ^ p = z ^ p\<close>
+      and not_cong_0 : \<open>[x \<noteq> 0] (mod p)\<close> \<open>[y \<noteq> 0] (mod p)\<close> \<open>[z \<noteq> 0] (mod p)\<close> by blast
+
+\<comment> \<open>We first assume w.l.o.g. that term\<open>x\<close>, term\<open>y\<close> and term\<open>z\<close> are setwise const\<open>coprime\<close>.\<close>
+  let ?gcd = \<open>Gcd {x, y, z}\<close>
+  wlog coprime : \<open>?gcd = 1\<close> goal False generalizing x y z keeping fermat not_cong_0
+    using FLT_setwise_coprime_reduction_mod_version[OF fermat not_cong_0]
+      hypothesis by blast
+
+\<comment> \<open>Then we can deduce that term\<open>x\<close>, term\<open>y\<close> and term\<open>z\<close> are pairwise const\<open>coprime\<close>.\<close>
+  have coprime_x_y : \<open>coprime x y\<close>
+    by (fact FLT_setwise_coprime_imp_pairwise_coprime
+        [OF SG_simps.nonzero[OF \<open>SG p\<close>] fermat coprime])
+  have coprime_y_z : \<open>coprime y z\<close>
+  proof (subst coprime_minus_right_iff[symmetric],
+      rule FLT_setwise_coprime_imp_pairwise_coprime[OF SG_simps.nonzero[OF \<open>SG p\<close>]])
+    from fermat \<open>odd p\<close> show \<open>y ^ p + (- z) ^ p = (- x) ^ p\<close> by simp
+  next
+    show \<open>Gcd {y, - z, - x} = 1\<close>
+      by (metis Gcd_insert coprime gcd_neg1_int insert_commute)
+  qed
+  have coprime_x_z : \<open>coprime x z\<close>
+  proof (subst coprime_minus_right_iff[symmetric],
+      rule FLT_setwise_coprime_imp_pairwise_coprime[OF SG_simps.nonzero[OF \<open>SG p\<close>]])
+    from fermat \<open>odd p\<close> show \<open>x ^ p + (- z) ^ p = (- y) ^ p\<close> by simp
+  next
+    show \<open>Gcd {x, - z, - y} = 1\<close> 
+      by (metis Gcd_insert coprime gcd_neg1_int insert_commute)
+  qed
+
+  let ?q = \<open>2 * p + 1\<close>
+    \<comment> \<open>From @{thm [show_question_marks = false] SG_simps.p_th_power_mod_q} we have that among term\<open>x\<close>, term\<open>y\<close>
+      and term\<open>z\<close>, one (and only one, see below) is a multiple of term\<open>?q\<close>.\<close>
+  have q_dvd_xyz : \<open>?q dvd x \<or> ?q dvd y \<or> ?q dvd z\<close>
+  proof (rule ccontr)
+    have cong_add_here : \<open>[x ^ p = n1] (mod ?q) \<Longrightarrow> [y ^ p = n2] (mod ?q) \<Longrightarrow> [z ^ p = n3] (mod ?q) \<Longrightarrow>
+                          [x ^ p + y ^ p + (- z) ^ p = n1 + n2 - n3] (mod ?q)\<close> for n1 n2 n3
+      by (simp add: cong_add cong_diff \<open>odd p\<close>)
+    assume \<open>\<not> (?q dvd x \<or> ?q dvd y \<or> ?q dvd z)\<close>
+    hence \<open>\<not> ?q dvd x\<close> \<open>\<not> ?q dvd y\<close> \<open>\<not> ?q dvd z\<close> by simp_all
+    from this[THEN SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]] cong_add_here \<open>odd p\<close>
+    have \<open>  [x ^ p + y ^ p + (- z) ^ p = - 3] (mod ?q) \<or> [x ^ p + y ^ p + (- z) ^ p = - 1] (mod ?q)
+          \<or> [x ^ p + y ^ p + (- z) ^ p = 1] (mod ?q) \<or> [x ^ p + y ^ p + (- z) ^ p = 3] (mod ?q)\<close> (is ?disj_congs)
+      by (elim disjE) fastforce+ (* eight cases *)
+    moreover from fermat \<open>odd p\<close> have \<open>[x ^ p + y ^ p + (- z) ^ p = 0] (mod ?q)\<close> by simp
+    ultimately show False by (metis cong_def SG_simps.notcong_zero[OF \<open>SG p\<close>])
+  qed
+
+\<comment> \<open>Without loss of generality, we can assume that term\<open>x\<close> is a multiple of term\<open>?q\<close>.\<close>
+  wlog \<open>?q dvd x\<close> goal False generalizing x y z
+    keeping fermat not_cong_0 coprime_x_y coprime_y_z coprime_x_z q_dvd_xyz
+  proof -
+    from negation q_dvd_xyz have \<open>?q dvd y \<or> ?q dvd z\<close> by simp
+    thus False
+    proof (elim disjE)
+      assume \<open>?q dvd y\<close>
+      thus False
+      proof (rule hypothesis[OF _ _ not_cong_0(2, 1, 3)])
+        from fermat show \<open>y ^ p + x ^ p = z ^ p\<close> by linarith
+      next
+        show \<open>coprime y x\<close> \<open>coprime x z\<close> \<open>coprime y z\<close>
+          by (simp_all add: coprime_commute coprime_x_y coprime_x_z coprime_y_z)
+      next
+        from q_dvd_xyz show \<open>?q dvd y \<or> ?q dvd x \<or> ?q dvd z\<close> by linarith
+      qed
+    next
+      assume \<open>?q dvd z\<close>
+      hence \<open>?q dvd - z\<close> by simp
+      thus False
+      proof (rule hypothesis)
+        from fermat \<open>odd p\<close> show \<open>(- z) ^ p + x ^ p = (- y) ^ p\<close> by simp
+      next
+        from \<open>[x \<noteq> 0] (mod p)\<close> \<open>[y \<noteq> 0] (mod p)\<close> \<open>[z \<noteq> 0] (mod p)\<close>
+        show \<open>[x \<noteq> 0] (mod p)\<close> \<open>[- y \<noteq> 0] (mod p)\<close> \<open>[- z \<noteq> 0] (mod p)\<close>
+          by (simp_all add: cong_0_iff)
+      next
+        show \<open>coprime (- z) x\<close> \<open>coprime x (- y)\<close> \<open>coprime (- z) (- y)\<close>
+          by (simp_all add: coprime_commute coprime_x_y coprime_x_z coprime_y_z)
+      next
+        from q_dvd_xyz show \<open>?q dvd - z \<or> ?q dvd x \<or> ?q dvd - y\<close> by auto
+      qed
+    qed
+  qed
+
+\<comment> \<open>Now we can use the lemma above.\<close>
+  let ?S = \<open>\<lambda>y z. \<Sum>k = 0..p - 1. (- z) ^ (p - 1 - k) * y ^ k\<close>
+  from fermat \<open>odd p\<close> have \<open>y ^ p + x ^ p + (- z) ^ p = 0\<close>
+    \<open>x ^ p + y ^ p + (- z) ^ p = 0\<close> \<open>(-z ) ^ p + x ^ p + y ^ p = 0\<close> by simp_all
+  from Sophie_Germain_lemma[OF SophGer_primeD(1-2)[OF \<open>SG p\<close>]
+      \<open>x ^ p + y ^ p + (- z) ^ p = 0\<close> \<open>[x \<noteq> 0] (mod p)\<close>]
+  obtain a \<alpha> where a_prop : \<open>y + (- z) = a ^ p\<close>
+    and \<alpha>_prop : \<open>?S y (-z) = \<alpha> ^ p\<close>
+    using coprime_minus_right_iff coprime_y_z by blast
+
+  from Sophie_Germain_lemma[OF SophGer_primeD(1-2)[OF \<open>SG p\<close>]
+      \<open>y ^ p + x ^ p + (- z) ^ p = 0\<close> \<open>[y \<noteq> 0] (mod p)\<close>]
+  obtain b where b_prop : \<open>x + - z = b ^ p\<close>
+    by (metis coprime_minus_right_iff coprime_x_z)
+  
+  from Sophie_Germain_lemma[OF SophGer_primeD(1-2)[OF \<open>SG p\<close>]
+      \<open>(-z ) ^ p + x ^ p + y ^ p = 0\<close>] coprime_x_z \<open>[z \<noteq> 0] (mod p)\<close> 
+  obtain c where c_prop : \<open>x + y = c ^ p\<close>
+    by (meson cong_0_iff coprime_x_y dvd_minus_iff)
+
+
+  from \<open>?q dvd x\<close> have \<open>\<not> ?q dvd y\<close> and \<open>\<not> ?q dvd z\<close>
+    using coprime_x_y coprime_x_z not_coprimeI not_prime_unit prime_nat_int_transfer
+    by (metis SophGer_primeD(3)[OF \<open>SG p\<close>] prime_nat_int_transfer)+
+
+  from b_prop \<open>?q dvd x\<close> have \<open>[b ^ p = - z] (mod ?q)\<close>
+    by (metis add_diff_cancel_right' cong_iff_dvd_diff)
+  with \<open>\<not> ?q dvd z\<close> cong_dvd_iff dvd_minus_iff have \<open>\<not> ?q dvd b ^ p\<close> by blast
+  with \<open>0 < p\<close> have \<open>\<not> ?q dvd b\<close> by (meson dvd_power dvd_trans)
+  with SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]
+  have cong1 : \<open>[b ^ p = 1] (mod ?q) \<or> [b ^ p = - 1] (mod ?q)\<close> by blast
+
+  from c_prop \<open>?q dvd x\<close> have \<open>[c ^ p = y] (mod ?q)\<close>
+    by (metis add_diff_cancel_right' cong_iff_dvd_diff)
+  with \<open>\<not> ?q dvd y\<close> cong_dvd_iff have \<open>\<not> ?q dvd c ^ p\<close> by blast
+  with \<open>0 < p\<close> have \<open>\<not> ?q dvd c\<close> by (meson dvd_power dvd_trans)
+  with SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]
+  have cong2 : \<open>[c ^ p = 1] (mod ?q) \<or> [c ^ p = - 1] (mod ?q)\<close> by blast
+
+
+  have \<open>?q dvd a\<close>
+  proof (rule ccontr)
+    have cong_add_here : \<open>[b ^ p = n1] (mod ?q) \<Longrightarrow> [c ^ p = n2] (mod ?q) \<Longrightarrow> [a ^ p = n3] (mod ?q) \<Longrightarrow>
+                          [b ^ p + c ^ p - a ^ p = n1 + n2 - n3] (mod ?q)\<close> for n1 n2 n3
+      by (intro cong_add cong_diff)
+    assume \<open>\<not> ?q dvd a\<close>
+    with SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]
+    have cong3 : \<open>[a ^ p = 1] (mod ?q) \<or> [a ^ p = - 1] (mod ?q)\<close> by blast
+    from cong1 cong2 cong3 cong_add_here
+    have \<open>  [b ^ p + c ^ p - a ^ p = - 3] (mod ?q) \<or> [b ^ p + c ^ p - a ^ p = - 1] (mod ?q)
+          \<or> [b ^ p + c ^ p - a ^ p = 1] (mod ?q) \<or> [b ^ p + c ^ p - a ^ p = 3] (mod ?q)\<close> (is ?disj_congs)
+      by (elim disjE) fastforce+ (* eight cases *)
+    have \<open>b ^ p + c ^ p - a ^ p = 2 * x\<close> by (fold a_prop b_prop c_prop) simp
+    also from \<open>?q dvd x\<close> cong_0_iff have \<open>[2 * x = 0] (mod ?q)\<close>
+      by (metis dvd_add_right_iff mult_2)
+    finally have \<open>[b ^ p + c ^ p - a ^ p = 0] (mod ?q)\<close> .
+    with \<open>?disj_congs\<close> show False by (metis cong_def SG_simps.notcong_zero[OF \<open>SG p\<close>])
+  qed
+  with oddE \<open>odd p\<close> have \<open>?q dvd a ^ p\<close> by fastforce
+  with a_prop have \<open>[y = z] (mod ?q)\<close> by (simp add: cong_iff_dvd_diff)
+  with cong_sym have \<open>[z = y] (mod ?q)\<close> by blast
+
+
+\<comment> \<open>It is now time to conclude the proof!\<close>
+  have \<open>\<alpha> ^ p = ?S y (-z)\<close> by (fact \<alpha>_prop[symmetric])
+  also from SG_simps.nonzero[OF \<open>SG p\<close>] \<open>[z = y] (mod ?q)\<close> cong_minus_minus_iff
+  have \<open>[?S y (-z) = p * y ^ (p - 1)] (mod ?q)\<close>
+    by (blast intro: Sophie_Germain_lemma_computation_cong_simp)
+  finally have \<open>[\<alpha> ^ p = p * y ^ (p - 1)] (mod ?q)\<close> .
+
+  from SG_simps.p_th_power_mod_q[OF \<open>SG p\<close> \<open>\<not> ?q dvd c ^ p\<close>] \<open>[c ^ p = y] (mod ?q)\<close>
+  have \<open>[y = 1] (mod ?q) \<or> [y = - 1] (mod ?q)\<close> by (metis cong2 cong_def)
+  hence \<open>[y ^ (p - 1) = 1] (mod ?q)\<close>
+    by (elim disjE) (drule cong_pow[where n = \<open>p - 1\<close>], simp add: \<open>odd p\<close>)+
+  from cong_trans[OF \<open>[\<alpha> ^ p = p * y ^ (p - 1)] (mod ?q)\<close> cong_mult[OF cong_refl this]]
+  have \<open>[\<alpha> ^ p = p] (mod ?q)\<close> by simp
+
+\<comment> \<open>But term\<open>\<alpha> ^ p\<close> is congruent to term\<open>- 1 :: int\<close>, term\<open>0 :: int\<close> or term\<open>1 :: int\<close>
+      modulo term\<open>?q\<close>, whereas term\<open>p\<close> can not be; hence the final contradiction.\<close>
+  moreover from SG_simps.p_th_power_mod_q[OF \<open>SG p\<close>]
+  have \<open>[\<alpha> ^ p = - 1] (mod ?q) \<or> [\<alpha> ^ p = 0] (mod ?q) \<or> [\<alpha> ^ p = 1] (mod ?q)\<close>
+    by (metis cong_0_iff cong_pow power_0_left)
+  ultimately show False by (metis SG_simps.notcong_p[OF \<open>SG p\<close>] cong_def)
+qed
+
+
+
+
+(*<*)
+end
+  (*>*)
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/document/root.bib
--- /dev/null
+++ thys/Sophie_Germain/document/root.bib
@@ -0,0 +1,15 @@
+@book{Francinou_Gianella_Nicolas_2014,
+  place={Paris},
+  title={Oraux X-ENS Algèbre 1},
+  publisher={Cassini},
+  author={Francinou, Serge and Gianella, Hervé and Nicolas, Serge},
+  year={2014}
+}
+
+
+@article{kiefer2012theoreme,
+  title={{Le th{\'e}or{\`e}me de Fermat vu par M. Le Blanc}},
+  author={Kiefer, Ann},
+  journal={Brussels Summer School of Mathematics, Notes de la cinqui{\`e}me BSSM},
+  year={2012}
+}
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 thys/Sophie_Germain/document/root.tex
--- /dev/null
+++ thys/Sophie_Germain/document/root.tex
@@ -0,0 +1,62 @@
+\documentclass[11pt,a4paper]{article}
+\usepackage[T1]{fontenc}
+\usepackage{isabelle,isabellesym}
+
+\usepackage{graphicx}
+
+% further packages required for unusual symbols (see also
+% isabellesym.sty), use only when needed
+
+\usepackage{amssymb}
+  %for \<leadsto>, \<box>, \<diamond>, \<sqsupset>, \<mho>, \<Join>,
+  %\<lhd>, \<lesssim>, \<greatersim>, \<lessapprox>, \<greaterapprox>,
+  %\<triangleq>, \<yen>, \<lozenge>
+
+%\usepackage{eurosym}
+  %for \<euro>
+
+%\usepackage[only,bigsqcap,bigparallel,fatsemi,interleave,sslash]{stmaryrd}
+  %for \<Sqinter>, \<Parallel>, \<Zsemi>, \<Parallel>, \<sslash>
+
+%\usepackage{eufrak}
+  %for \<AA> ... \<ZZ>, \<aa> ... \<zz> (also included in amssymb)
+
+%\usepackage{textcomp}
+  %for \<onequarter>, \<onehalf>, \<threequarters>, \<degree>, \<cent>,
+  %\<currency>
+
+% 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{Sophie Germain's Theorem}
+\author{Benoît Ballenghien\\Université Paris-Saclay, CNRS, ENS Paris-Saclay, LMF}
+\maketitle
+
+\tableofcontents
+
+% sane default for proof documents
+\parindent 0pt\parskip 0.5ex
+
+% generated text of all theories
+\input{session}
+
+% optional bibliography
+\bibliographystyle{abbrv}
+\bibliography{root}
+
+\end{document}
+
+%%% Local Variables:
+%%% mode: latex
+%%% TeX-master: t
+%%% End:
diff -r 848f2c00e57f -r 29846bc89de2 thys/Word_Lib/More_Divides.thy
--- thys/Word_Lib/More_Divides.thy
+++ thys/Word_Lib/More_Divides.thy
@@ -277,9 +277,8 @@
   proof (cases "m = 1")
     case False
     with that show ?thesis
-      using assms
-      unfolding mod_div_equality_div_eq
-      by (smt (verit, ccfv_SIG) dvd_eq_mod_eq_0 int_mod_ge' mod_diff_eq pos_mod_bound pos_mod_sign)
+      using assms mod_diff_left_eq[of n m "1"] zmod_minus1[of m]
+      unfolding mod_div_equality_div_eq by simp
   qed auto
   moreover
   have "\<not> m dvd n \<Longrightarrow> (n - 1) div m * m = n div m * m"
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/ballenghien/index.html
--- web/authors/ballenghien/index.html
+++ web/authors/ballenghien/index.html
@@ -86,6 +86,10 @@
   </div>
 </header>
       <div>
+    <a href="https://orcid.org/0009-0000-4941-187X">
+      <img alt="ORCID logo" src="https://info.orcid.org/wp-content/uploads/2019/11/orcid_16x16.png"
+           width="16" height="16" />0009-0000-4941-187X
+    </a>
     <h2>E-Mails 📧</h2>
     <ul>
         <li>
@@ -98,6 +102,81 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">Apr 10</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Sophie_Germain.html">Sophie Germain’s Theorem</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>
+        </div>
+        <span class="date">Apr 09</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Hilbert_Basis.html">A Proof of Hilbert Basis Theorem and            an Extension to Formal Power Series</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/ballenghien/index.xml
--- web/authors/ballenghien/index.xml
+++ web/authors/ballenghien/index.xml
@@ -4,6 +4,57 @@
     <title>Benoît Ballenghien</title>
     <link>https://isa-afp.org/authors/ballenghien/</link>
     <description>AFP entries of Benoît Ballenghien</description><item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>CSP Semantics over Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/HOL-CSP_RS.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>CSP Semantics over Restriction Spaces in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Definition and Elementary Properties of Ultrametric Spaces</title>
+  <link>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</link>
+  <pubDate>Thu, 10 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Definition and Elementary Properties of Ultrametric Spaces in the AFP</description>
+    <category>Mathematics/Topology</category>
+</item>
+<item>
+  <title>Sophie Germain’s Theorem</title>
+  <link>https://isa-afp.org/entries/Sophie_Germain.html</link>
+  <pubDate>Wed, 09 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Sophie Germain’s Theorem in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>A Proof of Hilbert Basis Theorem and            an Extension to Formal Power Series</title>
   <link>https://isa-afp.org/entries/Hilbert_Basis.html</link>
   <pubDate>Wed, 12 Feb 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/bordg/index.html
--- web/authors/bordg/index.html
+++ web/authors/bordg/index.html
@@ -102,6 +102,20 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/eberl/">Manuel Eberl</a>, 
+            <a href="../../authors/bordg/">Anthony Bordg</a>, 
+            <a href="../../authors/li/">Wenda Li</a> and 
+            <a href="../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Dedekind_Sums.html">Dedekind Sums</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/bordg/index.xml
--- web/authors/bordg/index.xml
+++ web/authors/bordg/index.xml
@@ -4,6 +4,13 @@
     <title>Anthony Bordg</title>
     <link>https://isa-afp.org/authors/bordg/</link>
     <description>AFP entries of Anthony Bordg</description><item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>Dedekind Sums</title>
   <link>https://isa-afp.org/entries/Dedekind_Sums.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/chevalier/index.html
--- web/authors/chevalier/index.html
+++ web/authors/chevalier/index.html
@@ -96,7 +96,7 @@
           </h5>
           <br>
           by
-            <a href="../../authors/danilkin/">Anton Danilkin</a> and 
+            <a href="../../authors/danilkin/">Anna Danilkin</a> and 
             <a href="../../authors/chevalier/">Loïc Chevalier</a>
         </div>
         <span class="date">May 03</span>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/danilkin/index.html
--- web/authors/danilkin/index.html
+++ web/authors/danilkin/index.html
@@ -3,11 +3,11 @@
   <meta charset="utf-8" />
   <meta http-equiv="X-UA-Compatible" content="IE=edge" />
   <meta name="viewport" content="width=device-width, initial-scale=1" />
-    <title>Anton Danilkin - Archive of Formal Proofs</title>
-    <meta name="description" content="Anton Danilkin in the Archive of Formal Proofs" />
-    <meta property="og:description" content="Anton Danilkin in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/danilkin/index.xml" title="Archive of Formal Proofs" />
+    <title>Anna Danilkin - Archive of Formal Proofs</title>
+    <meta name="description" content="Anna Danilkin in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Anna Danilkin in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/danilkin/index.xml" title="Archive of Formal Proofs" />
 
-  <meta property="og:title" content="Anton Danilkin" />
+  <meta property="og:title" content="Anna Danilkin" />
   <meta property="og:url" content="https://isa-afp.org/authors/danilkin/" />
   <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
     <meta property="og:type" content="profile" />
@@ -80,7 +80,7 @@
         </datalist>
       </div>
     </form>
-  <h1 ><span class='first'>A</span>nton <span class='first'>D</span>anilkin
+  <h1 ><span class='first'>A</span>nna <span class='first'>D</span>anilkin
   </h1>
   <div>
   </div>
@@ -96,7 +96,7 @@
           </h5>
           <br>
           by
-            <a href="../../authors/danilkin/">Anton Danilkin</a> and 
+            <a href="../../authors/danilkin/">Anna Danilkin</a> and 
             <a href="../../authors/chevalier/">Loïc Chevalier</a>
         </div>
         <span class="date">May 03</span>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/danilkin/index.xml
--- web/authors/danilkin/index.xml
+++ web/authors/danilkin/index.xml
@@ -1,9 +1,9 @@
 <?xml version="1.0" encoding="utf-8" standalone="yes"?>
 <rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
   <channel>
-    <title>Anton Danilkin</title>
+    <title>Anna Danilkin</title>
     <link>https://isa-afp.org/authors/danilkin/</link>
-    <description>AFP entries of Anton Danilkin</description><item>
+    <description>AFP entries of Anna Danilkin</description><item>
   <title>Three Squares Theorem</title>
   <link>https://isa-afp.org/entries/Three_Squares.html</link>
   <pubDate>Wed, 03 May 2023 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/eberl/index.html
--- web/authors/eberl/index.html
+++ web/authors/eberl/index.html
@@ -106,6 +106,20 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/eberl/">Manuel Eberl</a>, 
+            <a href="../../authors/bordg/">Anthony Bordg</a>, 
+            <a href="../../authors/li/">Wenda Li</a> and 
+            <a href="../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Pentagonal_Number_Theorem.html">The Partition Function and the Pentagonal Number Theorem</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/eberl/index.xml
--- web/authors/eberl/index.xml
+++ web/authors/eberl/index.xml
@@ -4,6 +4,13 @@
     <title>Manuel Eberl</title>
     <link>https://isa-afp.org/authors/eberl/</link>
     <description>AFP entries of Manuel Eberl</description><item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>The Partition Function and the Pentagonal Number Theorem</title>
   <link>https://isa-afp.org/entries/Pentagonal_Number_Theorem.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/fosters/index.html
--- web/authors/fosters/index.html
+++ web/authors/fosters/index.html
@@ -102,6 +102,17 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Shallow_Expressions.html">Shallow Expressions</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/fosters/">Simon Foster</a>
+        </div>
+        <span class="date">May 16</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Kolmogorov_Chentsov.html">The Kolmogorov-Chentsov theorem</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/fosters/index.xml
--- web/authors/fosters/index.xml
+++ web/authors/fosters/index.xml
@@ -4,6 +4,14 @@
     <title>Simon Foster</title>
     <link>https://isa-afp.org/authors/fosters/</link>
     <description>AFP entries of Simon Foster</description><item>
+  <title>Shallow Expressions</title>
+  <link>https://isa-afp.org/entries/Shallow_Expressions.html</link>
+  <pubDate>Fri, 16 May 2025 00:00:00 +0000</pubDate>
+  <description>Shallow Expressions in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Programming languages/Logics</category>
+</item>
+<item>
   <title>The Kolmogorov-Chentsov theorem</title>
   <link>https://isa-afp.org/entries/Kolmogorov_Chentsov.html</link>
   <pubDate>Sun, 06 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/gschossmann/index.html
--- /dev/null
+++ web/authors/gschossmann/index.html
@@ -0,0 +1,123 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Markus Gschoßmann - Archive of Formal Proofs</title>
+    <meta name="description" content="Markus Gschoßmann in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Markus Gschoßmann in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/gschossmann/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Markus Gschoßmann" />
+  <meta property="og:url" content="https://isa-afp.org/authors/gschossmann/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>M</span>arkus <span class='first'>G</span>schoßmann
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+    <a href="https://orcid.org/0009-0005-5058-5806">
+      <img alt="ORCID logo" src="https://info.orcid.org/wp-content/uploads/2019/11/orcid_16x16.png"
+           width="16" height="16" />0009-0005-5058-5806
+    </a>
+    <h2>E-Mails 📧</h2>
+    <ul>
+        <li>
+          <a class="obfuscated" data="eyJ1c2VyIjpbIm1hcmt1cyIsImdzY2hvc3NtYW5uIl0sImhvc3QiOlsidHVtIiwiZGUiXX0="><span class="rev">ed</span><span class="rev">.</span><span class="rev">mut</span>@<span class="rev">nnamssohcsg</span><span class="rev">.</span><span class="rev">sukram</span></a>
+        </li>
+    </ul>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/gschossmann/index.xml
--- /dev/null
+++ web/authors/gschossmann/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>Markus Gschoßmann</title>
+    <link>https://isa-afp.org/authors/gschossmann/</link>
+    <description>AFP entries of Markus Gschoßmann</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/index.html
--- web/authors/index.html
+++ web/authors/index.html
@@ -214,10 +214,10 @@
         <a href="../authors/stueber/">Anke Stüber</a>
       </li>
       <li>
-        <a href="../authors/bordg/">Anthony Bordg</a>
+        <a href="../authors/danilkin/">Anna Danilkin</a>
       </li>
       <li>
-        <a href="../authors/danilkin/">Anton Danilkin</a>
+        <a href="../authors/bordg/">Anthony Bordg</a>
       </li>
       <li>
         <a href="../authors/heller/">Armin Heller</a>
@@ -232,6 +232,9 @@
         <a href="../authors/keskin/">Ata Keskin</a>
       </li>
       <li>
+        <a href="../authors/stimpfle/">August Martin Stimpfle</a>
+      </li>
+      <li>
         <a href="../authors/bergstroem/">Axel Bergström</a>
       </li>
       <li>
@@ -310,6 +313,9 @@
         <a href="../authors/langenstein/">Bruno Langenstein</a>
       </li>
       <li>
+        <a href="../authors/philipp/">Bruno Philipp</a>
+      </li>
+      <li>
         <a href="../authors/paleo/">Bruno Woltzenlogel Paleo</a>
       </li>
       <li>
@@ -535,6 +541,9 @@
         <a href="../authors/immler/">Fabian Immler</a>
       </li>
       <li>
+        <a href="../authors/lehr/">Fabian Lehr</a>
+      </li>
+      <li>
         <a href="../authors/mitterwallner/">Fabian Mitterwallner</a>
       </li>
       <li>
@@ -547,6 +556,9 @@
         <a href="../authors/brandt/">Felix Brandt</a>
       </li>
       <li>
+        <a href="../authors/krayer/">Felix Krayer</a>
+      </li>
+      <li>
         <a href="../authors/maricf/">Filip Marić</a>
       </li>
       <li>
@@ -895,6 +907,9 @@
         <a href="../authors/liu/">Junyi Liu</a>
       </li>
       <li>
+        <a href="../authors/taskin/">Kaan Taskin</a>
+      </li>
+      <li>
         <a href="../authors/engelhardt/">Kai Engelhardt</a>
       </li>
       <li>
@@ -1036,6 +1051,9 @@
         <a href="../authors/wassell/">Mark Wassell</a>
       </li>
       <li>
+        <a href="../authors/gschossmann/">Markus Gschoßmann</a>
+      </li>
+      <li>
         <a href="../authors/olm/">Markus Müller-Olm</a>
       </li>
       <li>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/index.json
--- web/authors/index.json
+++ web/authors/index.json
@@ -168,14 +168,14 @@
   "name": "Anke Stüber"
  },
  {
+  "link": "/authors/danilkin/",
+  "name": "Anna Danilkin"
+ },
+ {
   "link": "/authors/bordg/",
   "name": "Anthony Bordg"
  },
  {
-  "link": "/authors/danilkin/",
-  "name": "Anton Danilkin"
- },
- {
   "link": "/authors/heller/",
   "name": "Armin Heller"
  },
@@ -192,6 +192,10 @@
   "name": "Ata Keskin"
  },
  {
+  "link": "/authors/stimpfle/",
+  "name": "August Martin Stimpfle"
+ },
+ {
   "link": "/authors/bergstroem/",
   "name": "Axel Bergström"
  },
@@ -296,6 +300,10 @@
   "name": "Bruno Langenstein"
  },
  {
+  "link": "/authors/philipp/",
+  "name": "Bruno Philipp"
+ },
+ {
   "link": "/authors/paleo/",
   "name": "Bruno Woltzenlogel Paleo"
  },
@@ -596,6 +604,10 @@
   "name": "Fabian Immler"
  },
  {
+  "link": "/authors/lehr/",
+  "name": "Fabian Lehr"
+ },
+ {
   "link": "/authors/mitterwallner/",
   "name": "Fabian Mitterwallner"
  },
@@ -612,6 +624,10 @@
   "name": "Felix Brandt"
  },
  {
+  "link": "/authors/krayer/",
+  "name": "Felix Krayer"
+ },
+ {
   "link": "/authors/maricf/",
   "name": "Filip Marić"
  },
@@ -1076,6 +1092,10 @@
   "name": "Junyi Liu"
  },
  {
+  "link": "/authors/taskin/",
+  "name": "Kaan Taskin"
+ },
+ {
   "link": "/authors/engelhardt/",
   "name": "Kai Engelhardt"
  },
@@ -1264,6 +1284,10 @@
   "name": "Mark Wassell"
  },
  {
+  "link": "/authors/gschossmann/",
+  "name": "Markus Gschoßmann"
+ },
+ {
   "link": "/authors/olm/",
   "name": "Markus Müller-Olm"
  },
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/krayer/index.html
--- /dev/null
+++ web/authors/krayer/index.html
@@ -0,0 +1,119 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Felix Krayer - Archive of Formal Proofs</title>
+    <meta name="description" content="Felix Krayer in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Felix Krayer in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/krayer/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Felix Krayer" />
+  <meta property="og:url" content="https://isa-afp.org/authors/krayer/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>F</span>elix <span class='first'>K</span>rayer
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+    <h2>E-Mails 📧</h2>
+    <ul>
+        <li>
+          <a class="obfuscated" data="eyJ1c2VyIjpbImZlbGl4Iiwia3JheWVyIl0sImhvc3QiOlsidHVtIiwiZGUiXX0="><span class="rev">ed</span><span class="rev">.</span><span class="rev">mut</span>@<span class="rev">reyark</span><span class="rev">.</span><span class="rev">xilef</span></a>
+        </li>
+    </ul>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/krayer/index.xml
--- /dev/null
+++ web/authors/krayer/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>Felix Krayer</title>
+    <link>https://isa-afp.org/authors/krayer/</link>
+    <description>AFP entries of Felix Krayer</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/lehr/index.html
--- /dev/null
+++ web/authors/lehr/index.html
@@ -0,0 +1,119 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Fabian Lehr - Archive of Formal Proofs</title>
+    <meta name="description" content="Fabian Lehr in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Fabian Lehr in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/lehr/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Fabian Lehr" />
+  <meta property="og:url" content="https://isa-afp.org/authors/lehr/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>F</span>abian <span class='first'>L</span>ehr
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+    <h2>E-Mails 📧</h2>
+    <ul>
+        <li>
+          <a class="obfuscated" data="eyJ1c2VyIjpbImZhYmlhbiIsImxlaHIiXSwiaG9zdCI6WyJ0dW0iLCJkZSJdfQ=="><span class="rev">ed</span><span class="rev">.</span><span class="rev">mut</span>@<span class="rev">rhel</span><span class="rev">.</span><span class="rev">naibaf</span></a>
+        </li>
+    </ul>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/lehr/index.xml
--- /dev/null
+++ web/authors/lehr/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>Fabian Lehr</title>
+    <link>https://isa-afp.org/authors/lehr/</link>
+    <description>AFP entries of Fabian Lehr</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/li/index.html
--- web/authors/li/index.html
+++ web/authors/li/index.html
@@ -102,6 +102,20 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/eberl/">Manuel Eberl</a>, 
+            <a href="../../authors/bordg/">Anthony Bordg</a>, 
+            <a href="../../authors/li/">Wenda Li</a> and 
+            <a href="../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Dedekind_Sums.html">Dedekind Sums</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/li/index.xml
--- web/authors/li/index.xml
+++ web/authors/li/index.xml
@@ -4,6 +4,13 @@
     <title>Wenda Li</title>
     <link>https://isa-afp.org/authors/li/</link>
     <description>AFP entries of Wenda Li</description><item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>Dedekind Sums</title>
   <link>https://isa-afp.org/entries/Dedekind_Sums.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/nipkow/index.html
--- web/authors/nipkow/index.html
+++ web/authors/nipkow/index.html
@@ -106,6 +106,24 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/List_Power.html">Power Operator for Lists</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/nipkow/index.xml
--- web/authors/nipkow/index.xml
+++ web/authors/nipkow/index.xml
@@ -4,6 +4,13 @@
     <title>Tobias Nipkow</title>
     <link>https://isa-afp.org/authors/nipkow/</link>
     <description>AFP entries of Tobias Nipkow</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+<item>
   <title>Power Operator for Lists</title>
   <link>https://isa-afp.org/entries/List_Power.html</link>
   <pubDate>Wed, 29 Jan 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/paulson/index.html
--- web/authors/paulson/index.html
+++ web/authors/paulson/index.html
@@ -86,6 +86,10 @@
   </div>
 </header>
       <div>
+    <a href="https://orcid.org/0000-0003-0288-4279">
+      <img alt="ORCID logo" src="https://info.orcid.org/wp-content/uploads/2019/11/orcid_16x16.png"
+           width="16" height="16" />0000-0003-0288-4279
+    </a>
     <h2>Homepages 🌐</h2>
     <ul>
         <li><a href="https://www.cl.cam.ac.uk/~lp15/">https://www.cl.cam.ac.uk/~lp15/</a></li>
@@ -102,6 +106,31 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/eberl/">Manuel Eberl</a>, 
+            <a href="../../authors/bordg/">Anthony Bordg</a>, 
+            <a href="../../authors/li/">Wenda Li</a> and 
+            <a href="../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Farey_Sequences.html">Farey Sequences and Ford Circles</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 15</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Dedekind_Sums.html">Dedekind Sums</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/paulson/index.xml
--- web/authors/paulson/index.xml
+++ web/authors/paulson/index.xml
@@ -4,6 +4,20 @@
     <title>Lawrence C. Paulson</title>
     <link>https://isa-afp.org/authors/paulson/</link>
     <description>AFP entries of Lawrence C. Paulson</description><item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
+  <title>Farey Sequences and Ford Circles</title>
+  <link>https://isa-afp.org/entries/Farey_Sequences.html</link>
+  <pubDate>Thu, 15 May 2025 00:00:00 +0000</pubDate>
+  <description>Farey Sequences and Ford Circles in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>Dedekind Sums</title>
   <link>https://isa-afp.org/entries/Dedekind_Sums.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/philipp/index.html
--- /dev/null
+++ web/authors/philipp/index.html
@@ -0,0 +1,119 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Bruno Philipp - Archive of Formal Proofs</title>
+    <meta name="description" content="Bruno Philipp in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Bruno Philipp in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/philipp/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Bruno Philipp" />
+  <meta property="og:url" content="https://isa-afp.org/authors/philipp/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>B</span>runo <span class='first'>P</span>hilipp
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+    <h2>E-Mails 📧</h2>
+    <ul>
+        <li>
+          <a class="obfuscated" data="eyJ1c2VyIjpbImdlOTVmZWsiXSwiaG9zdCI6WyJteXR1bSIsImRlIl19"><span class="rev">ed</span><span class="rev">.</span><span class="rev">mutym</span>@<span class="rev">kef59eg</span></a>
+        </li>
+    </ul>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/philipp/index.xml
--- /dev/null
+++ web/authors/philipp/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>Bruno Philipp</title>
+    <link>https://isa-afp.org/authors/philipp/</link>
+    <description>AFP entries of Bruno Philipp</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/puyobro/index.html
--- web/authors/puyobro/index.html
+++ web/authors/puyobro/index.html
@@ -86,6 +86,10 @@
   </div>
 </header>
       <div>
+    <a href="https://orcid.org/0009-0006-0294-9043">
+      <img alt="ORCID logo" src="https://info.orcid.org/wp-content/uploads/2019/11/orcid_16x16.png"
+           width="16" height="16" />0009-0006-0294-9043
+    </a>
     <h2>E-Mails 📧</h2>
     <ul>
         <li>
@@ -98,6 +102,58 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">Apr 10</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Morley_Theorem.html">Morley's Theorem</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/puyobro/index.xml
--- web/authors/puyobro/index.xml
+++ web/authors/puyobro/index.xml
@@ -4,6 +4,42 @@
     <title>Benjamin Puyobro</title>
     <link>https://isa-afp.org/authors/puyobro/</link>
     <description>AFP entries of Benjamin Puyobro</description><item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Definition and Elementary Properties of Ultrametric Spaces</title>
+  <link>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</link>
+  <pubDate>Thu, 10 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Definition and Elementary Properties of Ultrametric Spaces in the AFP</description>
+    <category>Mathematics/Topology</category>
+</item>
+<item>
   <title>Morley&#39;s Theorem</title>
   <link>https://isa-afp.org/entries/Morley_Theorem.html</link>
   <pubDate>Fri, 07 Mar 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/stimpfle/index.html
--- /dev/null
+++ web/authors/stimpfle/index.html
@@ -0,0 +1,113 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>August Martin Stimpfle - Archive of Formal Proofs</title>
+    <meta name="description" content="August Martin Stimpfle in the Archive of Formal Proofs" />
+    <meta property="og:description" content="August Martin Stimpfle in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/stimpfle/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="August Martin Stimpfle" />
+  <meta property="og:url" content="https://isa-afp.org/authors/stimpfle/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>A</span>ugust <span class='first'>M</span>artin <span class='first'>S</span>timpfle
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/stimpfle/index.xml
--- /dev/null
+++ web/authors/stimpfle/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>August Martin Stimpfle</title>
+    <link>https://isa-afp.org/authors/stimpfle/</link>
+    <description>AFP entries of August Martin Stimpfle</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/taskin/index.html
--- /dev/null
+++ web/authors/taskin/index.html
@@ -0,0 +1,119 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Kaan Taskin - Archive of Formal Proofs</title>
+    <meta name="description" content="Kaan Taskin in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Kaan Taskin in the Archive of Formal Proofs" /><link rel="alternate" type="application/rss+xml" href="https://isa-afp.org/authors/taskin/index.xml" title="Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Kaan Taskin" />
+  <meta property="og:url" content="https://isa-afp.org/authors/taskin/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="profile" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../../">
+            <li >Home</li>
+          </a>
+          <a href="../../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../../download/">
+            <li >Download</li>
+          </a>
+          <a href="../../help/">
+            <li >Help</li>
+          </a>
+          <a href="../../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>K</span>aan <span class='first'>T</span>askin
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+    <h2>E-Mails 📧</h2>
+    <ul>
+        <li>
+          <a class="obfuscated" data="eyJ1c2VyIjpbImthYW50YXNraW4yMDIyIl0sImhvc3QiOlsiZ21haWwiLCJjb20iXX0="><span class="rev">moc</span><span class="rev">.</span><span class="rev">liamg</span>@<span class="rev">2202niksatnaak</span></a>
+        </li>
+    </ul>
+
+  <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/taskin/index.xml
--- /dev/null
+++ web/authors/taskin/index.xml
@@ -0,0 +1,15 @@
+<?xml version="1.0" encoding="utf-8" standalone="yes"?>
+<rss version="2.0" xmlns:atom="http://www.w3.org/2005/Atom">
+  <channel>
+    <title>Kaan Taskin</title>
+    <link>https://isa-afp.org/authors/taskin/</link>
+    <description>AFP entries of Kaan Taskin</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+
+  </channel>
+</rss>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/wimmer/index.html
--- web/authors/wimmer/index.html
+++ web/authors/wimmer/index.html
@@ -98,6 +98,18 @@
     </ul>
 
   <h2>Entries</h2>
+    <h3 class="head">2025</h3>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/wimmer/">Simon Wimmer</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
     <h3 class="head">2024</h3>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/wimmer/index.xml
--- web/authors/wimmer/index.xml
+++ web/authors/wimmer/index.xml
@@ -4,6 +4,14 @@
     <title>Simon Wimmer</title>
     <link>https://isa-afp.org/authors/wimmer/</link>
     <description>AFP entries of Simon Wimmer</description><item>
+  <title>Munta: A Verified Model Checker for Timed Automata</title>
+  <link>https://isa-afp.org/entries/Munta_Model_Checker.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Munta: A Verified Model Checker for Timed Automata in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+    <category>Tools</category>
+</item>
+<item>
   <title>Difference Bound Matrices</title>
   <link>https://isa-afp.org/entries/Difference_Bound_Matrices.html</link>
   <pubDate>Wed, 14 Aug 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/wolff/index.html
--- web/authors/wolff/index.html
+++ web/authors/wolff/index.html
@@ -106,6 +106,70 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
+            <a class="title" href="../../entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">Apr 10</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/Hilbert_Basis.html">A Proof of Hilbert Basis Theorem and            an Extension to Formal Power Series</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/wolff/index.xml
--- web/authors/wolff/index.xml
+++ web/authors/wolff/index.xml
@@ -4,6 +4,50 @@
     <title>Burkhart Wolff</title>
     <link>https://isa-afp.org/authors/wolff/</link>
     <description>AFP entries of Burkhart Wolff</description><item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>CSP Semantics over Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/HOL-CSP_RS.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>CSP Semantics over Restriction Spaces in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Definition and Elementary Properties of Ultrametric Spaces</title>
+  <link>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</link>
+  <pubDate>Thu, 10 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Definition and Elementary Properties of Ultrametric Spaces in the AFP</description>
+    <category>Mathematics/Topology</category>
+</item>
+<item>
   <title>A Proof of Hilbert Basis Theorem and            an Extension to Formal Power Series</title>
   <link>https://isa-afp.org/entries/Hilbert_Basis.html</link>
   <pubDate>Wed, 12 Feb 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/yamada/index.html
--- web/authors/yamada/index.html
+++ web/authors/yamada/index.html
@@ -108,6 +108,24 @@
       <article class="entry">
         <div class="item-text">
           <h5>
+            <a class="title" href="../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+          </h5>
+          <br>
+          by
+            <a href="../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5>
             <a class="title" href="../../entries/First_Order_Rewriting.html">First-Order Rewriting</a>
           </h5>
           <br>
diff -r 848f2c00e57f -r 29846bc89de2 web/authors/yamada/index.xml
--- web/authors/yamada/index.xml
+++ web/authors/yamada/index.xml
@@ -4,6 +4,13 @@
     <title>Akihisa Yamada</title>
     <link>https://isa-afp.org/authors/yamada/</link>
     <description>AFP entries of Akihisa Yamada</description><item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+<item>
   <title>First-Order Rewriting</title>
   <link>https://isa-afp.org/entries/First_Order_Rewriting.html</link>
   <pubDate>Sun, 13 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/data/keywords.json
--- web/data/keywords.json
+++ web/data/keywords.json
@@ -192,6 +192,8 @@
 {"keyword":"adjoint functors"},
 {"keyword":"adjoint functors preserve limits"},
 {"keyword":"adjunctions"},
+{"keyword":"admissibility subgoals"},
+{"keyword":"admit natural notions"},
 {"keyword":"advanced algorithms"},
 {"keyword":"advanced binding constructs"},
 {"keyword":"advanced database systems"},
@@ -520,7 +522,9 @@
 {"keyword":"automatically extracted scala code"},
 {"keyword":"automatically generate proofs"},
 {"keyword":"automatically refines algorithms"},
+{"keyword":"automatically solve"},
 {"keyword":"automatically transferable"},
+{"keyword":"automating verification tasks"},
 {"keyword":"automation mechanisms"},
 {"keyword":"automation setup aimed"},
 {"keyword":"automatize canonical tasks"},
@@ -531,6 +535,7 @@
 {"keyword":"autonomous vehicle manufacturers"},
 {"keyword":"autoref tool"},
 {"keyword":"auxiliary labels"},
+{"keyword":"auxiliary primes"},
 {"keyword":"auxiliary triangle inequality"},
 {"keyword":"auxiliary type"},
 {"keyword":"average case"},
@@ -636,6 +641,8 @@
 {"keyword":"bayes theorem"},
 {"keyword":"bayesian regression presented"},
 {"keyword":"beautiful result"},
+{"keyword":"begin align"},
+{"keyword":"begin align dist"},
 {"keyword":"begin alignat 3"},
 {"keyword":"begin alignat 3 sum_"},
 {"keyword":"behavior trace assertions"},
@@ -811,6 +818,7 @@
 {"keyword":"called check monad"},
 {"keyword":"called complete sets"},
 {"keyword":"called concurrent transition systems"},
+{"keyword":"called distance"},
 {"keyword":"called eudoxus reals"},
 {"keyword":"called galois fields"},
 {"keyword":"called hol-csp 1"},
@@ -961,6 +969,7 @@
 {"keyword":"church-style simply-typed"},
 {"keyword":"cidr notation"},
 {"keyword":"ciphertext attacks"},
+{"keyword":"circles obtained"},
 {"keyword":"circular intuition"},
 {"keyword":"circus environment supports"},
 {"keyword":"circus language"},
@@ -1004,6 +1013,7 @@
 {"keyword":"classical theorem"},
 {"keyword":"classical theorem stating"},
 {"keyword":"classical two-sided matching scenarios"},
+{"keyword":"classical version"},
 {"keyword":"classical yun algorithm"},
 {"keyword":"classifies topological spaces"},
 {"keyword":"clausal consequences"},
@@ -1108,6 +1118,7 @@
 {"keyword":"combine stepwise refinement"},
 {"keyword":"combined factorization algorithm"},
 {"keyword":"combined result"},
+{"keyword":"combining adjacent elements"},
 {"keyword":"command"},
 {"keyword":"command mk_ide"},
 {"keyword":"command mk_ide enables"},
@@ -1368,6 +1379,7 @@
 {"keyword":"cone text arg"},
 {"keyword":"conference certified programs"},
 {"keyword":"conference interactive theorem proving"},
+{"keyword":"conference publications 4"},
 {"keyword":"confidential events"},
 {"keyword":"confidentiality guarantees"},
 {"keyword":"confidentiality properties"},
@@ -1554,6 +1566,7 @@
 {"keyword":"crytographic standards"},
 {"keyword":"csp noninterference security"},
 {"keyword":"csp noninterference security stated"},
+{"keyword":"csp process algebra"},
 {"keyword":"cubic equations"},
 {"keyword":"cubic space"},
 {"keyword":"current compression formats"},
@@ -1912,6 +1925,7 @@
 {"keyword":"dominated terms"},
 {"keyword":"dot-decimal notation"},
 {"keyword":"double commutant theorem"},
+{"keyword":"downarrow mathrm min"},
 {"keyword":"dprm theorem"},
 {"keyword":"dra targets similar applications"},
 {"keyword":"draft paper"},
@@ -2013,6 +2027,7 @@
 {"keyword":"efsm level"},
 {"keyword":"efsms execute traces"},
 {"keyword":"eigenvalue interlacing theorem"},
+{"keyword":"eisenstein series g_n"},
 {"keyword":"electronic proceedings"},
 {"keyword":"elegant encoding"},
 {"keyword":"elegant proof"},
@@ -2022,6 +2037,7 @@
 {"keyword":"elementary discrete inequality"},
 {"keyword":"elementary divisor rings"},
 {"keyword":"elementary facts"},
+{"keyword":"elementary formalization"},
 {"keyword":"elementary infrastructure"},
 {"keyword":"elementary measure theory"},
 {"keyword":"elementary methods"},
@@ -2041,6 +2057,7 @@
 {"keyword":"elliott mendelson"},
 {"keyword":"elliptic curve"},
 {"keyword":"elliptic curve cryptography"},
+{"keyword":"elliptic function"},
 {"keyword":"embedded logic"},
 {"keyword":"embedding path order"},
 {"keyword":"emphasising local spatial properties"},
@@ -2062,6 +2079,7 @@
 {"keyword":"encoding function"},
 {"keyword":"encourages code"},
 {"keyword":"encryption schemes"},
+{"keyword":"end align"},
 {"keyword":"end alignat"},
 {"keyword":"enforcing exclusive writes"},
 {"keyword":"engineering safety"},
@@ -2093,6 +2111,7 @@
 {"keyword":"entry lies"},
 {"keyword":"entry presents"},
 {"keyword":"entry strong security"},
+{"keyword":"entry supersedes earlier versions"},
 {"keyword":"entry vcg auctions"},
 {"keyword":"entry works"},
 {"keyword":"enumeration functions"},
@@ -2106,6 +2125,7 @@
 {"keyword":"equational axiomatisation"},
 {"keyword":"equational axioms"},
 {"keyword":"equational reasoning"},
+{"keyword":"equipping restriction spaces"},
 {"keyword":"equiv 3 pmod 8"},
 {"keyword":"equiv frac mathcal"},
 {"keyword":"equiv_ text"},
@@ -2225,6 +2245,7 @@
 {"keyword":"existing functionality"},
 {"keyword":"existing geometry libraries"},
 {"keyword":"existing hoare logic"},
+{"keyword":"existing holcf infrastructure"},
 {"keyword":"existing implementation"},
 {"keyword":"existing integration theory"},
 {"keyword":"existing libraries"},
@@ -2375,6 +2396,8 @@
 {"keyword":"famous result"},
 {"keyword":"famously short one-page proof"},
 {"keyword":"far-reaching impossibility theorem"},
+{"keyword":"farey fractions"},
+{"keyword":"farey sequence f_n"},
 {"keyword":"fast iterative algorithm"},
 {"keyword":"fast number theoretic transform"},
 {"keyword":"fast sat solver"},
@@ -2509,7 +2532,9 @@
 {"keyword":"fixed set"},
 {"keyword":"fixed time-unit"},
 {"keyword":"fixed upper bound"},
+{"keyword":"fixed-point operator"},
 {"keyword":"fixed-point theorem"},
+{"keyword":"fixed-point theorem established"},
 {"keyword":"fixed-width machine words"},
 {"keyword":"fixpoint operations lfp"},
 {"keyword":"fixpoint theorem"},
@@ -2539,6 +2564,7 @@
 {"keyword":"folder listinf"},
 {"keyword":"folklore results related"},
 {"keyword":"follow stefan schwoon"},
+{"keyword":"forbidden series g_2"},
 {"keyword":"foreach combinators"},
 {"keyword":"fork"},
 {"keyword":"form bigwedge_"},
@@ -2621,6 +2647,7 @@
 {"keyword":"formally verified security analysis"},
 {"keyword":"formally verified solver"},
 {"keyword":"formally verify gauss-seidel"},
+{"keyword":"forming assertions"},
 {"keyword":"formula mdp ta pta"},
 {"keyword":"formula represent propositional formulas"},
 {"keyword":"formulas"},
@@ -2759,6 +2786,7 @@
 {"keyword":"fundamental objects"},
 {"keyword":"fundamental problems"},
 {"keyword":"fundamental properties"},
+{"keyword":"fundamental region elliptic functions"},
 {"keyword":"fundamental result"},
 {"keyword":"fundamental solution"},
 {"keyword":"fundamental subspaces"},
@@ -2855,6 +2883,7 @@
 {"keyword":"generalized recurrence relation"},
 {"keyword":"generalized sumcheck protocol"},
 {"keyword":"generalized sylvester matrices"},
+{"keyword":"generalizes sophie germain primes"},
 {"keyword":"generalizes sutherland"},
 {"keyword":"generally detectable"},
 {"keyword":"generate"},
@@ -2908,6 +2937,7 @@
 {"keyword":"generic work-list algorithm"},
 {"keyword":"generic worklist algorithm"},
 {"keyword":"generic-deriving package"},
+{"keyword":"genuine complementarity"},
 {"keyword":"geocoq library"},
 {"keyword":"geodesic gromov-hyperbolic space"},
 {"keyword":"geodesic metric space"},
@@ -3176,6 +3206,7 @@
 {"keyword":"holcf package"},
 {"keyword":"holds"},
 {"keyword":"holomorphic automorphisms"},
+{"keyword":"holomorphic functions"},
 {"keyword":"holzf theory"},
 {"keyword":"hom embedding"},
 {"keyword":"homogeneous linear diophantine equations"},
@@ -3252,6 +3283,7 @@
 {"keyword":"implementation matches"},
 {"keyword":"implementation mixes"},
 {"keyword":"implementation relates pointer-based computation"},
+{"keyword":"implementation relies"},
 {"keyword":"implementation runs"},
 {"keyword":"implementation supports set membership"},
 {"keyword":"implemented multi-"},
@@ -3304,6 +3336,7 @@
 {"keyword":"increasing binary trees"},
 {"keyword":"increasing rational sequence r_n"},
 {"keyword":"increasingly important"},
+{"keyword":"increasingly refined type classes"},
 {"keyword":"incredible proof machine"},
 {"keyword":"incremental verification"},
 {"keyword":"incrementally check"},
@@ -3413,6 +3446,7 @@
 {"keyword":"infty infty frac"},
 {"keyword":"ingo van duijn"},
 {"keyword":"inherently based"},
+{"keyword":"inherited partial order structure"},
 {"keyword":"initial conversion"},
 {"keyword":"initial nonterminal"},
 {"keyword":"initial proof"},
@@ -3483,6 +3517,7 @@
 {"keyword":"interest distributed"},
 {"keyword":"interest rate"},
 {"keyword":"interesting case study"},
+{"keyword":"interesting consequences follow"},
 {"keyword":"interesting data structure"},
 {"keyword":"interesting formalization exercise"},
 {"keyword":"interesting format"},
@@ -3538,6 +3573,8 @@
 {"keyword":"intuitive arguments found"},
 {"keyword":"intuitive combinatorial proof"},
 {"keyword":"intuitive desired security policy"},
+{"keyword":"intuitive implementation"},
+{"keyword":"intuitive lifted expression syntax"},
 {"keyword":"intuitively secure programs"},
 {"keyword":"invariance"},
 {"keyword":"invariant based programming"},
@@ -3552,6 +3589,7 @@
 {"keyword":"inverse squares"},
 {"keyword":"inverse transform ifntt"},
 {"keyword":"inverse transform intt"},
+{"keyword":"inversion formulas"},
 {"keyword":"inversion rules"},
 {"keyword":"inversions"},
 {"keyword":"invertible lossless transformation"},
@@ -3659,6 +3697,7 @@
 {"keyword":"key confirmation"},
 {"keyword":"key construction"},
 {"keyword":"key contribution"},
+{"keyword":"key csp operators"},
 {"keyword":"key encapsulation mechanism"},
 {"keyword":"key establishment protocols"},
 {"keyword":"key proofs"},
@@ -3763,6 +3802,8 @@
 {"keyword":"latin rectangle"},
 {"keyword":"latin square"},
 {"keyword":"lattice basis reduction"},
+{"keyword":"lattice discriminant delta g_n"},
+{"keyword":"lattice invariants g_2"},
 {"keyword":"lattice ordered groups"},
 {"keyword":"lattice point"},
 {"keyword":"lattice supremum providing"},
@@ -3774,6 +3815,7 @@
 {"keyword":"laurent expansion"},
 {"keyword":"law"},
 {"keyword":"lawrence paulson"},
+{"keyword":"layered approach enables"},
 {"keyword":"lazy list"},
 {"keyword":"lazy sequences"},
 {"keyword":"lazy-list equality"},
@@ -3837,6 +3879,7 @@
 {"keyword":"lifting step"},
 {"keyword":"lifts resolution derivation steps"},
 {"keyword":"light-weight type system"},
+{"keyword":"lightweight integration"},
 {"keyword":"lim"},
 {"keyword":"limit process"},
 {"keyword":"limit process exists"},
@@ -3959,6 +4002,7 @@
 {"keyword":"longer valid"},
 {"keyword":"longest lyndon suffix"},
 {"keyword":"longest recognized substrings"},
+{"keyword":"longrightarrow exists"},
 {"keyword":"loop freedom"},
 {"keyword":"lov aacutesz local lemma"},
 {"keyword":"low edge probability"},
@@ -4019,6 +4063,7 @@
 {"keyword":"main order fully coincides"},
 {"keyword":"main premise"},
 {"keyword":"main result"},
+{"keyword":"main result establishes"},
 {"keyword":"main results verified"},
 {"keyword":"main routing table"},
 {"keyword":"main themes"},
@@ -4259,7 +4304,10 @@
 {"keyword":"modular arithmetic plays"},
 {"keyword":"modular assembly kit"},
 {"keyword":"modular exponentiation"},
+{"keyword":"modular forms"},
 {"keyword":"modular functions"},
+{"keyword":"modular group dedekind"},
+{"keyword":"modular group gamma"},
 {"keyword":"modular hierarchy"},
 {"keyword":"module development"},
 {"keyword":"modulo operation"},
@@ -4470,6 +4518,7 @@
 {"keyword":"noninterference theorem"},
 {"keyword":"nonnegative cubes"},
 {"keyword":"nonredundant clauses"},
+{"keyword":"nontrivial integer solutions"},
 {"keyword":"nontrivial size"},
 {"keyword":"nonzero rational number"},
 {"keyword":"nora szasz"},
@@ -4569,6 +4618,8 @@
 {"keyword":"omega omega"},
 {"keyword":"omega operation"},
 {"keyword":"omega-complete non-orders"},
+{"keyword":"omega_1 mathbb"},
+{"keyword":"omega_1 omega_2 notin mathbb"},
 {"keyword":"omnipresent foundational errors"},
 {"keyword":"one-complete computably enumerable set"},
 {"keyword":"one-dimensional case"},
@@ -4586,6 +4637,7 @@
 {"keyword":"open publishing association"},
 {"keyword":"open sets"},
 {"keyword":"operating system"},
+{"keyword":"operation called"},
 {"keyword":"operation results"},
 {"keyword":"operational"},
 {"keyword":"operational behavior derived"},
@@ -4634,6 +4686,7 @@
 {"keyword":"ordinal sum"},
 {"keyword":"ordinary assertional reasoning"},
 {"keyword":"ordinary category"},
+{"keyword":"ordinary differential equation satisfied"},
 {"keyword":"ordinary differential equations"},
 {"keyword":"ordinary functions"},
 {"keyword":"ordinary generating function"},
@@ -4827,6 +4880,7 @@
 {"keyword":"perfect square"},
 {"keyword":"perform stream fusion"},
 {"keyword":"perform update operations naively"},
+{"keyword":"performing substitutions"},
 {"keyword":"performs comparable"},
 {"keyword":"periodic arithmetic functions"},
 {"keyword":"periodic bernoulli polynomials"},
@@ -4876,6 +4930,7 @@
 {"keyword":"poincar disc model"},
 {"keyword":"poincar disc model development"},
 {"keyword":"point-wise reasoning"},
+{"keyword":"pointed complete partial order"},
 {"keyword":"points constructible"},
 {"keyword":"pointwise updates"},
 {"keyword":"polar form transformation"},
@@ -4959,6 +5014,7 @@
 {"keyword":"pragmatic interest"},
 {"keyword":"pragmatic reasons"},
 {"keyword":"precise algorithms"},
+{"keyword":"precise comparison"},
 {"keyword":"precise effect"},
 {"keyword":"precise proof"},
 {"keyword":"precise region"},
@@ -4981,6 +5037,7 @@
 {"keyword":"prefix order"},
 {"keyword":"preliminaries chapter"},
 {"keyword":"preliminary evaluations"},
+{"keyword":"preliminary results"},
 {"keyword":"preorder relations"},
 {"keyword":"presburger arithmetic"},
 {"keyword":"present"},
@@ -5101,6 +5158,7 @@
 {"keyword":"problem arithmetic progressions"},
 {"keyword":"problem reduction"},
 {"keyword":"problems"},
+{"keyword":"process algebra framework hol-csp"},
 {"keyword":"process calculi encodings"},
 {"keyword":"process composition based"},
 {"keyword":"process control"},
@@ -5189,6 +5247,7 @@
 {"keyword":"proof step"},
 {"keyword":"proof structure"},
 {"keyword":"proof system"},
+{"keyword":"proof tasks"},
 {"keyword":"proof techniques"},
 {"keyword":"proof technology"},
 {"keyword":"proof term"},
@@ -5268,6 +5327,7 @@
 {"keyword":"publisher component"},
 {"keyword":"publisher subscriber"},
 {"keyword":"publisher subscriber pattern"},
+{"keyword":"pumping lemmas"},
 {"keyword":"pure exchange economy"},
 {"keyword":"pure mathematical subjects"},
 {"keyword":"purely algebraic"},
@@ -5333,6 +5393,7 @@
 {"keyword":"quasi-fixed point"},
 {"keyword":"quelques probl"},
 {"keyword":"query evaluation"},
+{"keyword":"query functions"},
 {"keyword":"query optimization consisting"},
 {"keyword":"query suggestion mechanisms"},
 {"keyword":"queue data structures"},
@@ -5454,7 +5515,9 @@
 {"keyword":"recursive inseparability"},
 {"keyword":"recursive path order"},
 {"keyword":"recursive procedures"},
+{"keyword":"recursive process definitions"},
 {"keyword":"recursive programs based"},
+{"keyword":"recursive specifications"},
 {"keyword":"recursively enumerable set"},
 {"keyword":"recursively expressed"},
 {"keyword":"recursively inseparable"},
@@ -5562,6 +5625,7 @@
 {"keyword":"relevant proof methods"},
 {"keyword":"relevant standard"},
 {"keyword":"rely condition generalised"},
+{"keyword":"rely exclusively"},
 {"keyword":"rely guarantee reasoning"},
 {"keyword":"rely quotient"},
 {"keyword":"rely-guarantee-style reasoning"},
@@ -5637,6 +5701,10 @@
 {"keyword":"restricted solution space"},
 {"keyword":"restricted subset"},
 {"keyword":"restricted type"},
+{"keyword":"restriction spaces"},
+{"keyword":"restriction spaces coincides"},
+{"keyword":"restriction-based notions"},
+{"keyword":"restriction_spaces library works"},
 {"keyword":"restrictive definition"},
 {"keyword":"result due"},
 {"keyword":"result presented"},
@@ -5647,6 +5715,7 @@
 {"keyword":"resulting framework"},
 {"keyword":"resulting generalized counting sort"},
 {"keyword":"resulting hierarchy"},
+{"keyword":"resulting library"},
 {"keyword":"resulting logic"},
 {"keyword":"resulting proof system"},
 {"keyword":"resulting recursion induction rules"},
@@ -5660,6 +5729,7 @@
 {"keyword":"retain key properties"},
 {"keyword":"returns equi-satisfiable inequalities"},
 {"keyword":"reusable building blocks"},
+{"keyword":"reusable framework"},
 {"keyword":"reusable libraries"},
 {"keyword":"reusable modelling"},
 {"keyword":"reusable probability space framework"},
@@ -5687,6 +5757,7 @@
 {"keyword":"riemann zeta function"},
 {"keyword":"riesz-markov-kakutani representation theorem"},
 {"keyword":"right-hand side"},
+{"keyword":"right-linear grammars"},
 {"keyword":"rigorous definition"},
 {"keyword":"rigorous numerical algorithms"},
 {"keyword":"rigorous polynomial approximation"},
@@ -5872,6 +5943,7 @@
 {"keyword":"separation-logic based correctness proofs"},
 {"keyword":"separator smaller"},
 {"keyword":"sepref tool"},
+{"keyword":"sequence f_n"},
 {"keyword":"sequence preserves fairness"},
 {"keyword":"sequent calculus"},
 {"keyword":"sequent calculus prover"},
@@ -6130,6 +6202,7 @@
 {"keyword":"so-called hessenberg"},
 {"keyword":"so-called hyperproperties"},
 {"keyword":"so-called key equalities"},
+{"keyword":"so-called restriction"},
 {"keyword":"so-called sturm sequences"},
 {"keyword":"social decision schemes"},
 {"keyword":"social welfare"},
@@ -6154,6 +6227,8 @@
 {"keyword":"solving equations"},
 {"keyword":"solving linear programs"},
 {"keyword":"solving markov decision processes"},
+{"keyword":"sophie germain"},
+{"keyword":"sophie germain prime"},
 {"keyword":"sophie tourret"},
 {"keyword":"sophisticated languages"},
 {"keyword":"sophisticated object-oriented bytecode language"},
@@ -6175,6 +6250,8 @@
 {"keyword":"space complexity guarantees"},
 {"keyword":"space constraints"},
 {"keyword":"space usage"},
+{"keyword":"spaces equipped"},
+{"keyword":"spaces relying"},
 {"keyword":"spacetime location"},
 {"keyword":"spanning basic algorithms"},
 {"keyword":"spanning subhypergraphs"},
@@ -6263,6 +6340,8 @@
 {"keyword":"standard heron method"},
 {"keyword":"standard laws"},
 {"keyword":"standard logistic function"},
+{"keyword":"standard metric completeness"},
+{"keyword":"standard metric spaces"},
 {"keyword":"standard models"},
 {"keyword":"standard operators"},
 {"keyword":"standard overloaded syntax"},
@@ -6299,6 +6378,7 @@
 {"keyword":"state sigma_a"},
 {"keyword":"state space"},
 {"keyword":"state transformers"},
+{"keyword":"state type"},
 {"keyword":"state-based non-deterministic sequential computations"},
 {"keyword":"state-based semantics based"},
 {"keyword":"state-merging technique"},
@@ -6441,6 +6521,7 @@
 {"keyword":"style presented"},
 {"keyword":"sub-probability mass functions"},
 {"keyword":"subcategory comprised"},
+{"keyword":"subclass relationship"},
 {"keyword":"subject reduction property"},
 {"keyword":"sublists alternately extracted"},
 {"keyword":"submission"},
@@ -6482,6 +6563,7 @@
 {"keyword":"successor function"},
 {"keyword":"successor search"},
 {"keyword":"sufficient condition"},
+{"keyword":"sufficient criteria"},
 {"keyword":"sufficient criterion"},
 {"keyword":"sufficiently efficient"},
 {"keyword":"sufficiently good"},
@@ -6529,6 +6611,7 @@
 {"keyword":"support negative joins"},
 {"keyword":"support tostring functions"},
 {"keyword":"supported unicode characters"},
+{"keyword":"supporting automated reasoning"},
 {"keyword":"supporting automatic refinement"},
 {"keyword":"supports low-effort security proofs"},
 {"keyword":"supports mutual recursion"},
@@ -6570,6 +6653,7 @@
 {"keyword":"syntactic rewrite rules"},
 {"keyword":"syntactic sugar binding precedences"},
 {"keyword":"syntax syntax"},
+{"keyword":"syntax translations"},
 {"keyword":"syntax tree"},
 {"keyword":"syntax-independent logic infrastructure"},
 {"keyword":"synthesize imperative programs"},
@@ -6582,6 +6666,7 @@
 {"keyword":"system types"},
 {"keyword":"system verification"},
 {"keyword":"systematic development"},
+{"keyword":"systematically relate ultrametric"},
 {"keyword":"systems communication"},
 {"keyword":"systems communication patterns"},
 {"keyword":"systems communication plays"},
@@ -6713,6 +6798,7 @@
 {"keyword":"theoretical framework"},
 {"keyword":"theoretical insights"},
 {"keyword":"theoretically incomparable"},
+{"keyword":"theories guaranteeing existence"},
 {"keyword":"theories list"},
 {"keyword":"theories listinf"},
 {"keyword":"theories presented"},
@@ -6781,6 +6867,7 @@
 {"keyword":"topological groups"},
 {"keyword":"topological nature"},
 {"keyword":"topological proof"},
+{"keyword":"topological properties"},
 {"keyword":"topological space generated"},
 {"keyword":"topological spaces"},
 {"keyword":"torino group"},
@@ -6860,6 +6947,7 @@
 {"keyword":"triangle"},
 {"keyword":"triangle constructed"},
 {"keyword":"triangle counting lemma"},
+{"keyword":"triangle inequality"},
 {"keyword":"triangle removal lemma"},
 {"keyword":"triangular numbers"},
 {"keyword":"trick"},
@@ -6928,6 +7016,7 @@
 {"keyword":"uiuc"},
 {"keyword":"ultimately culminating"},
 {"keyword":"ultimately refutational completeness"},
+{"keyword":"ultrametric spaces"},
 {"keyword":"uml class diagrams"},
 {"keyword":"unbounded nondeterminism"},
 {"keyword":"unbounded sequences"},
@@ -6993,6 +7082,7 @@
 {"keyword":"unify computation models"},
 {"keyword":"unify correctness statements"},
 {"keyword":"unifying theories"},
+{"keyword":"unimodular relation bc"},
 {"keyword":"uninterpreted functions"},
 {"keyword":"union concatenation"},
 {"keyword":"union equals"},
@@ -7067,13 +7157,16 @@
 {"keyword":"updown scheme"},
 {"keyword":"upf emphasizes"},
 {"keyword":"upper bound"},
+{"keyword":"upper half plane"},
 {"keyword":"upper semicontinuous"},
 {"keyword":"upper triangular"},
 {"keyword":"upper-bound remains valid"},
 {"keyword":"usable framework"},
 {"keyword":"useless zero-reductions"},
 {"keyword":"user command"},
+{"keyword":"users write higher-order rules"},
 {"keyword":"usual definitions"},
+{"keyword":"usual metric structures"},
 {"keyword":"usual monad laws"},
 {"keyword":"usual programming constructs"},
 {"keyword":"usual redundancy criteria based"},
@@ -7100,6 +7193,7 @@
 {"keyword":"variable assignment"},
 {"keyword":"variable convention"},
 {"keyword":"variable-for-variable substitution"},
+{"keyword":"variant involving"},
 {"keyword":"variants"},
 {"keyword":"varphi_i vee mathbf"},
 {"keyword":"varying numbers"},
@@ -7111,6 +7205,7 @@
 {"keyword":"vector cross product"},
 {"keyword":"vector fields"},
 {"keyword":"vector space"},
+{"keyword":"verifiable side-conditions"},
 {"keyword":"verification back-ends"},
 {"keyword":"verification components"},
 {"keyword":"verification condition generation"},
@@ -7136,6 +7231,7 @@
 {"keyword":"verified heap functions"},
 {"keyword":"verified implementation"},
 {"keyword":"verified iptables firewall analysis"},
+{"keyword":"verified model checker"},
 {"keyword":"verified monitor"},
 {"keyword":"verified monitor implements"},
 {"keyword":"verified parser generator"},
@@ -7181,6 +7277,7 @@
 {"keyword":"vertex set"},
 {"keyword":"vertical composite"},
 {"keyword":"vertical residuation"},
+{"keyword":"vice versa"},
 {"keyword":"view sorted sets"},
 {"keyword":"vincent bloemen"},
 {"keyword":"vincent rahli"},
@@ -7232,6 +7329,7 @@
 {"keyword":"web components"},
 {"keyword":"web standards"},
 {"keyword":"webassembly language"},
+{"keyword":"weierstra elliptic function wp"},
 {"keyword":"weight-balanced trees"},
 {"keyword":"weighted arithmetic geometric"},
 {"keyword":"weighted graphs"},
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Algebraic_Numbers.html
--- web/entries/Algebraic_Numbers.html
+++ web/entries/Algebraic_Numbers.html
@@ -170,6 +170,7 @@
               <li><a href="../entries/Cubic_Quartic_Equations.html">Solving Cubic and Quartic Equations</a></li>
               <li><a href="../entries/Factor_Algebraic_Polynomial.html">Factorization of Polynomials with Algebraic Coefficients</a></li>
               <li><a href="../entries/Quantifier_Elimination_Hybrid.html">A First Complete Algorithm for Real Quantifier Elimination in Isabelle/HOL</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/AutoCorres2.html
--- web/entries/AutoCorres2.html
+++ web/entries/AutoCorres2.html
@@ -538,6 +538,12 @@
               <li><a href="../entries/Word_Lib.html">Finite Machine Word Library</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Certification_Monads.html
--- web/entries/Certification_Monads.html
+++ web/entries/Certification_Monads.html
@@ -144,6 +144,7 @@
               <li><a href="../entries/AI_Planning_Languages_Semantics.html">AI Planning Languages Semantics</a></li>
               <li><a href="../entries/Pattern_Completeness.html">Verifying a Decision Procedure for Pattern Completeness</a></li>
               <li><a href="../entries/First_Order_Rewriting.html">First-Order Rewriting</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Context_Free_Grammar.html
--- /dev/null
+++ web/entries/Context_Free_Grammar.html
@@ -0,0 +1,198 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Context-Free Grammars and Languages - Archive of Formal Proofs</title>
+    <meta name="description" content="Context-Free Grammars and Languages in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Context-Free Grammars and Languages in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Context-Free Grammars and Languages" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Context_Free_Grammar.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>C</span>ontext-<span class='first'>F</span>ree <span class='first'>G</span>rammars and <span class='first'>L</span>anguages
+  </h1>
+  <div>
+      <p><a href="../authors/nipkow/">Tobias Nipkow</a> <a class="obfuscated" data="eyJ1c2VyIjpbIm5pcGtvdyJdLCJob3N0IjpbImluIiwidHVtIiwiZGUiXX0=">📧</a>, <a href="../authors/gschossmann/">Markus Gschoßmann</a> <a class="obfuscated" data="eyJ1c2VyIjpbIm1hcmt1cyIsImdzY2hvc3NtYW5uIl0sImhvc3QiOlsidHVtIiwiZGUiXX0=">📧</a>, <a href="../authors/krayer/">Felix Krayer</a> <a class="obfuscated" data="eyJ1c2VyIjpbImZlbGl4Iiwia3JheWVyIl0sImhvc3QiOlsidHVtIiwiZGUiXX0=">📧</a>, <a href="../authors/lehr/">Fabian Lehr</a> <a class="obfuscated" data="eyJ1c2VyIjpbImZhYmlhbiIsImxlaHIiXSwiaG9zdCI6WyJ0dW0iLCJkZSJdfQ==">📧</a>, <a href="../authors/philipp/">Bruno Philipp</a> <a class="obfuscated" data="eyJ1c2VyIjpbImdlOTVmZWsiXSwiaG9zdCI6WyJteXR1bSIsImRlIl19">📧</a>, <a href="../authors/stimpfle/">August Martin Stimpfle</a>, <a href="../authors/taskin/">Kaan Taskin</a> <a class="obfuscated" data="eyJ1c2VyIjpbImthYW50YXNraW4yMDIyIl0sImhvc3QiOlsiZ21haWwiLCJjb20iXX0=">📧</a> and <a href="../authors/yamada/">Akihisa Yamada</a> <a class="obfuscated" data="eyJ1c2VyIjpbImFraWhpc2EiLCJ5YW1hZGEiXSwiaG9zdCI6WyJ1aWJrIiwiYWMiLCJhdCJdfQ==">📧</a>
+      </p>
+      <p class="date">May 21, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">This is a basic library of definitions and results about context-free grammars and languages.
+It includes context-free grammars and languages, parse trees, Chomsky normal form,
+pumping lemmas and the relationship of right-linear grammars to finite automata.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/computer-science/automata-and-formal-languages/">Computer science/Automata and formal languages</a></li>
+      </ul>
+      <h3>Session Context_Free_Grammar</h3>
+      <ul>
+          <li><a href="../sessions/context_free_grammar/#Context_Free_Grammar">Context_Free_Grammar</a></li>
+          <li><a href="../sessions/context_free_grammar/#Parse_Tree">Parse_Tree</a></li>
+          <li><a href="../sessions/context_free_grammar/#Renaming_CFG">Renaming_CFG</a></li>
+          <li><a href="../sessions/context_free_grammar/#Disjoint_Union_CFG">Disjoint_Union_CFG</a></li>
+          <li><a href="../sessions/context_free_grammar/#Context_Free_Language">Context_Free_Language</a></li>
+          <li><a href="../sessions/context_free_grammar/#Unit_Elimination">Unit_Elimination</a></li>
+          <li><a href="../sessions/context_free_grammar/#Epsilon_Elimination">Epsilon_Elimination</a></li>
+          <li><a href="../sessions/context_free_grammar/#Chomsky_Normal_Form">Chomsky_Normal_Form</a></li>
+          <li><a href="../sessions/context_free_grammar/#Pumping_Lemma_CFG">Pumping_Lemma_CFG</a></li>
+          <li><a href="../sessions/context_free_grammar/#AnBnCn_not_CFL">AnBnCn_not_CFL</a></li>
+          <li><a href="../sessions/context_free_grammar/#CFL_Not_Intersection_Closed">CFL_Not_Intersection_Closed</a></li>
+          <li><a href="../sessions/context_free_grammar/#Inlining1Prod">Inlining1Prod</a></li>
+          <li><a href="../sessions/context_free_grammar/#Binarize">Binarize</a></li>
+          <li><a href="../sessions/context_free_grammar/#Right_Linear">Right_Linear</a></li>
+          <li><a href="../sessions/context_free_grammar/#NDA_rlin2">NDA_rlin2</a></li>
+          <li><a href="../sessions/context_free_grammar/#Right_Linear_Automata">Right_Linear_Automata</a></li>
+          <li><a href="../sessions/context_free_grammar/#Pumping_Lemma_Regular">Pumping_Lemma_Regular</a></li>
+          <li><a href="../sessions/context_free_grammar/#AnBn_Not_Regular">AnBn_Not_Regular</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Finite_Automata_HF.html">Finite Automata in Hereditarily Finite Set Theory</a></li>
+              <li><a href="../entries/List_Power.html">Power Operator for Lists</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Context_Free_Grammar-AFP</p>
+        <pre id="copy-text">@article{Context_Free_Grammar-AFP,
+  author  = {Tobias Nipkow and Markus Gschoßmann and Felix Krayer and Fabian Lehr and Bruno Philipp and August Martin Stimpfle and Kaan Taskin and Akihisa Yamada},
+  title   = {Context-Free Grammars and Languages},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Context_Free_Grammar.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Context_Free_Grammar-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Cotangent_PFD_Formula.html
--- web/entries/Cotangent_PFD_Formula.html
+++ web/entries/Cotangent_PFD_Formula.html
@@ -130,6 +130,12 @@
       </ul>
 
     <div class="flex-wrap">
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Decreasing-Diagrams-II.html
--- web/entries/Decreasing-Diagrams-II.html
+++ web/entries/Decreasing-Diagrams-II.html
@@ -132,12 +132,6 @@
               <li><a href="../entries/Well_Quasi_Orders.html">Well-Quasi-Orders</a></li>
           </ul>
         </div>
-        <div>
-          <h3>Similar entries</h3>
-          <ul class="horizontal-list">
-              <li><a href="../entries/Decreasing-Diagrams.html">Decreasing Diagrams</a></li>
-          </ul>
-        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Decreasing-Diagrams.html
--- web/entries/Decreasing-Diagrams.html
+++ web/entries/Decreasing-Diagrams.html
@@ -135,12 +135,6 @@
               <li><a href="../entries/Completeness_Decreasing_Diagrams_for_N1.html">Completeness of Decreasing Diagrams for the Least Uncountable Cardinality</a></li>
           </ul>
         </div>
-        <div>
-          <h3>Similar entries</h3>
-          <ul class="horizontal-list">
-              <li><a href="../entries/Decreasing-Diagrams-II.html">Decreasing Diagrams II</a></li>
-          </ul>
-        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Dedekind_Sums.html
--- web/entries/Dedekind_Sums.html
+++ web/entries/Dedekind_Sums.html
@@ -143,6 +143,12 @@
               <li><a href="../entries/Bernoulli.html">Bernoulli Numbers</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Elementary_Ultrametric_Spaces.html
--- /dev/null
+++ web/entries/Elementary_Ultrametric_Spaces.html
@@ -0,0 +1,189 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Definition and Elementary Properties of Ultrametric Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="Definition and Elementary Properties of Ultrametric Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Definition and Elementary Properties of Ultrametric Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Definition and Elementary Properties of Ultrametric Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>D</span>efinition and <span class='first'>E</span>lementary <span class='first'>P</span>roperties of <span class='first'>U</span>ltrametric <span class='first'>S</span>paces
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a>, <a href="../authors/puyobro/">Benjamin Puyobro</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbmphbWluIiwicHV5b2JybyJdLCJob3N0IjpbImxhcG9zdGUiLCJuZXQiXX0=">📧</a> and <a href="../authors/wolff/">Burkhart Wolff</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJ1cmtoYXJ0Iiwid29sZmYiXSwiaG9zdCI6WyJsbWYiLCJjbnJzIiwiZnIiXX0=">📧</a>
+      </p>
+      <p class="date">April 10, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">An ultrametric space is a metric space in which the triangle inequality is strengthened
+by using the maximum instead of the sum. More formally, such a space is equipped with
+a real-valued application $dist$, called distance, verifying the four following conditions.
+\begin{align*}
+    & dist \ x \ y \ge 0\\
+    & dist \ x \ y = dist \ y \ x\\
+    & dist \ x \ y = 0 \longleftrightarrow x = y\\
+    & dist \ x \ z \le max \ (dist \ x \ y) \ (dist \ y \ z)
+\end{align*}
+In this entry, we present an elementary formalization of these spaces relying on axiomatic type classes.
+The connection with standard metric spaces is obtained through a subclass relationship,
+and fundamental properties of ultrametric spaces are formally established.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/mathematics/topology/">Mathematics/Topology</a></li>
+      </ul>
+      <h3>Session Elementary_Ultrametric_Spaces</h3>
+      <ul>
+          <li><a href="../sessions/elementary_ultrametric_spaces/#Elementary_Ultrametric_Spaces">Elementary_Ultrametric_Spaces</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Elementary_Ultrametric_Spaces/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Elementary_Ultrametric_Spaces/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Elementary_Ultrametric_Spaces/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Elementary_Ultrametric_Spaces-AFP</p>
+        <pre id="copy-text">@article{Elementary_Ultrametric_Spaces-AFP,
+  author  = {Benoît Ballenghien and Benjamin Puyobro and Burkhart Wolff},
+  title   = {Definition and Elementary Properties of Ultrametric Spaces},
+  journal = {Archive of Formal Proofs},
+  month   = {April},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Elementary_Ultrametric_Spaces-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Farey_Sequences.html
--- /dev/null
+++ web/entries/Farey_Sequences.html
@@ -0,0 +1,186 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Farey Sequences and Ford Circles - Archive of Formal Proofs</title>
+    <meta name="description" content="Farey Sequences and Ford Circles in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Farey Sequences and Ford Circles in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Farey Sequences and Ford Circles" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Farey_Sequences.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>F</span>arey <span class='first'>S</span>equences and <span class='first'>F</span>ord <span class='first'>C</span>ircles
+  </h1>
+  <div>
+      <p><a href="../authors/paulson/">Lawrence C. Paulson</a> <a class="obfuscated" data="eyJ1c2VyIjpbImxwMTUiXSwiaG9zdCI6WyJjYW0iLCJhYyIsInVrIl19">📧</a>
+      </p>
+      <p class="date">May 15, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">The sequence $F_n$ of Farey fractions of order $n$ 
+has the form 
+$$\frac{0}{1}, \frac{1}{n}, \frac{1}{n-1}, \ldots, \frac{n-1}{n}, \frac{1}{1}$$
+where the fractions appear in numerical order and have denominators at most $n$.
+The transformation from $F_n$ to $F_{n+1}$ can be effected by combining adjacent elements of 
+the sequence~$F_n$, using an operation called the mediant.
+Adjacent (reduced) fractions $(a/b) < (c/d)$ satisfy the unimodular
+relation $bc - ad = 1$ and their mediant is $\frac{a+c}{b+d}$.
+A Ford circle is specified by a rational number, and interesting consequences follow
+in the case of Ford circles obtained from some Farey sequence~$F_n$.
+The formalised material is drawn from Apostol's Modular Functions and Dirichlet Series in Number Theory.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/mathematics/number-theory/">Mathematics/Number theory</a></li>
+      </ul>
+      <h3>Related publications</h3>
+      <ul>
+        <li>Modular Functions and Dirichlet Series in Number Theory</li>
+      </ul>
+      <h3>Session Farey_Sequences</h3>
+      <ul>
+          <li><a href="../sessions/farey_sequences/#Farey_Ford">Farey_Ford</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Farey_Sequences/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Farey_Sequences/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Farey_Sequences/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Farey_Sequences-AFP</p>
+        <pre id="copy-text">@article{Farey_Sequences-AFP,
+  author  = {Lawrence C. Paulson},
+  title   = {Farey Sequences and Ford Circles},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Farey_Sequences.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Farey_Sequences-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Fermat3_4.html
--- web/entries/Fermat3_4.html
+++ web/entries/Fermat3_4.html
@@ -125,6 +125,12 @@
       </ul>
 
     <div class="flex-wrap">
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Sophie_Germain.html">Sophie Germain’s Theorem</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/FinFun.html
--- web/entries/FinFun.html
+++ web/entries/FinFun.html
@@ -146,6 +146,7 @@
               <li><a href="../entries/Nominal2.html">Nominal 2</a></li>
               <li><a href="../entries/Extended_Finite_State_Machines.html">A Formal Model of Extended Finite State Machines</a></li>
               <li><a href="../entries/Pattern_Completeness.html">Verifying a Decision Procedure for Pattern Completeness</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Finite_Automata_HF.html
--- web/entries/Finite_Automata_HF.html
+++ web/entries/Finite_Automata_HF.html
@@ -134,6 +134,12 @@
               <li><a href="../entries/Regular-Sets.html">Regular Sets and Expressions</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/First_Order_Clause.html
--- web/entries/First_Order_Clause.html
+++ web/entries/First_Order_Clause.html
@@ -120,14 +120,15 @@
       </ul>
       <h3>Session First_Order_Clause</h3>
       <ul>
+          <li><a href="../sessions/first_order_clause/#Context_Notation">Context_Notation</a></li>
           <li><a href="../sessions/first_order_clause/#Typing">Typing</a></li>
-          <li><a href="../sessions/first_order_clause/#Context_Extra">Context_Extra</a></li>
           <li><a href="../sessions/first_order_clause/#Term_Typing">Term_Typing</a></li>
-          <li><a href="../sessions/first_order_clause/#Ground_Term_Extra">Ground_Term_Extra</a></li>
-          <li><a href="../sessions/first_order_clause/#Ground_Context">Ground_Context</a></li>
+          <li><a href="../sessions/first_order_clause/#Generic_Context">Generic_Context</a></li>
+          <li><a href="../sessions/first_order_clause/#Parallel_Induct">Parallel_Induct</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Ground_Term">IsaFoR_Ground_Term</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Ground_Context">IsaFoR_Ground_Context</a></li>
           <li><a href="../sessions/first_order_clause/#Ground_Term_Typing">Ground_Term_Typing</a></li>
           <li><a href="../sessions/first_order_clause/#Nonground_Term">Nonground_Term</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Context">Nonground_Context</a></li>
           <li><a href="../sessions/first_order_clause/#Multiset_Extra">Multiset_Extra</a></li>
           <li><a href="../sessions/first_order_clause/#Multiset_Grounding_Lifting">Multiset_Grounding_Lifting</a></li>
           <li><a href="../sessions/first_order_clause/#Nonground_Clause_Generic">Nonground_Clause_Generic</a></li>
@@ -135,23 +136,15 @@
           <li><a href="../sessions/first_order_clause/#Nonground_Selection_Function">Nonground_Selection_Function</a></li>
           <li><a href="../sessions/first_order_clause/#HOL_Extra">HOL_Extra</a></li>
           <li><a href="../sessions/first_order_clause/#Grounded_Selection_Function">Grounded_Selection_Function</a></li>
-          <li><a href="../sessions/first_order_clause/#Collect_Extra">Collect_Extra</a></li>
-          <li><a href="../sessions/first_order_clause/#Infinite_Variables_Per_Type">Infinite_Variables_Per_Type</a></li>
-          <li><a href="../sessions/first_order_clause/#Typed_Functional_Substitution">Typed_Functional_Substitution</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Term_Typing">Nonground_Term_Typing</a></li>
-          <li><a href="../sessions/first_order_clause/#Monomorphic_Typing">Monomorphic_Typing</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Nonground_Term">IsaFoR_Nonground_Term</a></li>
+          <li><a href="../sessions/first_order_clause/#Literal_Functor">Literal_Functor</a></li>
           <li><a href="../sessions/first_order_clause/#Uprod_Extra">Uprod_Extra</a></li>
-          <li><a href="../sessions/first_order_clause/#Literal_Functor">Literal_Functor</a></li>
           <li><a href="../sessions/first_order_clause/#Uprod_Literal_Functor">Uprod_Literal_Functor</a></li>
-          <li><a href="../sessions/first_order_clause/#Ground_Clause">Ground_Clause</a></li>
           <li><a href="../sessions/first_order_clause/#Nonground_Clause_With_Equality">Nonground_Clause_With_Equality</a></li>
-          <li><a href="../sessions/first_order_clause/#Term_Rewrite_System">Term_Rewrite_System</a></li>
-          <li><a href="../sessions/first_order_clause/#Entailment_Lifting">Entailment_Lifting</a></li>
-          <li><a href="../sessions/first_order_clause/#Fold_Extra">Fold_Extra</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Entailment">Nonground_Entailment</a></li>
-          <li><a href="../sessions/first_order_clause/#Inference_Functor">Inference_Functor</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Inference">Nonground_Inference</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Clause">Nonground_Clause</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Nonground_Clause_With_Equality">IsaFoR_Nonground_Clause_With_Equality</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Context">Nonground_Context</a></li>
+          <li><a href="../sessions/first_order_clause/#Context_Functor">Context_Functor</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Nonground_Context">IsaFoR_Nonground_Context</a></li>
           <li><a href="../sessions/first_order_clause/#Restricted_Order">Restricted_Order</a></li>
           <li><a href="../sessions/first_order_clause/#Context_Compatible_Order">Context_Compatible_Order</a></li>
           <li><a href="../sessions/first_order_clause/#Term_Order_Notation">Term_Order_Notation</a></li>
@@ -166,10 +159,28 @@
           <li><a href="../sessions/first_order_clause/#Term_Order_Lifting">Term_Order_Lifting</a></li>
           <li><a href="../sessions/first_order_clause/#Ground_Order_Generic">Ground_Order_Generic</a></li>
           <li><a href="../sessions/first_order_clause/#Nonground_Order_Generic">Nonground_Order_Generic</a></li>
+          <li><a href="../sessions/first_order_clause/#Ground_Order_With_Equality">Ground_Order_With_Equality</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Order_With_Equality">Nonground_Order_With_Equality</a></li>
+          <li><a href="../sessions/first_order_clause/#Selection_Function_Select_Max">Selection_Function_Select_Max</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_KBO">IsaFoR_KBO</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Clause">Nonground_Clause</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Nonground_Clause">IsaFoR_Nonground_Clause</a></li>
+          <li><a href="../sessions/first_order_clause/#Ground_Term_Rewrite_System">Ground_Term_Rewrite_System</a></li>
+          <li><a href="../sessions/first_order_clause/#Ground_Critical_Pairs">Ground_Critical_Pairs</a></li>
+          <li><a href="../sessions/first_order_clause/#Ground_Critical_Peaks">Ground_Critical_Peaks</a></li>
+          <li><a href="../sessions/first_order_clause/#IsaFoR_Ground_Critical_Pairs">IsaFoR_Ground_Critical_Pairs</a></li>
+          <li><a href="../sessions/first_order_clause/#Collect_Extra">Collect_Extra</a></li>
+          <li><a href="../sessions/first_order_clause/#Infinite_Variables_Per_Type">Infinite_Variables_Per_Type</a></li>
+          <li><a href="../sessions/first_order_clause/#Typed_Functional_Substitution">Typed_Functional_Substitution</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Term_Typing">Nonground_Term_Typing</a></li>
+          <li><a href="../sessions/first_order_clause/#Monomorphic_Typing">Monomorphic_Typing</a></li>
+          <li><a href="../sessions/first_order_clause/#Entailment_Lifting">Entailment_Lifting</a></li>
+          <li><a href="../sessions/first_order_clause/#Fold_Extra">Fold_Extra</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Entailment">Nonground_Entailment</a></li>
+          <li><a href="../sessions/first_order_clause/#Inference_Functor">Inference_Functor</a></li>
+          <li><a href="../sessions/first_order_clause/#Nonground_Inference">Nonground_Inference</a></li>
           <li><a href="../sessions/first_order_clause/#Ground_Order">Ground_Order</a></li>
           <li><a href="../sessions/first_order_clause/#Nonground_Order">Nonground_Order</a></li>
-          <li><a href="../sessions/first_order_clause/#Ground_Order_With_Equality">Ground_Order_With_Equality</a></li>
-          <li><a href="../sessions/first_order_clause/#Nonground_Order_With_Equality">Nonground_Order_With_Equality</a></li>
           <li><a href="../sessions/first_order_clause/#Natural_Magma_Typing_Lifting">Natural_Magma_Typing_Lifting</a></li>
           <li><a href="../sessions/first_order_clause/#Multiset_Typing_Lifting">Multiset_Typing_Lifting</a></li>
           <li><a href="../sessions/first_order_clause/#Clause_Typing_Generic">Clause_Typing_Generic</a></li>
@@ -197,6 +208,7 @@
               <li><a href="../entries/Regular-Sets.html">Regular Sets and Expressions</a></li>
               <li><a href="../entries/Regular_Tree_Relations.html">Regular Tree Relations</a></li>
               <li><a href="../entries/Saturation_Framework_Extensions.html">Extensions to the Comprehensive Framework for Saturation Theorem Proving</a></li>
+              <li><a href="../entries/VeriComp.html">A Generic Framework for Verified Compilers</a></li>
           </ul>
         </div>
         <div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Gabow_SCC.html
--- web/entries/Gabow_SCC.html
+++ web/entries/Gabow_SCC.html
@@ -150,6 +150,7 @@
           <ul class="horizontal-list">
               <li><a href="../entries/CAVA_LTL_Modelchecker.html">A Fully Verified Executable LTL Model Checker</a></li>
               <li><a href="../entries/Transition_Systems_and_Automata.html">Transition Systems and Automata</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/HOL-CSP_OpSem.html
--- web/entries/HOL-CSP_OpSem.html
+++ web/entries/HOL-CSP_OpSem.html
@@ -147,6 +147,12 @@
               <li><a href="../entries/HOL-CSPM.html">HOL-CSPM - Architectural operators for HOL-CSP</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/HOL-CSP_RS.html
--- /dev/null
+++ web/entries/HOL-CSP_RS.html
@@ -0,0 +1,191 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>CSP Semantics over Restriction Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="CSP Semantics over Restriction Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="CSP Semantics over Restriction Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="CSP Semantics over Restriction Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/entries/HOL-CSP_RS.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>C</span><span class='first'>S</span><span class='first'>P</span> <span class='first'>S</span>emantics over <span class='first'>R</span>estriction <span class='first'>S</span>paces
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a> and <a href="../authors/wolff/">Burkhart Wolff</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJ1cmtoYXJ0Iiwid29sZmYiXSwiaG9zdCI6WyJsbWYiLCJjbnJzIiwiZnIiXX0=">📧</a>
+      </p>
+      <p class="date">May 7, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">We use the Restriction_Spaces library as a semantic foundation for the process algebra framework HOL-CSP, offering a complementary backend to the existing HOLCF infrastructure. The type of processes is instantiated as a restriction space, and we prove that it is complete in this setting. This enables the construction of fixed points for recursive process definitions without having to rely exclusively on a pointed complete partial order. Notably, some operators are constructive without being Scott-continuous, and vice versa, illustrating the genuine complementarity between the two approaches. We also show that key CSP operators are either constructive or non-destructive, and verify the admissibility of several predicates, thereby supporting automated reasoning over recursive specifications.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/computer-science/semantics-and-reasoning/">Computer science/Semantics and reasoning</a></li>
+          <li><a href="../topics/computer-science/concurrency/process-calculi/">Computer science/Concurrency/Process calculi</a></li>
+      </ul>
+      <h3>Session HOL-CSP_RS</h3>
+      <ul>
+          <li><a href="../sessions/hol-csp_rs/#Process_Restriction_Space">Process_Restriction_Space</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Prefixes_Constructive">Prefixes_Constructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Choices_Non_Destructive">Choices_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Renaming_Non_Destructive">Renaming_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Sequential_Composition_Non_Destructive">Sequential_Composition_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Synchronization_Product_Non_Destructive">Synchronization_Product_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Throw_Non_Destructive">Throw_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Interrupt_Non_Destructive">Interrupt_Non_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#After_Operator_Non_Too_Destructive">After_Operator_Non_Too_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#Hiding_Destructive">Hiding_Destructive</a></li>
+          <li><a href="../sessions/hol-csp_rs/#CSP_Restriction_Adm">CSP_Restriction_Adm</a></li>
+          <li><a href="../sessions/hol-csp_rs/#HOL-CSP_RS">HOL-CSP_RS</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/HOL-CSP_OpSem.html">Operational Semantics formally proven in HOL-CSP</a></li>
+              <li><a href="../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">HOL-CSP_RS-AFP</p>
+        <pre id="copy-text">@article{HOL-CSP_RS-AFP,
+  author  = {Benoît Ballenghien and Burkhart Wolff},
+  title   = {CSP Semantics over Restriction Spaces},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/HOL-CSP_RS.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-HOL-CSP_RS-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/LTL_Master_Theorem.html
--- web/entries/LTL_Master_Theorem.html
+++ web/entries/LTL_Master_Theorem.html
@@ -161,6 +161,7 @@
           <h3>Used by</h3>
           <ul class="horizontal-list">
               <li><a href="../entries/LTL_Normal_Form.html">An Efficient Normalisation Procedure for Linear Temporal Logic: Isabelle/HOL Formalisation</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Lambert_Series.html
--- web/entries/Lambert_Series.html
+++ web/entries/Lambert_Series.html
@@ -153,6 +153,7 @@
           <ul class="horizontal-list">
               <li><a href="../entries/Combinatorial_Q_Analogues.html">Combinatorial q-Analogues</a></li>
               <li><a href="../entries/Theta_Functions.html">Theta Functions</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/List_Power.html
--- web/entries/List_Power.html
+++ web/entries/List_Power.html
@@ -136,6 +136,7 @@
           <h3>Used by</h3>
           <ul class="horizontal-list">
               <li><a href="../entries/Combinatorics_Words.html">Combinatorics on Words Basics</a></li>
+              <li><a href="../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Munta_Model_Checker.html
--- /dev/null
+++ web/entries/Munta_Model_Checker.html
@@ -0,0 +1,247 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Munta: A Verified Model Checker for Timed Automata - Archive of Formal Proofs</title>
+    <meta name="description" content="Munta: A Verified Model Checker for Timed Automata in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Munta: A Verified Model Checker for Timed Automata in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Munta: A Verified Model Checker for Timed Automata" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Munta_Model_Checker.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>M</span>unta: <span class='first'>A</span> <span class='first'>V</span>erified <span class='first'>M</span>odel <span class='first'>C</span>hecker for <span class='first'>T</span>imed <span class='first'>A</span>utomata
+  </h1>
+  <div>
+      <p><a href="../authors/wimmer/">Simon Wimmer</a> <a class="obfuscated" data="eyJ1c2VyIjpbInNpbW9uIiwid2ltbWVyIl0sImhvc3QiOlsidHVtIiwiZGUiXX0=">📧</a>
+      </p>
+      <p class="date">May 22, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">Munta is a verified model checker for timed automata. It has been described in detail in a PhD thesis [3] and conference publications [4, 2]. This entry supersedes earlier versions of Munta hosted on GitHub.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/computer-science/automata-and-formal-languages/">Computer science/Automata and formal languages</a></li>
+          <li><a href="../topics/tools/">Tools</a></li>
+      </ul>
+      <h3>Related publications</h3>
+      <ul>
+        <li>Wimmer, S., & Lammich, P. (2018). Verified Model Checking of Timed Automata. Tools and Algorithms for the Construction and Analysis of Systems, 61–78. <a href="https://doi.org/10.1007/978-3-319-89960-2_4">https://doi.org/10.1007/978-3-319-89960-2_4</a>
+</li>
+        <li>Wimmer, S. (2019). Munta: A Verified Model Checker for Timed Automata. Formal Modeling and Analysis of Timed Systems, 236–243. <a href="https://doi.org/10.1007/978-3-030-29662-9_14">https://doi.org/10.1007/978-3-030-29662-9_14</a>
+</li>
+        <li>https://nbn-resolving.org/urn:nbn:de:bvb:91-diss-20201209-1576100-1-0</li>
+      </ul>
+      <h3>Session Munta_Base</h3>
+      <ul>
+          <li><a href="../sessions/munta_base/#TA_Syntax_Bundles">TA_Syntax_Bundles</a></li>
+          <li><a href="../sessions/munta_base/#ML_Util">ML_Util</a></li>
+          <li><a href="../sessions/munta_base/#Reordering_Quantifiers">Reordering_Quantifiers</a></li>
+          <li><a href="../sessions/munta_base/#More_Methods">More_Methods</a></li>
+          <li><a href="../sessions/munta_base/#Abstract_Term">Abstract_Term</a></li>
+          <li><a href="../sessions/munta_base/#Bijective_Embedding">Bijective_Embedding</a></li>
+          <li><a href="../sessions/munta_base/#Tracing">Tracing</a></li>
+          <li><a href="../sessions/munta_base/#Printing">Printing</a></li>
+          <li><a href="../sessions/munta_base/#Trace_Timing">Trace_Timing</a></li>
+          <li><a href="../sessions/munta_base/#Error_List_Monad">Error_List_Monad</a></li>
+          <li><a href="../sessions/munta_base/#Temporal_Logics">Temporal_Logics</a></li>
+          <li><a href="../sessions/munta_base/#Subsumption_Graphs">Subsumption_Graphs</a></li>
+          <li><a href="../sessions/munta_base/#Sepref_Acconstraint">Sepref_Acconstraint</a></li>
+          <li><a href="../sessions/munta_base/#TA_DBM_Operations_Impl">TA_DBM_Operations_Impl</a></li>
+          <li><a href="../sessions/munta_base/#TA_More">TA_More</a></li>
+          <li><a href="../sessions/munta_base/#Normalized_Zone_Semantics_Impl">Normalized_Zone_Semantics_Impl</a></li>
+          <li><a href="../sessions/munta_base/#TA_Impl_Misc">TA_Impl_Misc</a></li>
+          <li><a href="../sessions/munta_base/#Normalized_Zone_Semantics_Impl_Semantic_Refinement">Normalized_Zone_Semantics_Impl_Semantic_Refinement</a></li>
+          <li><a href="../sessions/munta_base/#Normalized_Zone_Semantics_Impl_Refine">Normalized_Zone_Semantics_Impl_Refine</a></li>
+          <li><a href="../sessions/munta_base/#Normalized_Zone_Semantics_Impl_Extra">Normalized_Zone_Semantics_Impl_Extra</a></li>
+      </ul>
+      <h3>Session Munta_Model_Checker</h3>
+      <ul>
+          <li><a href="../sessions/munta_model_checker/#Munta_Tagging">Munta_Tagging</a></li>
+          <li><a href="../sessions/munta_model_checker/#Munta_Error_Monad_Add">Munta_Error_Monad_Add</a></li>
+          <li><a href="../sessions/munta_model_checker/#Parser_Combinator">Parser_Combinator</a></li>
+          <li><a href="../sessions/munta_model_checker/#Lexer">Lexer</a></li>
+          <li><a href="../sessions/munta_model_checker/#JSON_Parsing">JSON_Parsing</a></li>
+          <li><a href="../sessions/munta_model_checker/#Networks">Networks</a></li>
+          <li><a href="../sessions/munta_model_checker/#State_Networks">State_Networks</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_Asm">UPPAAL_Asm</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_Asm_Clocks">UPPAAL_Asm_Clocks</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_State_Networks">UPPAAL_State_Networks</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_State_Networks_Impl">UPPAAL_State_Networks_Impl</a></li>
+          <li><a href="../sessions/munta_model_checker/#Program_Analysis">Program_Analysis</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_State_Networks_Impl_Refine">UPPAAL_State_Networks_Impl_Refine</a></li>
+          <li><a href="../sessions/munta_model_checker/#UPPAAL_Model_Checking">UPPAAL_Model_Checking</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Expressions">Simple_Expressions</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language">Simple_Network_Language</a></li>
+          <li><a href="../sessions/munta_model_checker/#TA_Impl_Misc2">TA_Impl_Misc2</a></li>
+          <li><a href="../sessions/munta_model_checker/#TA_More2">TA_More2</a></li>
+          <li><a href="../sessions/munta_model_checker/#TA_Equivalences">TA_Equivalences</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Impl">Simple_Network_Language_Impl</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Impl_Refine">Simple_Network_Language_Impl_Refine</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Model_Checking">Simple_Network_Language_Model_Checking</a></li>
+          <li><a href="../sessions/munta_model_checker/#Deadlock">Deadlock</a></li>
+          <li><a href="../sessions/munta_model_checker/#Deadlock_Impl">Deadlock_Impl</a></li>
+          <li><a href="../sessions/munta_model_checker/#Deadlock_Checking">Deadlock_Checking</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Renaming">Simple_Network_Language_Renaming</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Deadlock_Checking">Simple_Network_Language_Deadlock_Checking</a></li>
+          <li><a href="../sessions/munta_model_checker/#Shortest_SCC_Paths">Shortest_SCC_Paths</a></li>
+          <li><a href="../sessions/munta_model_checker/#Simple_Network_Language_Export_Code">Simple_Network_Language_Export_Code</a></li>
+          <li><a href="../sessions/munta_model_checker/#Munta_Compile_MLton">Munta_Compile_MLton</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/AutoCorres2.html">AutoCorres2</a></li>
+              <li><a href="../entries/Certification_Monads.html">Certification Monads</a></li>
+              <li><a href="../entries/FinFun.html">Code Generation for Functions as Data</a></li>
+              <li><a href="../entries/Gabow_SCC.html">Verified Efficient Implementation of Gabow's Strongly Connected Components Algorithm</a></li>
+              <li><a href="../entries/LTL_Master_Theorem.html">A Compositional and Unified Translation of LTL into ω-Automata</a></li>
+              <li><a href="../entries/Show.html">Haskell's Show Class in Isabelle/HOL</a></li>
+              <li><a href="../entries/Timed_Automata.html">Timed Automata</a></li>
+              <li><a href="../entries/Transition_Systems_and_Automata.html">Transition Systems and Automata</a></li>
+              <li><a href="../entries/Worklist_Algorithms.html">Worklist Algorithms</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Munta_Model_Checker-AFP</p>
+        <pre id="copy-text">@article{Munta_Model_Checker-AFP,
+  author  = {Simon Wimmer},
+  title   = {Munta: A Verified Model Checker for Timed Automata},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Munta_Model_Checker.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Munta_Model_Checker-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Optics.html
--- web/entries/Optics.html
+++ web/entries/Optics.html
@@ -184,6 +184,7 @@
           <ul class="horizontal-list">
               <li><a href="../entries/UTP.html">Isabelle/UTP: Mechanised Theory Engineering for Unifying Theories of Programming</a></li>
               <li><a href="../entries/ConcurrentHOL.html">ConcurrentHOL</a></li>
+              <li><a href="../entries/Shallow_Expressions.html">Shallow Expressions</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Pentagonal_Number_Theorem.html
--- web/entries/Pentagonal_Number_Theorem.html
+++ web/entries/Pentagonal_Number_Theorem.html
@@ -148,6 +148,12 @@
               <li><a href="../entries/Theta_Functions.html">Theta Functions</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Polylog.html
--- web/entries/Polylog.html
+++ web/entries/Polylog.html
@@ -143,6 +143,7 @@
           <h3>Used by</h3>
           <ul class="horizontal-list">
               <li><a href="../entries/Lambert_Series.html">Lambert Series</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Polynomial_Factorization.html
--- web/entries/Polynomial_Factorization.html
+++ web/entries/Polynomial_Factorization.html
@@ -187,6 +187,7 @@
               <li><a href="../entries/Amicable_Numbers.html">Amicable Numbers</a></li>
               <li><a href="../entries/Fishers_Inequality.html">Fisher's Inequality: Linear Algebraic Proof Techniques for Combinatorics</a></li>
               <li><a href="../entries/Continued_Fractions.html">Continued Fractions</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Restriction_Spaces-Examples.html
--- /dev/null
+++ web/entries/Restriction_Spaces-Examples.html
@@ -0,0 +1,204 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Examples of Restriction Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="Examples of Restriction Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Examples of Restriction Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Examples of Restriction Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Restriction_Spaces-Examples.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>E</span>xamples of <span class='first'>R</span>estriction <span class='first'>S</span>paces
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a>, <a href="../authors/puyobro/">Benjamin Puyobro</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbmphbWluIiwicHV5b2JybyJdLCJob3N0IjpbImxhcG9zdGUiLCJuZXQiXX0=">📧</a> and <a href="../authors/wolff/">Burkhart Wolff</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJ1cmtoYXJ0Iiwid29sZmYiXSwiaG9zdCI6WyJsbWYiLCJjbnJzIiwiZnIiXX0=">📧</a>
+      </p>
+      <p class="date">May 7, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">In this session, a number of examples are provided to illustrate how the Restriction_Spaces library works.
+The simple cases are, of course, covered: trivial restriction, booleans, integers, option type, and so on.
+But we also explore more elaborate constructions, such as formal power series and a trace model of the CSP process algebra.
+Additionally, we provide a lightweight integration with HOLCF,
+equipping restriction spaces with the inherited partial order structure when needed.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/mathematics/analysis/">Mathematics/Analysis</a></li>
+          <li><a href="../topics/mathematics/topology/">Mathematics/Topology</a></li>
+          <li><a href="../topics/mathematics/algebra/">Mathematics/Algebra</a></li>
+          <li><a href="../topics/computer-science/concurrency/process-calculi/">Computer science/Concurrency/Process calculi</a></li>
+      </ul>
+      <h3>Session Restriction_Spaces-Examples</h3>
+      <ul>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Any_Type">RS_Any_Type</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Bool">RS_Bool</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Nat">RS_Nat</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Int">RS_Int</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Option">RS_Option</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_List">RS_List</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Tree">RS_Tree</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Decimals">RS_Decimals</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Trace_Model_CSP">RS_Trace_Model_CSP</a></li>
+          <li><a href="../sessions/restriction_spaces-examples/#RS_Formal_Power_Series">RS_Formal_Power_Series</a></li>
+      </ul>
+      <h3>Session Restriction_Spaces-HOLCF</h3>
+      <ul>
+          <li><a href="../sessions/restriction_spaces-holcf/#Restriction_Spaces-HOLCF">Restriction_Spaces-HOLCF</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></li>
+          </ul>
+        </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Restriction_Spaces-Examples-AFP</p>
+        <pre id="copy-text">@article{Restriction_Spaces-Examples-AFP,
+  author  = {Benoît Ballenghien and Benjamin Puyobro and Burkhart Wolff},
+  title   = {Examples of Restriction Spaces},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Restriction_Spaces-Examples.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Restriction_Spaces-Examples-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Restriction_Spaces-Ultrametric.html
--- /dev/null
+++ web/entries/Restriction_Spaces-Ultrametric.html
@@ -0,0 +1,196 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Ultrametric Structure for Restriction Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="Ultrametric Structure for Restriction Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Ultrametric Structure for Restriction Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Ultrametric Structure for Restriction Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>U</span>ltrametric <span class='first'>S</span>tructure for <span class='first'>R</span>estriction <span class='first'>S</span>paces
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a>, <a href="../authors/puyobro/">Benjamin Puyobro</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbmphbWluIiwicHV5b2JybyJdLCJob3N0IjpbImxhcG9zdGUiLCJuZXQiXX0=">📧</a> and <a href="../authors/wolff/">Burkhart Wolff</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJ1cmtoYXJ0Iiwid29sZmYiXSwiaG9zdCI6WyJsbWYiLCJjbnJzIiwiZnIiXX0=">📧</a>
+      </p>
+      <p class="date">May 7, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">In this entry, we explore the relationship
+  between restriction spaces and usual metric
+  structures by instantiating the former as ultrametric spaces.
+  This is classically captured by defining the distance as
+  $$\mathrm{dist} \ x \ y = \inf_{x \downarrow n \ = \ y \downarrow n} \left(\frac{1}{2}\right)^ n$$
+  but we actually generalize this perspective by introducing a hierarchy of increasingly
+  refined type classes to systematically relate ultrametric and restriction-based notions.
+  This layered approach enables a precise comparison of structural and topological properties.
+  In the end, our main result establishes that completeness in the sense of restriction spaces
+  coincides with standard metric completeness, thus bridging the gap between
+  our framework and Banach's fixed-point theorem established in HOL-Analysis.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/mathematics/analysis/">Mathematics/Analysis</a></li>
+          <li><a href="../topics/mathematics/topology/">Mathematics/Topology</a></li>
+          <li><a href="../topics/computer-science/semantics-and-reasoning/">Computer science/Semantics and reasoning</a></li>
+      </ul>
+      <h3>Session Restriction_Spaces-Ultrametric</h3>
+      <ul>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Banach_Theorem_Extension">Banach_Theorem_Extension</a></li>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Ultrametric_Restriction_Spaces_Locale">Ultrametric_Restriction_Spaces_Locale</a></li>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Ultrametric_Restriction_Spaces">Ultrametric_Restriction_Spaces</a></li>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Fun_Ultrametric_Restriction_Spaces">Fun_Ultrametric_Restriction_Spaces</a></li>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Prod_Ultrametric_Restriction_Spaces">Prod_Ultrametric_Restriction_Spaces</a></li>
+          <li><a href="../sessions/restriction_spaces-ultrametric/#Restriction_Spaces-Ultrametric">Restriction_Spaces-Ultrametric</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a></li>
+              <li><a href="../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Restriction_Spaces-Ultrametric-AFP</p>
+        <pre id="copy-text">@article{Restriction_Spaces-Ultrametric-AFP,
+  author  = {Benoît Ballenghien and Benjamin Puyobro and Burkhart Wolff},
+  title   = {Ultrametric Structure for Restriction Spaces},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Restriction_Spaces-Ultrametric-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Restriction_Spaces.html
--- /dev/null
+++ web/entries/Restriction_Spaces.html
@@ -0,0 +1,208 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Restriction Spaces: a Fixed-Point Theory - Archive of Formal Proofs</title>
+    <meta name="description" content="Restriction Spaces: a Fixed-Point Theory in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Restriction Spaces: a Fixed-Point Theory in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Restriction Spaces: a Fixed-Point Theory" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Restriction_Spaces.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>R</span>estriction <span class='first'>S</span>paces: a <span class='first'>F</span>ixed-<span class='first'>P</span>oint <span class='first'>T</span>heory
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a>, <a href="../authors/puyobro/">Benjamin Puyobro</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbmphbWluIiwicHV5b2JybyJdLCJob3N0IjpbImxhcG9zdGUiLCJuZXQiXX0=">📧</a> and <a href="../authors/wolff/">Burkhart Wolff</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJ1cmtoYXJ0Iiwid29sZmYiXSwiaG9zdCI6WyJsbWYiLCJjbnJzIiwiZnIiXX0=">📧</a>
+      </p>
+      <p class="date">May 7, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">Fixed-point constructions are fundamental to defining recursive and co-recursive functions.
+  However, a general axiom $Y f = f (Y f)$ leads to inconsistency, and definitions
+  must therefore be based on theories guaranteeing existence under suitable conditions.
+  In Isabelle/HOL, these constructions are typically based on sets, well-founded orders
+  or domain-theoretic models such as for example HOLCF.
+  In this submission we introduce a formalization of restriction spaces i.e. spaces
+  equipped with a so-called restriction, denoted by $\downarrow$,
+  satifying three properties:
+  \begin{align*}
+    & x \downarrow 0 = y \downarrow 0\\
+    & x \downarrow n \downarrow m = x \downarrow \mathrm{min} \ n \ m\\
+    & x \not = y \Longrightarrow \exists n. \ x \downarrow n \not = y \downarrow n
+  \end{align*}
+  They turn out to be cartesian closed and admit natural notions of constructiveness and
+  completeness, enabling the definition of a fixed-point operator under verifiable side-conditions.
+  This is achieved in our entry, from topological definitions to induction principles.
+  Additionally, we configure the simplifier so that it can automatically solve both
+  constructiveness and admissibility subgoals, as long as users write higher-order rules
+  for their operators.
+  Since our implementation relies on axiomatic type classes, the resulting library
+  is a fully abstract, flexible and reusable framework.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/tools/">Tools</a></li>
+          <li><a href="../topics/mathematics/analysis/">Mathematics/Analysis</a></li>
+          <li><a href="../topics/mathematics/topology/">Mathematics/Topology</a></li>
+          <li><a href="../topics/computer-science/semantics-and-reasoning/">Computer science/Semantics and reasoning</a></li>
+      </ul>
+      <h3>Session Restriction_Spaces</h3>
+      <ul>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Locales">Restriction_Spaces_Locales</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Classes">Restriction_Spaces_Classes</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Prod">Restriction_Spaces_Prod</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Fun">Restriction_Spaces_Fun</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Topology">Restriction_Spaces_Topology</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces_Induction">Restriction_Spaces_Induction</a></li>
+          <li><a href="../sessions/restriction_spaces/#Restriction_Spaces">Restriction_Spaces</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></li>
+              <li><a href="../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Restriction_Spaces-AFP</p>
+        <pre id="copy-text">@article{Restriction_Spaces-AFP,
+  author  = {Benoît Ballenghien and Benjamin Puyobro and Burkhart Wolff},
+  title   = {Restriction Spaces: a Fixed-Point Theory},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Restriction_Spaces.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Restriction_Spaces-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Shallow_Expressions.html
--- /dev/null
+++ web/entries/Shallow_Expressions.html
@@ -0,0 +1,198 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Shallow Expressions - Archive of Formal Proofs</title>
+    <meta name="description" content="Shallow Expressions in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Shallow Expressions in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Shallow Expressions" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Shallow_Expressions.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>S</span>hallow <span class='first'>E</span>xpressions
+  </h1>
+  <div>
+      <p><a href="../authors/fosters/">Simon Foster</a> <a class="obfuscated" data="eyJ1c2VyIjpbInNpbW9uIiwiZm9zdGVyIl0sImhvc3QiOlsieW9yayIsImFjIiwidWsiXX0=">📧</a>
+      </p>
+      <p class="date">May 16, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">Most verification techniques use expressions, for example when assigning to variables or forming assertions over the state. Deep embeddings provide such expressions using a datatype, which allows queries over the syntax, such as calculating the free variables, and performing substitutions. Shallow embeddings, in contrast, model expressions as query functions over the state type, and are more amenable to automating verification tasks. The goal of this library is provide an intuitive implementation of shallow expressions, which nevertheless provides many of the benefits of a deep embedding. We harness the Optics library to provide an algebraic semantics for state variables, and use syntax translations to provide an intuitive lifted expression syntax. Furthermore, we provide a variety of meta-logic-style queries on expressions, such as dependencies on a state variable, and substitution of a variable for an expression. We also provide proof methods, based on the simplifier, to automate the associated proof tasks.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/computer-science/semantics-and-reasoning/">Computer science/Semantics and reasoning</a></li>
+          <li><a href="../topics/computer-science/programming-languages/logics/">Computer science/Programming languages/Logics</a></li>
+      </ul>
+      <h3>Related publications</h3>
+      <ul>
+        <li>Huerta y Munive, J. J., Foster, S., Gleirscher, M., Struth, G., Pardillo Laursen, C., & Hickman, T. (2024). IsaVODEs: Interactive Verification of Cyber-Physical Systems at Scale. Journal of Automated Reasoning, 68(4). <a href="https://doi.org/10.1007/s10817-024-09709-2">https://doi.org/10.1007/s10817-024-09709-2</a>
+</li>
+        <li>Foster, S., Baxter, J., Cavalcanti, A., Woodcock, J., & Zeyda, F. (2020). Unifying semantic foundations for automated verification tools in Isabelle/UTP. Science of Computer Programming, 197, 102510. <a href="https://doi.org/10.1016/j.scico.2020.102510">https://doi.org/10.1016/j.scico.2020.102510</a>
+</li>
+      </ul>
+      <h3>Session Shallow_Expressions</h3>
+      <ul>
+          <li><a href="../sessions/shallow_expressions/#Variables">Variables</a></li>
+          <li><a href="../sessions/shallow_expressions/#Expressions">Expressions</a></li>
+          <li><a href="../sessions/shallow_expressions/#Unrestriction">Unrestriction</a></li>
+          <li><a href="../sessions/shallow_expressions/#Substitutions">Substitutions</a></li>
+          <li><a href="../sessions/shallow_expressions/#Extension">Extension</a></li>
+          <li><a href="../sessions/shallow_expressions/#Liberation">Liberation</a></li>
+          <li><a href="../sessions/shallow_expressions/#Quantifiers">Quantifiers</a></li>
+          <li><a href="../sessions/shallow_expressions/#Collections">Collections</a></li>
+          <li><a href="../sessions/shallow_expressions/#Named_Expressions">Named_Expressions</a></li>
+          <li><a href="../sessions/shallow_expressions/#Local_State">Local_State</a></li>
+          <li><a href="../sessions/shallow_expressions/#Shallow_Expressions">Shallow_Expressions</a></li>
+          <li><a href="../sessions/shallow_expressions/#Expressions_Tests">Expressions_Tests</a></li>
+          <li><a href="../sessions/shallow_expressions/#Shallow_Expressions_Examples">Shallow_Expressions_Examples</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Optics.html">Optics</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Shallow_Expressions-AFP</p>
+        <pre id="copy-text">@article{Shallow_Expressions-AFP,
+  author  = {Simon Foster},
+  title   = {Shallow Expressions},
+  journal = {Archive of Formal Proofs},
+  month   = {May},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Shallow_Expressions.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Shallow_Expressions-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Show.html
--- web/entries/Show.html
+++ web/entries/Show.html
@@ -181,6 +181,7 @@
               <li><a href="../entries/LL1_Parser.html">LL(1) Parser Generator</a></li>
               <li><a href="../entries/Difference_Bound_Matrices.html">Difference Bound Matrices</a></li>
               <li><a href="../entries/Mission_Time_LTL_Language_Partition.html">Language Partitioning for Mission-time Linear Temporal Logic</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Sophie_Germain.html
--- /dev/null
+++ web/entries/Sophie_Germain.html
@@ -0,0 +1,186 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Sophie Germain’s Theorem - Archive of Formal Proofs</title>
+    <meta name="description" content="Sophie Germain’s Theorem in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Sophie Germain’s Theorem in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Sophie Germain’s Theorem" />
+  <meta property="og:url" content="https://isa-afp.org/entries/Sophie_Germain.html" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="article" />
+  <link rel="stylesheet" type="text/css" href="../css/front.min.css">
+
+  <link rel="icon" href="../images/favicon.ico" type="image/icon">
+    
+    <script>
+      MathJax = {
+        tex: {
+          inlineMath: [["$", "$"], ["\\(", "\\)"]]
+        },
+        processEscapes: true,
+        svg: {
+          fontCache: "global"
+        }
+      };
+    </script>
+    <script id="MathJax-script" async src="../js/mathjax/es5/tex-mml-chtml.js">
+    </script>
+    <script src="../js/entries.js"></script>
+
+  <script src="../js/obfuscate.js"></script>
+  <script src="../js/flexsearch.bundle.js"></script>
+  <script src="../js/scroll-spy.js"></script>
+  <script src="../js/theory.js"></script>
+  <script src="../js/util.js"></script>
+    <script src="../js/header-search.js"></script>
+  <script src="../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../" class="logo-link">
+    <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../" class="logo-link">
+        <img src="../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+      <ul>
+          <a href="../">
+            <li >Home</li>
+          </a>
+          <a href="../topics/">
+            <li >Topics</li>
+          </a>
+          <a href="../download/">
+            <li >Download</li>
+          </a>
+          <a href="../help/">
+            <li >Help</li>
+          </a>
+          <a href="../submission/">
+            <li >Submission</li>
+          </a>
+          <a href="../statistics/">
+            <li >Statistics</li>
+          </a>
+          <a href="../about/">
+            <li >About</li>
+          </a>
+      </ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content entries"><header>
+    <form autocomplete="off" action="../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>S</span>ophie <span class='first'>G</span>ermain’s <span class='first'>T</span>heorem
+  </h1>
+  <div>
+      <p><a href="../authors/ballenghien/">Benoît Ballenghien</a> <a class="obfuscated" data="eyJ1c2VyIjpbImJlbm9pdCIsImJhbGxlbmdoaWVuIl0sImhvc3QiOlsidW5pdmVyc2l0ZS1wYXJpcy1zYWNsYXkiLCJmciJdfQ==">📧</a>
+      </p>
+      <p class="date">April 9, 2025</p>
+  </div>
+</header>
+      <div>
+  <main>
+
+    <h3>Abstract</h3>
+    <div class="abstract mathjax_process">Sophie Germain's theorem states that if $p$ is a prime number such that $2 p + 1$ is also prime (with $p$ then called a Sophie Germain prime), a weak version of Fermat’s Last Theorem (FLT) holds for the exponent $p$. That is, there exist no nontrivial integer solutions $x, y, z \in \mathbb{Z}$ to the equation
+$$x ^ p + y ^ p = z ^ p$$
+such that $p$ divides neither $x$, $y$ nor $z$.
+After some preliminary results, mainly on sufficient criteria for FLT, we propose a formalization of the classical version of Sophie Germain's theorem as well as a variant involving the concept of auxiliary primes that generalizes Sophie Germain primes.</div>
+
+    <h3>License</h3>
+    <div>
+        <a href="https://isa-afp.org/LICENSE">BSD License</a>
+    </div>
+      <h3>Topics</h3>
+      <ul>
+          <li><a href="../topics/mathematics/number-theory/">Mathematics/Number theory</a></li>
+      </ul>
+      <h3>Session Sophie_Germain</h3>
+      <ul>
+          <li><a href="../sessions/sophie_germain/#SG_Introduction">SG_Introduction</a></li>
+          <li><a href="../sessions/sophie_germain/#SG_Preliminaries">SG_Preliminaries</a></li>
+          <li><a href="../sessions/sophie_germain/#FLT_Sufficient_Conditions">FLT_Sufficient_Conditions</a></li>
+          <li><a href="../sessions/sophie_germain/#SG_Theorem">SG_Theorem</a></li>
+          <li><a href="../sessions/sophie_germain/#SG_Generalization">SG_Generalization</a></li>
+      </ul>
+
+    <div class="flex-wrap">
+        <div>
+          <h3>Depends on</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Fermat3_4.html">Fermat's Last Theorem for Exponents 3 and 4 and the Parametrisation of Pythagorean Triples</a></li>
+              <li><a href="../entries/Wlog.html">Without Loss of Generality</a></li>
+          </ul>
+        </div>
+    </div>
+  </main>
+
+  <nav class="links">
+    <a class="popup-button" href="#cite-popup">Cite</a>
+    <a class="popup-button" href="#download-popup">Download</a>
+    <h4>PDFs</h4>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/outline.pdf">Proof outline</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/document.pdf">Proof document</a>
+    <a href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/session_graph.pdf">Dependencies</a>
+  </nav>
+
+  <div id="cite-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Cite</h2>
+      <a class="close" href="#">&times;</a>
+      <div>
+        <p style="display:none;" id="bibtex-filename">Sophie_Germain-AFP</p>
+        <pre id="copy-text">@article{Sophie_Germain-AFP,
+  author  = {Benoît Ballenghien},
+  title   = {Sophie Germain’s Theorem},
+  journal = {Archive of Formal Proofs},
+  month   = {April},
+  year    = {2025},
+  note    = {\url{https://isa-afp.org/entries/Sophie_Germain.html},
+             Formal proof development},
+  ISSN    = {2150-914x},
+}</pre>
+        <button id="copy-bibtex">Copy</button> <a id="download-bibtex">Download</a>
+      </div>
+    </div>
+  </div>
+
+  <div id="download-popup" class="overlay">
+    <a class="cancel" href="#"></a>
+    <div class="popup">
+      <h2>Download</h2>
+      <a class="close" href="#">&times;</a>
+      <a href="https://www.isa-afp.org/release/afp-Sophie_Germain-current.tar.gz" download>
+        Download latest</a>
+    </div>
+  </div>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Sum_Of_Squares_Count.html
--- web/entries/Sum_Of_Squares_Count.html
+++ web/entries/Sum_Of_Squares_Count.html
@@ -115,7 +115,7 @@
 </p>
 <p>
 Using this, I then formalise the main result: Jacobi's two-square theorem, which states that for $n > 0$
-we have \[r_2(n) = 4(d_1(3) - d_3(n))\ ,\] where $d_i(n)$ denotes the number of divisors $m$ of $n$
+we have \[r_2(n) = 4(d_1(n) - d_3(n))\ ,\] where $d_i(n)$ denotes the number of divisors $m$ of $n$
 such that $m = i\ (\text{mod}\ 4)$.
 </p>
 <p>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Superposition_Calculus.html
--- web/entries/Superposition_Calculus.html
+++ web/entries/Superposition_Calculus.html
@@ -123,19 +123,19 @@
       </ul>
       <h3>Session Superposition_Calculus</h3>
       <ul>
-          <li><a href="../sessions/superposition_calculus/#Ground_Critical_Pairs">Ground_Critical_Pairs</a></li>
           <li><a href="../sessions/superposition_calculus/#Ground_Superposition">Ground_Superposition</a></li>
           <li><a href="../sessions/superposition_calculus/#Abstract_Rewriting_Extra">Abstract_Rewriting_Extra</a></li>
           <li><a href="../sessions/superposition_calculus/#Relation_Extra">Relation_Extra</a></li>
           <li><a href="../sessions/superposition_calculus/#Ground_Superposition_Completeness">Ground_Superposition_Completeness</a></li>
           <li><a href="../sessions/superposition_calculus/#Ground_Superposition_Soundness">Ground_Superposition_Soundness</a></li>
-          <li><a href="../sessions/superposition_calculus/#Ground_Superposition_Welltypedness_Preservation">Ground_Superposition_Welltypedness_Preservation</a></li>
+          <li><a href="../sessions/superposition_calculus/#Typed_Ground_Superposition">Typed_Ground_Superposition</a></li>
           <li><a href="../sessions/superposition_calculus/#Superposition">Superposition</a></li>
+          <li><a href="../sessions/superposition_calculus/#Monomorphic_Superposition">Monomorphic_Superposition</a></li>
+          <li><a href="../sessions/superposition_calculus/#Ground_Superposition_Example">Ground_Superposition_Example</a></li>
           <li><a href="../sessions/superposition_calculus/#Grounded_Superposition">Grounded_Superposition</a></li>
           <li><a href="../sessions/superposition_calculus/#Superposition_Welltypedness_Preservation">Superposition_Welltypedness_Preservation</a></li>
           <li><a href="../sessions/superposition_calculus/#Superposition_Completeness">Superposition_Completeness</a></li>
           <li><a href="../sessions/superposition_calculus/#Superposition_Soundness">Superposition_Soundness</a></li>
-          <li><a href="../sessions/superposition_calculus/#IsaFoR_Term_Copy">IsaFoR_Term_Copy</a></li>
           <li><a href="../sessions/superposition_calculus/#Superposition_Example">Superposition_Example</a></li>
           <li><a href="../sessions/superposition_calculus/#Untyped_Superposition">Untyped_Superposition</a></li>
           <li><a href="../sessions/superposition_calculus/#Grounded_Untyped_Superposition">Grounded_Untyped_Superposition</a></li>
@@ -150,7 +150,6 @@
               <li><a href="../entries/First_Order_Clause.html">First Order Clause</a></li>
               <li><a href="../entries/Saturation_Framework.html">A Comprehensive Framework for Saturation Theorem Proving</a></li>
               <li><a href="../entries/Saturation_Framework_Extensions.html">Extensions to the Comprehensive Framework for Saturation Theorem Proving</a></li>
-              <li><a href="../entries/VeriComp.html">A Generic Framework for Verified Compilers</a></li>
           </ul>
         </div>
         <div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Theta_Functions.html
--- web/entries/Theta_Functions.html
+++ web/entries/Theta_Functions.html
@@ -155,6 +155,7 @@
           <ul class="horizontal-list">
               <li><a href="../entries/Rogers_Ramanujan.html">The Rogers–Ramanujan Identities</a></li>
               <li><a href="../entries/Pentagonal_Number_Theorem.html">The Partition Function and the Pentagonal Number Theorem</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Three_Squares.html
--- web/entries/Three_Squares.html
+++ web/entries/Three_Squares.html
@@ -98,7 +98,7 @@
   <h1 ><span class='first'>T</span>hree <span class='first'>S</span>quares <span class='first'>T</span>heorem
   </h1>
   <div>
-      <p><a href="../authors/danilkin/">Anton Danilkin</a> and <a href="../authors/chevalier/">Loïc Chevalier</a>
+      <p><a href="../authors/danilkin/">Anna Danilkin</a> and <a href="../authors/chevalier/">Loïc Chevalier</a>
       </p>
       <p class="date">May 3, 2023</p>
   </div>
@@ -184,7 +184,7 @@
       <div>
         <p style="display:none;" id="bibtex-filename">Three_Squares-AFP</p>
         <pre id="copy-text">@article{Three_Squares-AFP,
-  author  = {Anton Danilkin and Loïc Chevalier},
+  author  = {Anna Danilkin and Loïc Chevalier},
   title   = {Three Squares Theorem},
   journal = {Archive of Formal Proofs},
   month   = {May},
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Timed_Automata.html
--- web/entries/Timed_Automata.html
+++ web/entries/Timed_Automata.html
@@ -176,6 +176,7 @@
           <h3>Used by</h3>
           <ul class="horizontal-list">
               <li><a href="../entries/Probabilistic_Timed_Automata.html">Probabilistic Timed Automata</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Transition_Systems_and_Automata.html
--- web/entries/Transition_Systems_and_Automata.html
+++ web/entries/Transition_Systems_and_Automata.html
@@ -198,6 +198,7 @@
               <li><a href="../entries/Partial_Order_Reduction.html">Partial Order Reduction</a></li>
               <li><a href="../entries/LTL_Master_Theorem.html">A Compositional and Unified Translation of LTL into ω-Automata</a></li>
               <li><a href="../entries/Adaptive_State_Counting.html">Formalisation of an Adaptive State Counting Algorithm</a></li>
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/VeriComp.html
--- web/entries/VeriComp.html
+++ web/entries/VeriComp.html
@@ -149,8 +149,8 @@
           <h3>Used by</h3>
           <ul class="horizontal-list">
               <li><a href="../entries/Interpreter_Optimizations.html">Inline Caching and Unboxing Optimization for Interpreters</a></li>
-              <li><a href="../entries/Superposition_Calculus.html">A Modular Formalization of Superposition</a></li>
               <li><a href="../entries/SCL_Simulates_Ground_Resolution.html">SCL Simulates Nonredundant Ground Resolution</a></li>
+              <li><a href="../entries/First_Order_Clause.html">First Order Clause</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Wlog.html
--- web/entries/Wlog.html
+++ web/entries/Wlog.html
@@ -135,6 +135,7 @@
           <ul class="horizontal-list">
               <li><a href="../entries/Complex_Bounded_Operators.html">Complex Bounded Operators</a></li>
               <li><a href="../entries/Hilbert_Space_Tensor_Product.html">The Tensor Product on Hilbert Spaces</a></li>
+              <li><a href="../entries/Sophie_Germain.html">Sophie Germain’s Theorem</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Worklist_Algorithms.html
--- web/entries/Worklist_Algorithms.html
+++ web/entries/Worklist_Algorithms.html
@@ -154,6 +154,12 @@
               <li><a href="../entries/Refine_Imperative_HOL.html">The Imperative Refinement Framework</a></li>
           </ul>
         </div>
+        <div>
+          <h3>Used by</h3>
+          <ul class="horizontal-list">
+              <li><a href="../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></li>
+          </ul>
+        </div>
     </div>
   </main>
 
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/Zeta_Function.html
--- web/entries/Zeta_Function.html
+++ web/entries/Zeta_Function.html
@@ -165,6 +165,7 @@
               <li><a href="../entries/Executable_Randomized_Algorithms.html">Executable Randomized Algorithms</a></li>
               <li><a href="../entries/Kummer_Congruence.html">Kummer's congruence</a></li>
               <li><a href="../entries/Chebyshev_Prime_Bounds.html">Concrete bounds for Chebyshev’s prime counting functions</a></li>
+              <li><a href="../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></li>
           </ul>
         </div>
     </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/entries/index.json
--- web/entries/index.json
+++ web/entries/index.json
@@ -1,5 +1,188 @@
 [
  {
+  "abstract": " This entry defines complex lattices, i.e. $\\Lambda(\\omega_1,\\omega_2) = \\mathbb{Z}\\omega_1 + \\mathbb{Z}\\omega_2$ where $\\omega_1/\\omega_2\\notin\\mathbb{R}$. Based on this, various other related topics are covered:\nthe modular group $\\Gamma$ and its fundamental region elliptic functions and their basic properties the Weierstraß elliptic function $\\wp$ and the fact that every elliptic function can be written in terms of $\\wp$ the Eisenstein series $G_n$ (including the forbidden series $G_2$) the ordinary differential equation satisfied by $\\wp$, the recurrence relation for $G_n$, and the addition and duplication theorems for $\\wp$ the lattice invariants $g_2$, $g_3$, and Klein's $J$ invariant the non-vanishing of the lattice discriminant $\\Delta$ $G_n$, $\\Delta$, $J$ as holomorphic functions in the upper half plane the Fourier expansion of $G_n(z)$ for $z\\to i\\infty$ the functional equations of $G_n$, $\\Delta$, $J$, and $\\eta$ w.r.t. the modular group Dedekind's $\\eta$ function the inversion formulas for the Jacobi $\\theta$ functions In particular, this entry contains most of Chapters 1 and 3 from Apostol's Modular Functions and Dirichlet Series in Number Theory and parts of Chapter 2.\nThe purpose of this entry is to provide a foundation for further formalisation of modular functions and modular forms.\n",
+  "authors": [
+   "Manuel Eberl",
+   "Anthony Bordg",
+   "Wenda Li",
+   "Lawrence C. Paulson"
+  ],
+  "link": "/entries/Elliptic_Functions.html",
+  "shortname": "Elliptic_Functions",
+  "title": "Complex Lattices, Elliptic Functions, and the Modular Group",
+  "topic_links": [
+   "/topics/mathematics/number-theory/"
+  ],
+  "topics": [
+   "Mathematics/Number theory"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "Munta is a verified model checker for timed automata. It has been described in detail in a PhD thesis [3] and conference publications [4, 2]. This entry supersedes earlier versions of Munta hosted on GitHub.",
+  "authors": [
+   "Simon Wimmer"
+  ],
+  "link": "/entries/Munta_Model_Checker.html",
+  "shortname": "Munta_Model_Checker",
+  "title": "Munta: A Verified Model Checker for Timed Automata",
+  "topic_links": [
+   "/topics/computer-science/automata-and-formal-languages/",
+   "/topics/tools/"
+  ],
+  "topics": [
+   "Computer science/Automata and formal languages",
+   "Tools"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "This is a basic library of definitions and results about context-free grammars and languages. It includes context-free grammars and languages, parse trees, Chomsky normal form, pumping lemmas and the relationship of right-linear grammars to finite automata.",
+  "authors": [
+   "Tobias Nipkow",
+   "Markus Gschoßmann",
+   "Felix Krayer",
+   "Fabian Lehr",
+   "Bruno Philipp",
+   "August Martin Stimpfle",
+   "Kaan Taskin",
+   "Akihisa Yamada"
+  ],
+  "link": "/entries/Context_Free_Grammar.html",
+  "shortname": "Context_Free_Grammar",
+  "title": "Context-Free Grammars and Languages",
+  "topic_links": [
+   "/topics/computer-science/automata-and-formal-languages/"
+  ],
+  "topics": [
+   "Computer science/Automata and formal languages"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "Most verification techniques use expressions, for example when assigning to variables or forming assertions over the state. Deep embeddings provide such expressions using a datatype, which allows queries over the syntax, such as calculating the free variables, and performing substitutions. Shallow embeddings, in contrast, model expressions as query functions over the state type, and are more amenable to automating verification tasks. The goal of this library is provide an intuitive implementation of shallow expressions, which nevertheless provides many of the benefits of a deep embedding. We harness the Optics library to provide an algebraic semantics for state variables, and use syntax translations to provide an intuitive lifted expression syntax. Furthermore, we provide a variety of meta-logic-style queries on expressions, such as dependencies on a state variable, and substitution of a variable for an expression. We also provide proof methods, based on the simplifier, to automate the associated proof tasks.",
+  "authors": [
+   "Simon Foster"
+  ],
+  "link": "/entries/Shallow_Expressions.html",
+  "shortname": "Shallow_Expressions",
+  "title": "Shallow Expressions",
+  "topic_links": [
+   "/topics/computer-science/semantics-and-reasoning/",
+   "/topics/computer-science/programming-languages/logics/"
+  ],
+  "topics": [
+   "Computer science/Semantics and reasoning",
+   "Computer science/Programming languages/Logics"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "The sequence $F_n$ of Farey fractions of order $n$ has the form $$\\frac{0}{1}, \\frac{1}{n}, \\frac{1}{n-1}, \\ldots, \\frac{n-1}{n}, \\frac{1}{1}$$ where the fractions appear in numerical order and have denominators at most $n$. The transformation from $F_n$ to $F_{n+1}$ can be effected by combining adjacent elements of the sequence~$F_n$, using an operation called the mediant. Adjacent (reduced) fractions $(a/b) \u003c (c/d)$ satisfy the unimodular relation $bc - ad = 1$ and their mediant is $\\frac{a+c}{b+d}$. A Ford circle is specified by a rational number, and interesting consequences follow in the case of Ford circles obtained from some Farey sequence~$F_n$. The formalised material is drawn from Apostol's Modular Functions and Dirichlet Series in Number Theory.",
+  "authors": [
+   "Lawrence C. Paulson"
+  ],
+  "link": "/entries/Farey_Sequences.html",
+  "shortname": "Farey_Sequences",
+  "title": "Farey Sequences and Ford Circles",
+  "topic_links": [
+   "/topics/mathematics/number-theory/"
+  ],
+  "topics": [
+   "Mathematics/Number theory"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "We use the Restriction_Spaces library as a semantic foundation for the process algebra framework HOL-CSP, offering a complementary backend to the existing HOLCF infrastructure. The type of processes is instantiated as a restriction space, and we prove that it is complete in this setting. This enables the construction of fixed points for recursive process definitions without having to rely exclusively on a pointed complete partial order. Notably, some operators are constructive without being Scott-continuous, and vice versa, illustrating the genuine complementarity between the two approaches. We also show that key CSP operators are either constructive or non-destructive, and verify the admissibility of several predicates, thereby supporting automated reasoning over recursive specifications.",
+  "authors": [
+   "Benoît Ballenghien",
+   "Burkhart Wolff"
+  ],
+  "link": "/entries/HOL-CSP_RS.html",
+  "shortname": "HOL-CSP_RS",
+  "title": "CSP Semantics over Restriction Spaces",
+  "topic_links": [
+   "/topics/computer-science/semantics-and-reasoning/",
+   "/topics/computer-science/concurrency/process-calculi/"
+  ],
+  "topics": [
+   "Computer science/Semantics and reasoning",
+   "Computer science/Concurrency/Process calculi"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "In this session, a number of examples are provided to illustrate how the Restriction_Spaces library works. The simple cases are, of course, covered: trivial restriction, booleans, integers, option type, and so on. But we also explore more elaborate constructions, such as formal power series and a trace model of the CSP process algebra. Additionally, we provide a lightweight integration with HOLCF, equipping restriction spaces with the inherited partial order structure when needed.",
+  "authors": [
+   "Benoît Ballenghien",
+   "Benjamin Puyobro",
+   "Burkhart Wolff"
+  ],
+  "link": "/entries/Restriction_Spaces-Examples.html",
+  "shortname": "Restriction_Spaces-Examples",
+  "title": "Examples of Restriction Spaces",
+  "topic_links": [
+   "/topics/mathematics/analysis/",
+   "/topics/mathematics/topology/",
+   "/topics/mathematics/algebra/",
+   "/topics/computer-science/concurrency/process-calculi/"
+  ],
+  "topics": [
+   "Mathematics/Analysis",
+   "Mathematics/Topology",
+   "Mathematics/Algebra",
+   "Computer science/Concurrency/Process calculi"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "Fixed-point constructions are fundamental to defining recursive and co-recursive functions.   However, a general axiom $Y f = f (Y f)$ leads to inconsistency, and definitions   must therefore be based on theories guaranteeing existence under suitable conditions.   In Isabelle/HOL, these constructions are typically based on sets, well-founded orders   or domain-theoretic models such as for example HOLCF.   In this submission we introduce a formalization of restriction spaces i.e. spaces   equipped with a so-called restriction, denoted by $\\downarrow$,   satifying three properties:   \\begin{align*}     \u0026 x \\downarrow 0 = y \\downarrow 0\\\\     \u0026 x \\downarrow n \\downarrow m = x \\downarrow \\mathrm{min} \\ n \\ m\\\\     \u0026 x \\not = y \\Longrightarrow \\exists n. \\ x \\downarrow n \\not = y \\downarrow n   \\end{align*}   They turn out to be cartesian closed and admit natural notions of constructiveness and   completeness, enabling the definition of a fixed-point operator under verifiable side-conditions.   This is achieved in our entry, from topological definitions to induction principles.   Additionally, we configure the simplifier so that it can automatically solve both   constructiveness and admissibility subgoals, as long as users write higher-order rules   for their operators.   Since our implementation relies on axiomatic type classes, the resulting library   is a fully abstract, flexible and reusable framework.",
+  "authors": [
+   "Benoît Ballenghien",
+   "Benjamin Puyobro",
+   "Burkhart Wolff"
+  ],
+  "link": "/entries/Restriction_Spaces.html",
+  "shortname": "Restriction_Spaces",
+  "title": "Restriction Spaces: a Fixed-Point Theory",
+  "topic_links": [
+   "/topics/tools/",
+   "/topics/mathematics/analysis/",
+   "/topics/mathematics/topology/",
+   "/topics/computer-science/semantics-and-reasoning/"
+  ],
+  "topics": [
+   "Tools",
+   "Mathematics/Analysis",
+   "Mathematics/Topology",
+   "Computer science/Semantics and reasoning"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "In this entry, we explore the relationship   between restriction spaces and usual metric   structures by instantiating the former as ultrametric spaces.   This is classically captured by defining the distance as   $$\\mathrm{dist} \\ x \\ y = \\inf_{x \\downarrow n \\ = \\ y \\downarrow n} \\left(\\frac{1}{2}\\right)^ n$$   but we actually generalize this perspective by introducing a hierarchy of increasingly   refined type classes to systematically relate ultrametric and restriction-based notions.   This layered approach enables a precise comparison of structural and topological properties.   In the end, our main result establishes that completeness in the sense of restriction spaces   coincides with standard metric completeness, thus bridging the gap between   our framework and Banach's fixed-point theorem established in HOL-Analysis.",
+  "authors": [
+   "Benoît Ballenghien",
+   "Benjamin Puyobro",
+   "Burkhart Wolff"
+  ],
+  "link": "/entries/Restriction_Spaces-Ultrametric.html",
+  "shortname": "Restriction_Spaces-Ultrametric",
+  "title": "Ultrametric Structure for Restriction Spaces",
+  "topic_links": [
+   "/topics/mathematics/analysis/",
+   "/topics/mathematics/topology/",
+   "/topics/computer-science/semantics-and-reasoning/"
+  ],
+  "topics": [
+   "Mathematics/Analysis",
+   "Mathematics/Topology",
+   "Computer science/Semantics and reasoning"
+  ],
+  "year": "2025"
+ },
+ {
   "abstract": "For integers $h$, $k$, the Dedekind sum is defined as \\[s(h,k) = \\sum_{r=1}^{k-1} \\frac{r}{k} \\left(\\left\\{\\frac{hr}{k}\\right\\} - \\frac{1}{2}\\right) \\] where $\\{x\\} = x - \\lfloor x\\rfloor$ denotes the fractional part of $x$.\nThese sums occur in various contexts in analytic number theory, e.g. in the functional equation of the Dedekind $\\eta$ function or in the study of modular forms.\nWe give the definition of $s(h,k)$ and prove its basic properties, including the reciprocity law \\[s(h,k) + s(k,h) = \\frac{1}{12hk} + \\frac{h}{12k} + \\frac{k}{12h} - \\frac{1}{4} \\] and various congruence results.\nOur formalisation follows Chapter 3 of Apostol's Modular Functions and Dirichlet Series in Number Theory and contains all facts related to Dedekind sums from it (without the exercises).\n",
   "authors": [
    "Manuel Eberl",
@@ -91,6 +274,40 @@
   "year": "2025"
  },
  {
+  "abstract": "An ultrametric space is a metric space in which the triangle inequality is strengthened by using the maximum instead of the sum. More formally, such a space is equipped with a real-valued application $dist$, called distance, verifying the four following conditions. \\begin{align*}     \u0026 dist \\ x \\ y \\ge 0\\\\     \u0026 dist \\ x \\ y = dist \\ y \\ x\\\\     \u0026 dist \\ x \\ y = 0 \\longleftrightarrow x = y\\\\     \u0026 dist \\ x \\ z \\le max \\ (dist \\ x \\ y) \\ (dist \\ y \\ z) \\end{align*} In this entry, we present an elementary formalization of these spaces relying on axiomatic type classes. The connection with standard metric spaces is obtained through a subclass relationship, and fundamental properties of ultrametric spaces are formally established.",
+  "authors": [
+   "Benoît Ballenghien",
+   "Benjamin Puyobro",
+   "Burkhart Wolff"
+  ],
+  "link": "/entries/Elementary_Ultrametric_Spaces.html",
+  "shortname": "Elementary_Ultrametric_Spaces",
+  "title": "Definition and Elementary Properties of Ultrametric Spaces",
+  "topic_links": [
+   "/topics/mathematics/topology/"
+  ],
+  "topics": [
+   "Mathematics/Topology"
+  ],
+  "year": "2025"
+ },
+ {
+  "abstract": "Sophie Germain's theorem states that if $p$ is a prime number such that $2 p + 1$ is also prime (with $p$ then called a Sophie Germain prime), a weak version of Fermat’s Last Theorem (FLT) holds for the exponent $p$. That is, there exist no nontrivial integer solutions $x, y, z \\in \\mathbb{Z}$ to the equation $$x ^ p + y ^ p = z ^ p$$ such that $p$ divides neither $x$, $y$ nor $z$. After some preliminary results, mainly on sufficient criteria for FLT, we propose a formalization of the classical version of Sophie Germain's theorem as well as a variant involving the concept of auxiliary primes that generalizes Sophie Germain primes.",
+  "authors": [
+   "Benoît Ballenghien"
+  ],
+  "link": "/entries/Sophie_Germain.html",
+  "shortname": "Sophie_Germain",
+  "title": "Sophie Germain’s Theorem",
+  "topic_links": [
+   "/topics/mathematics/number-theory/"
+  ],
+  "topics": [
+   "Mathematics/Number theory"
+  ],
+  "year": "2025"
+ },
+ {
   "abstract": " Continuous-time stochastic processes often carry the condition of having almost-surely continuous paths. If some process \\(X\\) satisfies certain bounds on its expectation, then the Kolmogorov-Chentsov theorem lets us construct a modification of \\(X\\), i.e. a process \\(X'\\) such that \\(\\forall t. X_t = X'_t\\) almost surely, that has Hölder continuous paths. In this work, we mechanise the Kolmogorov-Chentsov theorem. To get there, we develop a theory of stochastic processes, together with Hölder continuity, convergence in measure, and arbitrary intervals of dyadic rationals. With this, we pave the way towards a construction of Brownian motion. The work is based on the exposition in Achim Klenke's probability theory text. ",
   "authors": [
    "Christian Pardillo-Laursen",
@@ -524,7 +741,7 @@
   "year": "2024"
  },
  {
-  "abstract": " This entry defines the sum-of-squares function $r_k(n)$, which counts the number of ways to write a natural number $n$ as a sum of $k$ squares of integers. Signs and permutations of these integers are taken into account, such that e.g. $1^2+2^2$, $2^2+1^2$, and $(-1)^2+2^2$ are all different decompositions of $5$. Using this, I then formalise the main result: Jacobi's two-square theorem, which states that for $n \u003e 0$ we have \\[r_2(n) = 4(d_1(3) - d_3(n))\\ ,\\] where $d_i(n)$ denotes the number of divisors $m$ of $n$ such that $m = i\\ (\\text{mod}\\ 4)$. Corollaries include the identities $r_2(2n) = r_2(n)$ and $r_2(p^2n) = r_2(n)$ if $p = 3\\ (\\text{mod}\\ 4)$ and the well-known theorem that $r_2(n) = 0$ iff $n$ has a prime factor $p$ of odd multiplicity with $p = 3\\ (\\text{mod}\\ 4)$. ",
+  "abstract": " This entry defines the sum-of-squares function $r_k(n)$, which counts the number of ways to write a natural number $n$ as a sum of $k$ squares of integers. Signs and permutations of these integers are taken into account, such that e.g. $1^2+2^2$, $2^2+1^2$, and $(-1)^2+2^2$ are all different decompositions of $5$. Using this, I then formalise the main result: Jacobi's two-square theorem, which states that for $n \u003e 0$ we have \\[r_2(n) = 4(d_1(n) - d_3(n))\\ ,\\] where $d_i(n)$ denotes the number of divisors $m$ of $n$ such that $m = i\\ (\\text{mod}\\ 4)$. Corollaries include the identities $r_2(2n) = r_2(n)$ and $r_2(p^2n) = r_2(n)$ if $p = 3\\ (\\text{mod}\\ 4)$ and the well-known theorem that $r_2(n) = 0$ iff $n$ has a prime factor $p$ of odd multiplicity with $p = 3\\ (\\text{mod}\\ 4)$. ",
   "authors": [
    "Manuel Eberl"
   ],
@@ -2778,7 +2995,7 @@
  {
   "abstract": "We formalize the Legendre's three squares theorem and its consequences, in particular the following results: A natural number can be represented as the sum of three squares of natural numbers if and only if it is not of the form $4^a (8 k + 7)$, where $a$ and $k$ are natural numbers. If $n$ is a natural number such that $n \\equiv 3 \\pmod{8}$, then $n$ can be be represented as the sum of three squares of odd natural numbers. Consequences include the following: An integer $n$ can be written as $n = x^2 + y^2 + z^2 + z$, where $x$, $y$, $z$ are integers, if and only if $n \\geq 0$. The Legendre's four squares theorem: any natural number can be represented as the sum of four squares of natural numbers. ",
   "authors": [
-   "Anton Danilkin",
+   "Anna Danilkin",
    "Loïc Chevalier"
   ],
   "link": "/entries/Three_Squares.html",
diff -r 848f2c00e57f -r 29846bc89de2 web/index.html
--- web/index.html
+++ web/index.html
@@ -102,6 +102,122 @@
         <article class="entry">
           <div class="item-text">
             <h5>
+              <a class="title" href="./entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/wimmer/">Simon Wimmer</a>
+          </div>
+          <span class="date">May 22</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/eberl/">Manuel Eberl</a>, 
+              <a href="./authors/bordg/">Anthony Bordg</a>, 
+              <a href="./authors/li/">Wenda Li</a> and 
+              <a href="./authors/paulson/">Lawrence C. Paulson</a>
+          </div>
+          <span class="date">May 22</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/nipkow/">Tobias Nipkow</a>, 
+              <a href="./authors/gschossmann/">Markus Gschoßmann</a>, 
+              <a href="./authors/krayer/">Felix Krayer</a>, 
+              <a href="./authors/lehr/">Fabian Lehr</a>, 
+              <a href="./authors/philipp/">Bruno Philipp</a>, 
+              <a href="./authors/stimpfle/">August Martin Stimpfle</a>, 
+              <a href="./authors/taskin/">Kaan Taskin</a> and 
+              <a href="./authors/yamada/">Akihisa Yamada</a>
+          </div>
+          <span class="date">May 21</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Shallow_Expressions.html">Shallow Expressions</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/fosters/">Simon Foster</a>
+          </div>
+          <span class="date">May 16</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Farey_Sequences.html">Farey Sequences and Ford Circles</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/paulson/">Lawrence C. Paulson</a>
+          </div>
+          <span class="date">May 15</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a>, 
+              <a href="./authors/puyobro/">Benjamin Puyobro</a> and 
+              <a href="./authors/wolff/">Burkhart Wolff</a>
+          </div>
+          <span class="date">May 07</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a>, 
+              <a href="./authors/puyobro/">Benjamin Puyobro</a> and 
+              <a href="./authors/wolff/">Burkhart Wolff</a>
+          </div>
+          <span class="date">May 07</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a>, 
+              <a href="./authors/puyobro/">Benjamin Puyobro</a> and 
+              <a href="./authors/wolff/">Burkhart Wolff</a>
+          </div>
+          <span class="date">May 07</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a> and 
+              <a href="./authors/wolff/">Burkhart Wolff</a>
+          </div>
+          <span class="date">May 07</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
               <a class="title" href="./entries/Pentagonal_Number_Theorem.html">The Partition Function and the Pentagonal Number Theorem</a>
             </h5>
             <br>
@@ -168,6 +284,30 @@
         <article class="entry">
           <div class="item-text">
             <h5>
+              <a class="title" href="./entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a>, 
+              <a href="./authors/puyobro/">Benjamin Puyobro</a> and 
+              <a href="./authors/wolff/">Burkhart Wolff</a>
+          </div>
+          <span class="date">Apr 10</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
+              <a class="title" href="./entries/Sophie_Germain.html">Sophie Germain’s Theorem</a>
+            </h5>
+            <br>
+            by
+              <a href="./authors/ballenghien/">Benoît Ballenghien</a>
+          </div>
+          <span class="date">Apr 09</span>
+        </article>
+        <article class="entry">
+          <div class="item-text">
+            <h5>
               <a class="title" href="./entries/Kolmogorov_Chentsov.html">The Kolmogorov-Chentsov theorem</a>
             </h5>
             <br>
@@ -1951,7 +2091,7 @@
             </h5>
             <br>
             by
-              <a href="./authors/danilkin/">Anton Danilkin</a> and 
+              <a href="./authors/danilkin/">Anna Danilkin</a> and 
               <a href="./authors/chevalier/">Loïc Chevalier</a>
           </div>
           <span class="date">May 03</span>
diff -r 848f2c00e57f -r 29846bc89de2 web/index.xml
--- web/index.xml
+++ web/index.xml
@@ -4,6 +4,80 @@
     <title>Archive of Formal Proofs</title>
     <link>https://isa-afp.org/</link>
     <description>Formal developments in the AFP</description><item>
+  <title>Munta: A Verified Model Checker for Timed Automata</title>
+  <link>https://isa-afp.org/entries/Munta_Model_Checker.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Munta: A Verified Model Checker for Timed Automata in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+    <category>Tools</category>
+</item>
+<item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+<item>
+  <title>Shallow Expressions</title>
+  <link>https://isa-afp.org/entries/Shallow_Expressions.html</link>
+  <pubDate>Fri, 16 May 2025 00:00:00 +0000</pubDate>
+  <description>Shallow Expressions in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Programming languages/Logics</category>
+</item>
+<item>
+  <title>Farey Sequences and Ford Circles</title>
+  <link>https://isa-afp.org/entries/Farey_Sequences.html</link>
+  <pubDate>Thu, 15 May 2025 00:00:00 +0000</pubDate>
+  <description>Farey Sequences and Ford Circles in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>CSP Semantics over Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/HOL-CSP_RS.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>CSP Semantics over Restriction Spaces in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
   <title>The Partition Function and the Pentagonal Number Theorem</title>
   <link>https://isa-afp.org/entries/Pentagonal_Number_Theorem.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
@@ -39,6 +113,20 @@
     <category>Logic/Rewriting</category>
 </item>
 <item>
+  <title>Definition and Elementary Properties of Ultrametric Spaces</title>
+  <link>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</link>
+  <pubDate>Thu, 10 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Definition and Elementary Properties of Ultrametric Spaces in the AFP</description>
+    <category>Mathematics/Topology</category>
+</item>
+<item>
+  <title>Sophie Germain’s Theorem</title>
+  <link>https://isa-afp.org/entries/Sophie_Germain.html</link>
+  <pubDate>Wed, 09 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Sophie Germain’s Theorem in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>The Kolmogorov-Chentsov theorem</title>
   <link>https://isa-afp.org/entries/Kolmogorov_Chentsov.html</link>
   <pubDate>Sun, 06 Apr 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/context_free_grammar/index.html
--- /dev/null
+++ web/sessions/context_free_grammar/index.html
@@ -0,0 +1,133 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Context_Free_Grammar - Archive of Formal Proofs</title>
+    <meta name="description" content="Context_Free_Grammar in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Context_Free_Grammar in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Context_Free_Grammar" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/context_free_grammar/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Context_Free_Grammar.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>C</span>ontext_<span class='first'>F</span>ree_<span class='first'>G</span>rammar
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Context_Free_Grammar" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Context_Free_Grammar.html">
+        <h2>Context_Free_Grammar</h2>
+      </a>
+      <a id="Parse_Tree" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Parse_Tree.html">
+        <h2>Parse_Tree</h2>
+      </a>
+      <a id="Renaming_CFG" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Renaming_CFG.html">
+        <h2>Renaming_CFG</h2>
+      </a>
+      <a id="Disjoint_Union_CFG" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Disjoint_Union_CFG.html">
+        <h2>Disjoint_Union_CFG</h2>
+      </a>
+      <a id="Context_Free_Language" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Context_Free_Language.html">
+        <h2>Context_Free_Language</h2>
+      </a>
+      <a id="Unit_Elimination" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Unit_Elimination.html">
+        <h2>Unit_Elimination</h2>
+      </a>
+      <a id="Epsilon_Elimination" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Epsilon_Elimination.html">
+        <h2>Epsilon_Elimination</h2>
+      </a>
+      <a id="Chomsky_Normal_Form" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Chomsky_Normal_Form.html">
+        <h2>Chomsky_Normal_Form</h2>
+      </a>
+      <a id="Pumping_Lemma_CFG" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Pumping_Lemma_CFG.html">
+        <h2>Pumping_Lemma_CFG</h2>
+      </a>
+      <a id="AnBnCn_not_CFL" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/AnBnCn_not_CFL.html">
+        <h2>AnBnCn_not_CFL</h2>
+      </a>
+      <a id="CFL_Not_Intersection_Closed" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/CFL_Not_Intersection_Closed.html">
+        <h2>CFL_Not_Intersection_Closed</h2>
+      </a>
+      <a id="Inlining1Prod" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Inlining1Prod.html">
+        <h2>Inlining1Prod</h2>
+      </a>
+      <a id="Binarize" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Binarize.html">
+        <h2>Binarize</h2>
+      </a>
+      <a id="Right_Linear" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Right_Linear.html">
+        <h2>Right_Linear</h2>
+      </a>
+      <a id="NDA_rlin2" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/NDA_rlin2.html">
+        <h2>NDA_rlin2</h2>
+      </a>
+      <a id="Right_Linear_Automata" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Right_Linear_Automata.html">
+        <h2>Right_Linear_Automata</h2>
+      </a>
+      <a id="Pumping_Lemma_Regular" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/Pumping_Lemma_Regular.html">
+        <h2>Pumping_Lemma_Regular</h2>
+      </a>
+      <a id="AnBn_Not_Regular" href="https://www.isa-afp.org/browser_info/current/AFP/Context_Free_Grammar/AnBn_Not_Regular.html">
+        <h2>AnBn_Not_Regular</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/elementary_ultrametric_spaces/index.html
--- /dev/null
+++ web/sessions/elementary_ultrametric_spaces/index.html
@@ -0,0 +1,82 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Elementary_Ultrametric_Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="Elementary_Ultrametric_Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Elementary_Ultrametric_Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Elementary_Ultrametric_Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/elementary_ultrametric_spaces/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Elementary_Ultrametric_Spaces.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>E</span>lementary_<span class='first'>U</span>ltrametric_<span class='first'>S</span>paces
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Elementary_Ultrametric_Spaces" href="https://www.isa-afp.org/browser_info/current/AFP/Elementary_Ultrametric_Spaces/Elementary_Ultrametric_Spaces.html">
+        <h2>Elementary_Ultrametric_Spaces</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/farey_sequences/index.html
--- /dev/null
+++ web/sessions/farey_sequences/index.html
@@ -0,0 +1,82 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Farey_Sequences - Archive of Formal Proofs</title>
+    <meta name="description" content="Farey_Sequences in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Farey_Sequences in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Farey_Sequences" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/farey_sequences/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Farey_Sequences.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>F</span>arey_<span class='first'>S</span>equences
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Farey_Ford" href="https://www.isa-afp.org/browser_info/current/AFP/Farey_Sequences/Farey_Ford.html">
+        <h2>Farey_Ford</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/first_order_clause/index.html
--- web/sessions/first_order_clause/index.html
+++ web/sessions/first_order_clause/index.html
@@ -72,20 +72,26 @@
 </header>
       <div>
   <main id="theories">
+      <a id="Context_Notation" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Context_Notation.html">
+        <h2>Context_Notation</h2>
+      </a>
       <a id="Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Typing.html">
         <h2>Typing</h2>
       </a>
-      <a id="Context_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Context_Extra.html">
-        <h2>Context_Extra</h2>
-      </a>
       <a id="Term_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Term_Typing.html">
         <h2>Term_Typing</h2>
       </a>
-      <a id="Ground_Term_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Term_Extra.html">
-        <h2>Ground_Term_Extra</h2>
+      <a id="Generic_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Generic_Context.html">
+        <h2>Generic_Context</h2>
+      </a>
+      <a id="Parallel_Induct" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Parallel_Induct.html">
+        <h2>Parallel_Induct</h2>
       </a>
-      <a id="Ground_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Context.html">
-        <h2>Ground_Context</h2>
+      <a id="IsaFoR_Ground_Term" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Ground_Term.html">
+        <h2>IsaFoR_Ground_Term</h2>
+      </a>
+      <a id="IsaFoR_Ground_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Ground_Context.html">
+        <h2>IsaFoR_Ground_Context</h2>
       </a>
       <a id="Ground_Term_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Term_Typing.html">
         <h2>Ground_Term_Typing</h2>
@@ -93,9 +99,6 @@
       <a id="Nonground_Term" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Term.html">
         <h2>Nonground_Term</h2>
       </a>
-      <a id="Nonground_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Context.html">
-        <h2>Nonground_Context</h2>
-      </a>
       <a id="Multiset_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Multiset_Extra.html">
         <h2>Multiset_Extra</h2>
       </a>
@@ -117,56 +120,32 @@
       <a id="Grounded_Selection_Function" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Grounded_Selection_Function.html">
         <h2>Grounded_Selection_Function</h2>
       </a>
-      <a id="Collect_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Collect_Extra.html">
-        <h2>Collect_Extra</h2>
-      </a>
-      <a id="Infinite_Variables_Per_Type" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Infinite_Variables_Per_Type.html">
-        <h2>Infinite_Variables_Per_Type</h2>
+      <a id="IsaFoR_Nonground_Term" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Nonground_Term.html">
+        <h2>IsaFoR_Nonground_Term</h2>
       </a>
-      <a id="Typed_Functional_Substitution" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Typed_Functional_Substitution.html">
-        <h2>Typed_Functional_Substitution</h2>
-      </a>
-      <a id="Nonground_Term_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Term_Typing.html">
-        <h2>Nonground_Term_Typing</h2>
-      </a>
-      <a id="Monomorphic_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Monomorphic_Typing.html">
-        <h2>Monomorphic_Typing</h2>
+      <a id="Literal_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Literal_Functor.html">
+        <h2>Literal_Functor</h2>
       </a>
       <a id="Uprod_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Uprod_Extra.html">
         <h2>Uprod_Extra</h2>
       </a>
-      <a id="Literal_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Literal_Functor.html">
-        <h2>Literal_Functor</h2>
-      </a>
       <a id="Uprod_Literal_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Uprod_Literal_Functor.html">
         <h2>Uprod_Literal_Functor</h2>
       </a>
-      <a id="Ground_Clause" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Clause.html">
-        <h2>Ground_Clause</h2>
-      </a>
       <a id="Nonground_Clause_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Clause_With_Equality.html">
         <h2>Nonground_Clause_With_Equality</h2>
       </a>
-      <a id="Term_Rewrite_System" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Term_Rewrite_System.html">
-        <h2>Term_Rewrite_System</h2>
+      <a id="IsaFoR_Nonground_Clause_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Nonground_Clause_With_Equality.html">
+        <h2>IsaFoR_Nonground_Clause_With_Equality</h2>
       </a>
-      <a id="Entailment_Lifting" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Entailment_Lifting.html">
-        <h2>Entailment_Lifting</h2>
-      </a>
-      <a id="Fold_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Fold_Extra.html">
-        <h2>Fold_Extra</h2>
+      <a id="Nonground_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Context.html">
+        <h2>Nonground_Context</h2>
       </a>
-      <a id="Nonground_Entailment" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Entailment.html">
-        <h2>Nonground_Entailment</h2>
-      </a>
-      <a id="Inference_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Inference_Functor.html">
-        <h2>Inference_Functor</h2>
+      <a id="Context_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Context_Functor.html">
+        <h2>Context_Functor</h2>
       </a>
-      <a id="Nonground_Inference" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Inference.html">
-        <h2>Nonground_Inference</h2>
-      </a>
-      <a id="Nonground_Clause" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Clause.html">
-        <h2>Nonground_Clause</h2>
+      <a id="IsaFoR_Nonground_Context" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Nonground_Context.html">
+        <h2>IsaFoR_Nonground_Context</h2>
       </a>
       <a id="Restricted_Order" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Restricted_Order.html">
         <h2>Restricted_Order</h2>
@@ -210,18 +189,72 @@
       <a id="Nonground_Order_Generic" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Order_Generic.html">
         <h2>Nonground_Order_Generic</h2>
       </a>
+      <a id="Ground_Order_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Order_With_Equality.html">
+        <h2>Ground_Order_With_Equality</h2>
+      </a>
+      <a id="Nonground_Order_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Order_With_Equality.html">
+        <h2>Nonground_Order_With_Equality</h2>
+      </a>
+      <a id="Selection_Function_Select_Max" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Selection_Function_Select_Max.html">
+        <h2>Selection_Function_Select_Max</h2>
+      </a>
+      <a id="IsaFoR_KBO" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_KBO.html">
+        <h2>IsaFoR_KBO</h2>
+      </a>
+      <a id="Nonground_Clause" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Clause.html">
+        <h2>Nonground_Clause</h2>
+      </a>
+      <a id="IsaFoR_Nonground_Clause" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Nonground_Clause.html">
+        <h2>IsaFoR_Nonground_Clause</h2>
+      </a>
+      <a id="Ground_Term_Rewrite_System" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Term_Rewrite_System.html">
+        <h2>Ground_Term_Rewrite_System</h2>
+      </a>
+      <a id="Ground_Critical_Pairs" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Critical_Pairs.html">
+        <h2>Ground_Critical_Pairs</h2>
+      </a>
+      <a id="Ground_Critical_Peaks" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Critical_Peaks.html">
+        <h2>Ground_Critical_Peaks</h2>
+      </a>
+      <a id="IsaFoR_Ground_Critical_Pairs" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/IsaFoR_Ground_Critical_Pairs.html">
+        <h2>IsaFoR_Ground_Critical_Pairs</h2>
+      </a>
+      <a id="Collect_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Collect_Extra.html">
+        <h2>Collect_Extra</h2>
+      </a>
+      <a id="Infinite_Variables_Per_Type" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Infinite_Variables_Per_Type.html">
+        <h2>Infinite_Variables_Per_Type</h2>
+      </a>
+      <a id="Typed_Functional_Substitution" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Typed_Functional_Substitution.html">
+        <h2>Typed_Functional_Substitution</h2>
+      </a>
+      <a id="Nonground_Term_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Term_Typing.html">
+        <h2>Nonground_Term_Typing</h2>
+      </a>
+      <a id="Monomorphic_Typing" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Monomorphic_Typing.html">
+        <h2>Monomorphic_Typing</h2>
+      </a>
+      <a id="Entailment_Lifting" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Entailment_Lifting.html">
+        <h2>Entailment_Lifting</h2>
+      </a>
+      <a id="Fold_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Fold_Extra.html">
+        <h2>Fold_Extra</h2>
+      </a>
+      <a id="Nonground_Entailment" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Entailment.html">
+        <h2>Nonground_Entailment</h2>
+      </a>
+      <a id="Inference_Functor" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Inference_Functor.html">
+        <h2>Inference_Functor</h2>
+      </a>
+      <a id="Nonground_Inference" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Inference.html">
+        <h2>Nonground_Inference</h2>
+      </a>
       <a id="Ground_Order" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Order.html">
         <h2>Ground_Order</h2>
       </a>
       <a id="Nonground_Order" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Order.html">
         <h2>Nonground_Order</h2>
       </a>
-      <a id="Ground_Order_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Ground_Order_With_Equality.html">
-        <h2>Ground_Order_With_Equality</h2>
-      </a>
-      <a id="Nonground_Order_With_Equality" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Nonground_Order_With_Equality.html">
-        <h2>Nonground_Order_With_Equality</h2>
-      </a>
       <a id="Natural_Magma_Typing_Lifting" href="https://www.isa-afp.org/browser_info/current/AFP/First_Order_Clause/Natural_Magma_Typing_Lifting.html">
         <h2>Natural_Magma_Typing_Lifting</h2>
       </a>
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/hol-csp_rs/index.html
--- /dev/null
+++ web/sessions/hol-csp_rs/index.html
@@ -0,0 +1,115 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>HOL-CSP_RS - Archive of Formal Proofs</title>
+    <meta name="description" content="HOL-CSP_RS in the Archive of Formal Proofs" />
+    <meta property="og:description" content="HOL-CSP_RS in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="HOL-CSP_RS" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/hol-csp_rs/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/HOL-CSP_RS.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>H</span><span class='first'>O</span><span class='first'>L</span>-<span class='first'>C</span><span class='first'>S</span><span class='first'>P</span>_<span class='first'>R</span><span class='first'>S</span>
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Process_Restriction_Space" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Process_Restriction_Space.html">
+        <h2>Process_Restriction_Space</h2>
+      </a>
+      <a id="Prefixes_Constructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Prefixes_Constructive.html">
+        <h2>Prefixes_Constructive</h2>
+      </a>
+      <a id="Choices_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Choices_Non_Destructive.html">
+        <h2>Choices_Non_Destructive</h2>
+      </a>
+      <a id="Renaming_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Renaming_Non_Destructive.html">
+        <h2>Renaming_Non_Destructive</h2>
+      </a>
+      <a id="Sequential_Composition_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Sequential_Composition_Non_Destructive.html">
+        <h2>Sequential_Composition_Non_Destructive</h2>
+      </a>
+      <a id="Synchronization_Product_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Synchronization_Product_Non_Destructive.html">
+        <h2>Synchronization_Product_Non_Destructive</h2>
+      </a>
+      <a id="Throw_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Throw_Non_Destructive.html">
+        <h2>Throw_Non_Destructive</h2>
+      </a>
+      <a id="Interrupt_Non_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Interrupt_Non_Destructive.html">
+        <h2>Interrupt_Non_Destructive</h2>
+      </a>
+      <a id="After_Operator_Non_Too_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/After_Operator_Non_Too_Destructive.html">
+        <h2>After_Operator_Non_Too_Destructive</h2>
+      </a>
+      <a id="Hiding_Destructive" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/Hiding_Destructive.html">
+        <h2>Hiding_Destructive</h2>
+      </a>
+      <a id="CSP_Restriction_Adm" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/CSP_Restriction_Adm.html">
+        <h2>CSP_Restriction_Adm</h2>
+      </a>
+      <a id="HOL-CSP_RS" href="https://www.isa-afp.org/browser_info/current/AFP/HOL-CSP_RS/HOL-CSP_RS.html">
+        <h2>HOL-CSP_RS</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/munta_base/index.html
--- /dev/null
+++ web/sessions/munta_base/index.html
@@ -0,0 +1,139 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Munta_Base - Archive of Formal Proofs</title>
+    <meta name="description" content="Munta_Base in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Munta_Base in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Munta_Base" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/munta_base/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Munta_Model_Checker.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>M</span>unta_<span class='first'>B</span>ase
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="TA_Syntax_Bundles" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/TA_Syntax_Bundles.html">
+        <h2>TA_Syntax_Bundles</h2>
+      </a>
+      <a id="ML_Util" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/ML_Util.html">
+        <h2>ML_Util</h2>
+      </a>
+      <a id="Reordering_Quantifiers" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Reordering_Quantifiers.html">
+        <h2>Reordering_Quantifiers</h2>
+      </a>
+      <a id="More_Methods" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/More_Methods.html">
+        <h2>More_Methods</h2>
+      </a>
+      <a id="Abstract_Term" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Abstract_Term.html">
+        <h2>Abstract_Term</h2>
+      </a>
+      <a id="Bijective_Embedding" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Bijective_Embedding.html">
+        <h2>Bijective_Embedding</h2>
+      </a>
+      <a id="Tracing" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Tracing.html">
+        <h2>Tracing</h2>
+      </a>
+      <a id="Printing" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Printing.html">
+        <h2>Printing</h2>
+      </a>
+      <a id="Trace_Timing" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Trace_Timing.html">
+        <h2>Trace_Timing</h2>
+      </a>
+      <a id="Error_List_Monad" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Error_List_Monad.html">
+        <h2>Error_List_Monad</h2>
+      </a>
+      <a id="Temporal_Logics" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Temporal_Logics.html">
+        <h2>Temporal_Logics</h2>
+      </a>
+      <a id="Subsumption_Graphs" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Subsumption_Graphs.html">
+        <h2>Subsumption_Graphs</h2>
+      </a>
+      <a id="Sepref_Acconstraint" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Sepref_Acconstraint.html">
+        <h2>Sepref_Acconstraint</h2>
+      </a>
+      <a id="TA_DBM_Operations_Impl" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/TA_DBM_Operations_Impl.html">
+        <h2>TA_DBM_Operations_Impl</h2>
+      </a>
+      <a id="TA_More" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/TA_More.html">
+        <h2>TA_More</h2>
+      </a>
+      <a id="Normalized_Zone_Semantics_Impl" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Normalized_Zone_Semantics_Impl.html">
+        <h2>Normalized_Zone_Semantics_Impl</h2>
+      </a>
+      <a id="TA_Impl_Misc" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/TA_Impl_Misc.html">
+        <h2>TA_Impl_Misc</h2>
+      </a>
+      <a id="Normalized_Zone_Semantics_Impl_Semantic_Refinement" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Normalized_Zone_Semantics_Impl_Semantic_Refinement.html">
+        <h2>Normalized_Zone_Semantics_Impl_Semantic_Refinement</h2>
+      </a>
+      <a id="Normalized_Zone_Semantics_Impl_Refine" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Normalized_Zone_Semantics_Impl_Refine.html">
+        <h2>Normalized_Zone_Semantics_Impl_Refine</h2>
+      </a>
+      <a id="Normalized_Zone_Semantics_Impl_Extra" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Base/Normalized_Zone_Semantics_Impl_Extra.html">
+        <h2>Normalized_Zone_Semantics_Impl_Extra</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/munta_model_checker/index.html
--- /dev/null
+++ web/sessions/munta_model_checker/index.html
@@ -0,0 +1,169 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Munta_Model_Checker - Archive of Formal Proofs</title>
+    <meta name="description" content="Munta_Model_Checker in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Munta_Model_Checker in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Munta_Model_Checker" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/munta_model_checker/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Munta_Model_Checker.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>M</span>unta_<span class='first'>M</span>odel_<span class='first'>C</span>hecker
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Munta_Tagging" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Munta_Tagging.html">
+        <h2>Munta_Tagging</h2>
+      </a>
+      <a id="Munta_Error_Monad_Add" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Munta_Error_Monad_Add.html">
+        <h2>Munta_Error_Monad_Add</h2>
+      </a>
+      <a id="Parser_Combinator" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Parser_Combinator.html">
+        <h2>Parser_Combinator</h2>
+      </a>
+      <a id="Lexer" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Lexer.html">
+        <h2>Lexer</h2>
+      </a>
+      <a id="JSON_Parsing" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/JSON_Parsing.html">
+        <h2>JSON_Parsing</h2>
+      </a>
+      <a id="Networks" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Networks.html">
+        <h2>Networks</h2>
+      </a>
+      <a id="State_Networks" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/State_Networks.html">
+        <h2>State_Networks</h2>
+      </a>
+      <a id="UPPAAL_Asm" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_Asm.html">
+        <h2>UPPAAL_Asm</h2>
+      </a>
+      <a id="UPPAAL_Asm_Clocks" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_Asm_Clocks.html">
+        <h2>UPPAAL_Asm_Clocks</h2>
+      </a>
+      <a id="UPPAAL_State_Networks" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_State_Networks.html">
+        <h2>UPPAAL_State_Networks</h2>
+      </a>
+      <a id="UPPAAL_State_Networks_Impl" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_State_Networks_Impl.html">
+        <h2>UPPAAL_State_Networks_Impl</h2>
+      </a>
+      <a id="Program_Analysis" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Program_Analysis.html">
+        <h2>Program_Analysis</h2>
+      </a>
+      <a id="UPPAAL_State_Networks_Impl_Refine" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_State_Networks_Impl_Refine.html">
+        <h2>UPPAAL_State_Networks_Impl_Refine</h2>
+      </a>
+      <a id="UPPAAL_Model_Checking" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/UPPAAL_Model_Checking.html">
+        <h2>UPPAAL_Model_Checking</h2>
+      </a>
+      <a id="Simple_Expressions" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Expressions.html">
+        <h2>Simple_Expressions</h2>
+      </a>
+      <a id="Simple_Network_Language" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language.html">
+        <h2>Simple_Network_Language</h2>
+      </a>
+      <a id="TA_Impl_Misc2" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/TA_Impl_Misc2.html">
+        <h2>TA_Impl_Misc2</h2>
+      </a>
+      <a id="TA_More2" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/TA_More2.html">
+        <h2>TA_More2</h2>
+      </a>
+      <a id="TA_Equivalences" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/TA_Equivalences.html">
+        <h2>TA_Equivalences</h2>
+      </a>
+      <a id="Simple_Network_Language_Impl" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Impl.html">
+        <h2>Simple_Network_Language_Impl</h2>
+      </a>
+      <a id="Simple_Network_Language_Impl_Refine" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Impl_Refine.html">
+        <h2>Simple_Network_Language_Impl_Refine</h2>
+      </a>
+      <a id="Simple_Network_Language_Model_Checking" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Model_Checking.html">
+        <h2>Simple_Network_Language_Model_Checking</h2>
+      </a>
+      <a id="Deadlock" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Deadlock.html">
+        <h2>Deadlock</h2>
+      </a>
+      <a id="Deadlock_Impl" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Deadlock_Impl.html">
+        <h2>Deadlock_Impl</h2>
+      </a>
+      <a id="Deadlock_Checking" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Deadlock_Checking.html">
+        <h2>Deadlock_Checking</h2>
+      </a>
+      <a id="Simple_Network_Language_Renaming" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Renaming.html">
+        <h2>Simple_Network_Language_Renaming</h2>
+      </a>
+      <a id="Simple_Network_Language_Deadlock_Checking" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Deadlock_Checking.html">
+        <h2>Simple_Network_Language_Deadlock_Checking</h2>
+      </a>
+      <a id="Shortest_SCC_Paths" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Shortest_SCC_Paths.html">
+        <h2>Shortest_SCC_Paths</h2>
+      </a>
+      <a id="Simple_Network_Language_Export_Code" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Simple_Network_Language_Export_Code.html">
+        <h2>Simple_Network_Language_Export_Code</h2>
+      </a>
+      <a id="Munta_Compile_MLton" href="https://www.isa-afp.org/browser_info/current/AFP/Munta_Model_Checker/Munta_Compile_MLton.html">
+        <h2>Munta_Compile_MLton</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/restriction_spaces-examples/index.html
--- /dev/null
+++ web/sessions/restriction_spaces-examples/index.html
@@ -0,0 +1,109 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Restriction_Spaces-Examples - Archive of Formal Proofs</title>
+    <meta name="description" content="Restriction_Spaces-Examples in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Restriction_Spaces-Examples in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Restriction_Spaces-Examples" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/restriction_spaces-examples/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Restriction_Spaces-Examples.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>R</span>estriction_<span class='first'>S</span>paces-<span class='first'>E</span>xamples
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="RS_Any_Type" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Any_Type.html">
+        <h2>RS_Any_Type</h2>
+      </a>
+      <a id="RS_Bool" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Bool.html">
+        <h2>RS_Bool</h2>
+      </a>
+      <a id="RS_Nat" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Nat.html">
+        <h2>RS_Nat</h2>
+      </a>
+      <a id="RS_Int" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Int.html">
+        <h2>RS_Int</h2>
+      </a>
+      <a id="RS_Option" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Option.html">
+        <h2>RS_Option</h2>
+      </a>
+      <a id="RS_List" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_List.html">
+        <h2>RS_List</h2>
+      </a>
+      <a id="RS_Tree" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Tree.html">
+        <h2>RS_Tree</h2>
+      </a>
+      <a id="RS_Decimals" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Decimals.html">
+        <h2>RS_Decimals</h2>
+      </a>
+      <a id="RS_Trace_Model_CSP" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Trace_Model_CSP.html">
+        <h2>RS_Trace_Model_CSP</h2>
+      </a>
+      <a id="RS_Formal_Power_Series" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Examples/RS_Formal_Power_Series.html">
+        <h2>RS_Formal_Power_Series</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/restriction_spaces-holcf/index.html
--- /dev/null
+++ web/sessions/restriction_spaces-holcf/index.html
@@ -0,0 +1,82 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Restriction_Spaces-HOLCF - Archive of Formal Proofs</title>
+    <meta name="description" content="Restriction_Spaces-HOLCF in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Restriction_Spaces-HOLCF in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Restriction_Spaces-HOLCF" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/restriction_spaces-holcf/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Restriction_Spaces-Examples.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>R</span>estriction_<span class='first'>S</span>paces-<span class='first'>H</span><span class='first'>O</span><span class='first'>L</span><span class='first'>C</span><span class='first'>F</span>
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Restriction_Spaces-HOLCF" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-HOLCF/Restriction_Spaces-HOLCF.html">
+        <h2>Restriction_Spaces-HOLCF</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/restriction_spaces-ultrametric/index.html
--- /dev/null
+++ web/sessions/restriction_spaces-ultrametric/index.html
@@ -0,0 +1,97 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Restriction_Spaces-Ultrametric - Archive of Formal Proofs</title>
+    <meta name="description" content="Restriction_Spaces-Ultrametric in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Restriction_Spaces-Ultrametric in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Restriction_Spaces-Ultrametric" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/restriction_spaces-ultrametric/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Restriction_Spaces-Ultrametric.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>R</span>estriction_<span class='first'>S</span>paces-<span class='first'>U</span>ltrametric
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Banach_Theorem_Extension" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Banach_Theorem_Extension.html">
+        <h2>Banach_Theorem_Extension</h2>
+      </a>
+      <a id="Ultrametric_Restriction_Spaces_Locale" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces_Locale.html">
+        <h2>Ultrametric_Restriction_Spaces_Locale</h2>
+      </a>
+      <a id="Ultrametric_Restriction_Spaces" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Ultrametric_Restriction_Spaces.html">
+        <h2>Ultrametric_Restriction_Spaces</h2>
+      </a>
+      <a id="Fun_Ultrametric_Restriction_Spaces" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Fun_Ultrametric_Restriction_Spaces.html">
+        <h2>Fun_Ultrametric_Restriction_Spaces</h2>
+      </a>
+      <a id="Prod_Ultrametric_Restriction_Spaces" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Prod_Ultrametric_Restriction_Spaces.html">
+        <h2>Prod_Ultrametric_Restriction_Spaces</h2>
+      </a>
+      <a id="Restriction_Spaces-Ultrametric" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces-Ultrametric/Restriction_Spaces-Ultrametric.html">
+        <h2>Restriction_Spaces-Ultrametric</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/restriction_spaces/index.html
--- /dev/null
+++ web/sessions/restriction_spaces/index.html
@@ -0,0 +1,100 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Restriction_Spaces - Archive of Formal Proofs</title>
+    <meta name="description" content="Restriction_Spaces in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Restriction_Spaces in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Restriction_Spaces" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/restriction_spaces/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Restriction_Spaces.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>R</span>estriction_<span class='first'>S</span>paces
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Restriction_Spaces_Locales" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Locales.html">
+        <h2>Restriction_Spaces_Locales</h2>
+      </a>
+      <a id="Restriction_Spaces_Classes" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Classes.html">
+        <h2>Restriction_Spaces_Classes</h2>
+      </a>
+      <a id="Restriction_Spaces_Prod" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Prod.html">
+        <h2>Restriction_Spaces_Prod</h2>
+      </a>
+      <a id="Restriction_Spaces_Fun" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Fun.html">
+        <h2>Restriction_Spaces_Fun</h2>
+      </a>
+      <a id="Restriction_Spaces_Topology" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Topology.html">
+        <h2>Restriction_Spaces_Topology</h2>
+      </a>
+      <a id="Restriction_Spaces_Induction" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces_Induction.html">
+        <h2>Restriction_Spaces_Induction</h2>
+      </a>
+      <a id="Restriction_Spaces" href="https://www.isa-afp.org/browser_info/current/AFP/Restriction_Spaces/Restriction_Spaces.html">
+        <h2>Restriction_Spaces</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/shallow_expressions/index.html
--- /dev/null
+++ web/sessions/shallow_expressions/index.html
@@ -0,0 +1,118 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Shallow_Expressions - Archive of Formal Proofs</title>
+    <meta name="description" content="Shallow_Expressions in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Shallow_Expressions in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Shallow_Expressions" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/shallow_expressions/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Shallow_Expressions.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>S</span>hallow_<span class='first'>E</span>xpressions
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="Variables" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Variables.html">
+        <h2>Variables</h2>
+      </a>
+      <a id="Expressions" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Expressions.html">
+        <h2>Expressions</h2>
+      </a>
+      <a id="Unrestriction" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Unrestriction.html">
+        <h2>Unrestriction</h2>
+      </a>
+      <a id="Substitutions" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Substitutions.html">
+        <h2>Substitutions</h2>
+      </a>
+      <a id="Extension" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Extension.html">
+        <h2>Extension</h2>
+      </a>
+      <a id="Liberation" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Liberation.html">
+        <h2>Liberation</h2>
+      </a>
+      <a id="Quantifiers" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Quantifiers.html">
+        <h2>Quantifiers</h2>
+      </a>
+      <a id="Collections" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Collections.html">
+        <h2>Collections</h2>
+      </a>
+      <a id="Named_Expressions" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Named_Expressions.html">
+        <h2>Named_Expressions</h2>
+      </a>
+      <a id="Local_State" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Local_State.html">
+        <h2>Local_State</h2>
+      </a>
+      <a id="Shallow_Expressions" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Shallow_Expressions.html">
+        <h2>Shallow_Expressions</h2>
+      </a>
+      <a id="Expressions_Tests" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Expressions_Tests.html">
+        <h2>Expressions_Tests</h2>
+      </a>
+      <a id="Shallow_Expressions_Examples" href="https://www.isa-afp.org/browser_info/current/AFP/Shallow_Expressions/Shallow_Expressions_Examples.html">
+        <h2>Shallow_Expressions_Examples</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/sophie_germain/index.html
--- /dev/null
+++ web/sessions/sophie_germain/index.html
@@ -0,0 +1,94 @@
+<!DOCTYPE html>
+<html lang="en"><head>
+  <meta charset="utf-8" />
+  <meta http-equiv="X-UA-Compatible" content="IE=edge" />
+  <meta name="viewport" content="width=device-width, initial-scale=1" />
+    <title>Sophie_Germain - Archive of Formal Proofs</title>
+    <meta name="description" content="Sophie_Germain in the Archive of Formal Proofs" />
+    <meta property="og:description" content="Sophie_Germain in the Archive of Formal Proofs" />
+
+  <meta property="og:title" content="Sophie_Germain" />
+  <meta property="og:url" content="https://isa-afp.org/sessions/sophie_germain/" />
+  <meta property="og:image" content="https://isa-afp.org/images/afp.png" />
+    <meta property="og:type" content="website" />
+  <link rel="stylesheet" type="text/css" href="../../css/front.min.css">
+    <link rel="stylesheet" type="text/css" href="../../css/isabelle.css">
+
+  <link rel="icon" href="../../images/favicon.ico" type="image/icon">
+
+  <script src="../../js/obfuscate.js"></script>
+  <script src="../../js/flexsearch.bundle.js"></script>
+  <script src="../../js/scroll-spy.js"></script>
+  <script src="../../js/theory.js"></script>
+  <script src="../../js/util.js"></script>
+    <script src="../../js/header-search.js"></script>
+  <script src="../../js/search-autocomplete.js"></script>
+</head>
+
+  <body class="mathjax_ignore theories">
+    <aside><div id="menu-toggle">
+  <input id="toggle" type="checkbox" />
+  <label for="toggle">
+    <span>menu</span>
+    <img src="../../images/menu.svg" alt="Menu" />
+  </label>
+
+  <a href="../../" class="logo-link">
+    <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+  </a>
+
+  <nav id="menu">
+    <div>
+      <a href="../../" class="logo-link">
+        <img src="../../images/afp.png" alt="Logo of the Archive of Formal Proofs" class="logo">
+      </a>
+        <ul id="return">
+          <li>
+            <a href="../../entries/Sophie_Germain.html">Return to entry</a>
+          </li>
+        </ul>
+      <hr>
+      <ul id="theory-navbar" class="list-group"></ul>
+    </div>
+  </nav>
+</div>
+    </aside>
+
+    <div class="content"><header>
+    <form autocomplete="off" action="../../search">
+      <div class="form-container">
+        <input id="search-input" type="search" size="31" maxlength="255" value=""
+               aria-label="Search the AFP" list="autocomplete"><button id="search-button" type="button">
+          <img src="../../images/search.svg" alt="Search" />
+        </button>
+        <datalist id="autocomplete">
+        </datalist>
+      </div>
+    </form>
+  <h1 ><span class='first'>S</span>ophie_<span class='first'>G</span>ermain
+  </h1>
+  <div>
+  </div>
+</header>
+      <div>
+  <main id="theories">
+      <a id="SG_Introduction" href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/SG_Introduction.html">
+        <h2>SG_Introduction</h2>
+      </a>
+      <a id="SG_Preliminaries" href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/SG_Preliminaries.html">
+        <h2>SG_Preliminaries</h2>
+      </a>
+      <a id="FLT_Sufficient_Conditions" href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/FLT_Sufficient_Conditions.html">
+        <h2>FLT_Sufficient_Conditions</h2>
+      </a>
+      <a id="SG_Theorem" href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/SG_Theorem.html">
+        <h2>SG_Theorem</h2>
+      </a>
+      <a id="SG_Generalization" href="https://www.isa-afp.org/browser_info/current/AFP/Sophie_Germain/SG_Generalization.html">
+        <h2>SG_Generalization</h2>
+      </a>
+  </main>
+      </div>
+    </div>
+  </body>
+</html>
\ No newline at end of file
diff -r 848f2c00e57f -r 29846bc89de2 web/sessions/superposition_calculus/index.html
--- web/sessions/superposition_calculus/index.html
+++ web/sessions/superposition_calculus/index.html
@@ -72,9 +72,6 @@
 </header>
       <div>
   <main id="theories">
-      <a id="Ground_Critical_Pairs" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Ground_Critical_Pairs.html">
-        <h2>Ground_Critical_Pairs</h2>
-      </a>
       <a id="Ground_Superposition" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Ground_Superposition.html">
         <h2>Ground_Superposition</h2>
       </a>
@@ -90,12 +87,18 @@
       <a id="Ground_Superposition_Soundness" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Ground_Superposition_Soundness.html">
         <h2>Ground_Superposition_Soundness</h2>
       </a>
-      <a id="Ground_Superposition_Welltypedness_Preservation" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Ground_Superposition_Welltypedness_Preservation.html">
-        <h2>Ground_Superposition_Welltypedness_Preservation</h2>
+      <a id="Typed_Ground_Superposition" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Typed_Ground_Superposition.html">
+        <h2>Typed_Ground_Superposition</h2>
       </a>
       <a id="Superposition" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Superposition.html">
         <h2>Superposition</h2>
       </a>
+      <a id="Monomorphic_Superposition" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Monomorphic_Superposition.html">
+        <h2>Monomorphic_Superposition</h2>
+      </a>
+      <a id="Ground_Superposition_Example" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Ground_Superposition_Example.html">
+        <h2>Ground_Superposition_Example</h2>
+      </a>
       <a id="Grounded_Superposition" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Grounded_Superposition.html">
         <h2>Grounded_Superposition</h2>
       </a>
@@ -108,9 +111,6 @@
       <a id="Superposition_Soundness" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Superposition_Soundness.html">
         <h2>Superposition_Soundness</h2>
       </a>
-      <a id="IsaFoR_Term_Copy" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/IsaFoR_Term_Copy.html">
-        <h2>IsaFoR_Term_Copy</h2>
-      </a>
       <a id="Superposition_Example" href="https://www.isa-afp.org/browser_info/current/AFP/Superposition_Calculus/Superposition_Example.html">
         <h2>Superposition_Example</h2>
       </a>
diff -r 848f2c00e57f -r 29846bc89de2 web/sitemap.xml
--- web/sitemap.xml
+++ web/sitemap.xml
@@ -3,10 +3,37 @@
   xmlns:xhtml="http://www.w3.org/1999/xhtml">
   <url>
     <loc>https://isa-afp.org/entries/index.json</loc>
-    <lastmod>2025-04-23T00:00:00+00:00</lastmod>
+    <lastmod>2025-05-22T00:00:00+00:00</lastmod>
   </url><url>
     <loc>https://isa-afp.org/</loc>
-    <lastmod>2025-04-23T00:00:00+00:00</lastmod>
+    <lastmod>2025-05-22T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Elliptic_Functions.html</loc>
+    <lastmod>2025-05-22T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Munta_Model_Checker.html</loc>
+    <lastmod>2025-05-22T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Context_Free_Grammar.html</loc>
+    <lastmod>2025-05-21T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Shallow_Expressions.html</loc>
+    <lastmod>2025-05-16T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Farey_Sequences.html</loc>
+    <lastmod>2025-05-15T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/HOL-CSP_RS.html</loc>
+    <lastmod>2025-05-07T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</loc>
+    <lastmod>2025-05-07T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Restriction_Spaces.html</loc>
+    <lastmod>2025-05-07T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</loc>
+    <lastmod>2025-05-07T00:00:00+00:00</lastmod>
   </url><url>
     <loc>https://isa-afp.org/entries/Dedekind_Sums.html</loc>
     <lastmod>2025-04-23T00:00:00+00:00</lastmod>
@@ -23,6 +50,12 @@
     <loc>https://isa-afp.org/entries/Completeness_Decreasing_Diagrams_for_N1.html</loc>
     <lastmod>2025-04-12T00:00:00+00:00</lastmod>
   </url><url>
+    <loc>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</loc>
+    <lastmod>2025-04-10T00:00:00+00:00</lastmod>
+  </url><url>
+    <loc>https://isa-afp.org/entries/Sophie_Germain.html</loc>
+    <lastmod>2025-04-09T00:00:00+00:00</lastmod>
+  </url><url>
     <loc>https://isa-afp.org/entries/Kolmogorov_Chentsov.html</loc>
     <lastmod>2025-04-06T00:00:00+00:00</lastmod>
   </url><url>
@@ -2844,12 +2877,12 @@
   </url><url>
     <loc>https://isa-afp.org/authors/stueber/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/danilkin/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/anselmgod/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/bordg/</loc>
   </url><url>
-    <loc>https://isa-afp.org/authors/danilkin/</loc>
-  </url><url>
     <loc>https://isa-afp.org/sessions/aodv/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/aot/</loc>
@@ -2878,6 +2911,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/attack_trees/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/stimpfle/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/auto2_hol/</loc>
@@ -3012,6 +3047,8 @@
   </url><url>
     <loc>https://isa-afp.org/authors/langenstein/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/philipp/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/paleo/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/btree/</loc>
@@ -3300,6 +3337,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/containers-benchmarks/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/context_free_grammar/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/continued_fractions/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/cook_levin/</loc>
@@ -3546,6 +3585,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/efficient-mergesort/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/elementary_ultrametric_spaces/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/wenninger/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/elimination_of_repeated_factors/</loc>
@@ -3554,6 +3595,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/elliptic_curves_group_law/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/elliptic_functions/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/karayel/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/gunther/</loc>
@@ -3604,6 +3647,8 @@
   </url><url>
     <loc>https://isa-afp.org/authors/immler/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/lehr/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/mitterwallner/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/factor_algebraic_polynomial/</loc>
@@ -3616,6 +3661,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/falling_factorial_sum/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/farey_sequences/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/farkas/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/featherweight_ocl/</loc>
@@ -3624,6 +3671,8 @@
   </url><url>
     <loc>https://isa-afp.org/authors/brandt/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/krayer/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/fermat3_4/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/fft/</loc>
@@ -3942,6 +3991,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/hol-csp_opsem/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/hol-csp_rs/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/hol-cspm/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/hol-data_structures/</loc>
@@ -4264,6 +4315,8 @@
   </url><url>
     <loc>https://isa-afp.org/authors/liu/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/taskin/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/kad/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/engelhardt/</loc>
@@ -4550,6 +4603,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/markov_models/</loc>
   </url><url>
+    <loc>https://isa-afp.org/authors/gschossmann/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/olm/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/rabe/</loc>
@@ -4778,6 +4833,10 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/multitape_to_singletape_tm/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/munta_base/</loc>
+  </url><url>
+    <loc>https://isa-afp.org/sessions/munta_model_checker/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/myhill-nerode/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/zhann/</loc>
@@ -5216,6 +5275,14 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/resolution_fol/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/restriction_spaces/</loc>
+  </url><url>
+    <loc>https://isa-afp.org/sessions/restriction_spaces-examples/</loc>
+  </url><url>
+    <loc>https://isa-afp.org/sessions/restriction_spaces-holcf/</loc>
+  </url><url>
+    <loc>https://isa-afp.org/sessions/restriction_spaces-ultrametric/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/rewrite_properties_reduction/</loc>
   </url><url>
     <loc>https://isa-afp.org/sessions/rewriting_z/</loc>
@@ -5375,6 +5442,8 @@
   </url><url>
     <loc>https://isa-afp.org/sessions/shadow_sc_dom/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/shallow_expressions/</loc>
+  </url><url>
     <loc>https://isa-afp.org/authors/ying/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/pillai/</loc>
@@ -5443,6 +5512,8 @@
   </url><url>
     <loc>https://isa-afp.org/authors/tourret/</loc>
   </url><url>
+    <loc>https://isa-afp.org/sessions/sophie_germain/</loc>
+  </url><url>
     <loc>https://isa-afp.org/sessions/sophomores_dream/</loc>
   </url><url>
     <loc>https://isa-afp.org/authors/debois/</loc>
diff -r 848f2c00e57f -r 29846bc89de2 web/statistics/index.html
--- web/statistics/index.html
+++ web/statistics/index.html
@@ -88,22 +88,22 @@
       <div>
   <table>
     <tr>
-      <td class="statsnumber">901</td>
+      <td class="statsnumber">912</td>
       <td><a href="../">Entries</a></td>
     </tr>
     <tr>
-      <td class="statsnumber">537</td>
+      <td class="statsnumber">543</td>
       <td><a href="../authors/">Authors</a></td>
     </tr>
     <tr>
       <td class="statsnumber">
-        ~289,900
+        ~294,900
       </td>
       <td>Lemmas</td>
     </tr>
     <tr>
       <td class="statsnumber">
-        ~4,802,100
+        ~4,872,000
       </td>
       <td>Lines of Code</td>
     </tr>
@@ -124,7 +124,7 @@
       <tr>
         <td>2.</td>
         <td><a href="../entries/Show.html">Haskell's Show Class in Isabelle/HOL</a></td>
-        <td>20</td>
+        <td>21</td>
       </tr>
       <tr>
         <td>3.</td>
@@ -143,25 +143,25 @@
       </tr>
       <tr>
         <td>6.</td>
-        <td><a href="../entries/Deriving.html">Deriving class instances for datatypes</a></td>
+        <td><a href="../entries/Polynomial_Factorization.html">Polynomial Factorization</a></td>
         <td>15</td>
       </tr>
       <tr>
         <td>7.</td>
+        <td><a href="../entries/Deriving.html">Deriving class instances for datatypes</a></td>
+        <td>15</td>
+      </tr>
+      <tr>
+        <td>8.</td>
         <td><a href="../entries/Regular-Sets.html">Regular Sets and Expressions</a></td>
         <td>14</td>
       </tr>
       <tr>
-        <td>8.</td>
+        <td>9.</td>
         <td><a href="../entries/Polynomial_Interpolation.html">Polynomial Interpolation</a></td>
         <td>14</td>
       </tr>
       <tr>
-        <td>9.</td>
-        <td><a href="../entries/Polynomial_Factorization.html">Polynomial Factorization</a></td>
-        <td>14</td>
-      </tr>
-      <tr>
         <td>10.</td>
         <td><a href="../entries/Refine_Imperative_HOL.html">The Imperative Refinement Framework</a></td>
         <td>13</td>
@@ -212,29 +212,29 @@
 721,
 792,
 875,
-901]
-    const no_loc = [59041,
-94446,
-128626,
-235978,
-351521,
-433780,
-515160,
-566223,
-736026,
-823120,
-1035113,
-1215762,
-1598656,
-1854867,
-2130127,
-2453790,
-2854773,
-3454528,
-3818491,
-4163168,
-4695661,
-4802073 ]
+912]
+    const no_loc = [59006,
+94411,
+128591,
+235943,
+351486,
+433745,
+515124,
+566187,
+735990,
+823083,
+1035076,
+1215735,
+1598663,
+1854879,
+2130140,
+2453838,
+2854822,
+3454569,
+3818526,
+4163196,
+4695384,
+4872036 ]
     const no_authors = [14,
 11,
 6,
@@ -256,7 +256,7 @@
 31,
 28,
 45,
-18]
+24]
     const no_authors_series = [14,
 25,
 31,
@@ -278,7 +278,7 @@
 442,
 470,
 515,
-533]
+539]
     const all_articles = ['AVL-Trees',
 'MiniML',
 'Functional-Automata',
@@ -1175,11 +1175,22 @@
 'GyrovectorSpaces',
 'Dilworth',
 'Kolmogorov_Chentsov',
+'Sophie_Germain',
+'Elementary_Ultrametric_Spaces',
 'Completeness_Decreasing_Diagrams_for_N1',
 'First_Order_Rewriting',
 'Proof_Terms_Term_Rewriting',
 'Dedekind_Sums',
-'Pentagonal_Number_Theorem']
+'Pentagonal_Number_Theorem',
+'HOL-CSP_RS',
+'Restriction_Spaces',
+'Restriction_Spaces-Examples',
+'Restriction_Spaces-Ultrametric',
+'Farey_Sequences',
+'Shallow_Expressions',
+'Context_Free_Grammar',
+'Elliptic_Functions',
+'Munta_Model_Checker']
     const article_years_unique = ['2004',
 '',
 '',
@@ -2080,10 +2091,21 @@
 '',
 '',
 '',
+'',
+'',
+'',
+'',
+'',
+'',
+'',
+'',
+'',
+'',
+'',
 '']
     const loc_articles = [839,
 1249,
-1457,
+1422,
 528,
 1055,
 2419,
@@ -2145,7 +2167,7 @@
 5025,
 4380,
 208,
-9538,
+9537,
 495,
 2380,
 3399,
@@ -2213,7 +2235,7 @@
 4355,
 1251,
 1917,
-6218,
+6217,
 4974,
 10114,
 7252,
@@ -2292,15 +2314,15 @@
 2796,
 1116,
 5256,
-4867,
-8846,
+4868,
+8843,
 1356,
 5986,
 527,
-6653,
+6654,
 2584,
 1772,
-5327,
+5338,
 1092,
 4123,
 952,
@@ -2337,9 +2359,9 @@
 838,
 3622,
 4616,
-6332,
+6333,
 3995,
-8166,
+8178,
 12084,
 3188,
 518,
@@ -2385,7 +2407,7 @@
 2459,
 4344,
 9356,
-20052,
+20073,
 3725,
 3419,
 319,
@@ -2438,7 +2460,7 @@
 2103,
 639,
 14044,
-3156,
+3158,
 3930,
 4869,
 468,
@@ -2459,7 +2481,7 @@
 250,
 10663,
 822,
-7432,
+7435,
 3123,
 5219,
 2782,
@@ -2483,7 +2505,7 @@
 6489,
 9963,
 787,
-9570,
+9572,
 3249,
 1815,
 8443,
@@ -2518,7 +2540,7 @@
 9770,
 2765,
 641,
-11918,
+11917,
 2206,
 1743,
 7909,
@@ -2527,7 +2549,7 @@
 685,
 1812,
 1235,
-3559,
+3558,
 3579,
 2954,
 2218,
@@ -2535,10 +2557,10 @@
 5182,
 4968,
 2767,
-17369,
+17370,
 34348,
 3235,
-6016,
+6015,
 1891,
 373,
 10016,
@@ -2564,7 +2586,7 @@
 471,
 4916,
 3199,
-13320,
+13319,
 987,
 1002,
 4455,
@@ -2577,7 +2599,7 @@
 2090,
 3734,
 5783,
-2350,
+2387,
 4324,
 3804,
 176,
@@ -2607,7 +2629,7 @@
 2564,
 333,
 28070,
-24138,
+24140,
 10943,
 3049,
 744,
@@ -2635,8 +2657,8 @@
 5787,
 12873,
 9015,
-4257,
-4710,
+4260,
+4709,
 426,
 11477,
 3546,
@@ -2644,13 +2666,13 @@
 8100,
 16378,
 3523,
-7792,
+7794,
 12655,
 15363,
 650,
 1770,
 2352,
-16435,
+16430,
 1295,
 7697,
 3998,
@@ -2669,7 +2691,7 @@
 4039,
 3764,
 4958,
-16839,
+16827,
 8442,
 15823,
 6578,
@@ -2685,7 +2707,7 @@
 828,
 666,
 840,
-19737,
+19736,
 1088,
 232,
 4435,
@@ -2709,9 +2731,9 @@
 169,
 17517,
 4706,
-19445,
+19446,
 1306,
-21686,
+21690,
 13533,
 2627,
 1324,
@@ -2756,7 +2778,7 @@
 816,
 401,
 1313,
-1820,
+1819,
 5167,
 2337,
 2169,
@@ -2773,7 +2795,7 @@
 5782,
 16082,
 1040,
-53725,
+53723,
 351,
 3174,
 1961,
@@ -2795,12 +2817,12 @@
 764,
 2052,
 419,
-4086,
+4085,
 23189,
 1844,
 135,
 888,
-3535,
+3533,
 8765,
 2566,
 1285,
@@ -2809,7 +2831,7 @@
 1105,
 2001,
 1036,
-6113,
+6117,
 3226,
 4767,
 6361,
@@ -2817,10 +2839,10 @@
 1739,
 7872,
 8371,
-9379,
+9403,
 818,
 5861,
-7930,
+7896,
 6198,
 9145,
 1022,
@@ -2849,7 +2871,7 @@
 1551,
 1749,
 259,
-3531,
+3530,
 3346,
 1178,
 17080,
@@ -2895,7 +2917,7 @@
 7752,
 22884,
 8247,
-98618,
+98615,
 6811,
 1014,
 8009,
@@ -2929,24 +2951,24 @@
 862,
 713,
 99,
-10108,
+10107,
 9148,
 1307,
 17604,
 1636,
 253,
-1582,
+1589,
 2647,
 1180,
 908,
-21914,
+21919,
 1999,
 457,
 13232,
 4758,
 1143,
-5589,
-24951,
+5413,
+24953,
 3101,
 761,
 772,
@@ -2955,21 +2977,21 @@
 1367,
 3325,
 749,
-3754,
+3615,
 381,
 8570,
 1157,
 4164,
 3442,
 3650,
-5078,
+6279,
 233,
 1091,
 7406,
 4596,
 260,
 431,
-1379,
+1380,
 3758,
 3622,
 6805,
@@ -2977,11 +2999,22 @@
 7509,
 1590,
 3428,
+1643,
+339,
 13238,
 9872,
 10037,
 1304,
-1543]
+1543,
+3432,
+2719,
+1894,
+1877,
+1177,
+1534,
+5918,
+13604,
+34901]
   </script>
   <h4>Growth in number of entries:</h4>
   <script src="../js/Chart.js"></script>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/automata-and-formal-languages/index.html
--- web/topics/computer-science/automata-and-formal-languages/index.html
+++ web/topics/computer-science/automata-and-formal-languages/index.html
@@ -92,6 +92,31 @@
     <h2 class="head">2025</h2>
       <article class="entry">
         <div class="item-text">
+          <h5><a class="title" href="../../../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></h5>
+          <br>
+          by
+            <a href="../../../authors/wimmer/">Simon Wimmer</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Context_Free_Grammar.html">Context-Free Grammars and Languages</a></h5>
+          <br>
+          by
+            <a href="../../../authors/nipkow/">Tobias Nipkow</a>, 
+            <a href="../../../authors/gschossmann/">Markus Gschoßmann</a>, 
+            <a href="../../../authors/krayer/">Felix Krayer</a>, 
+            <a href="../../../authors/lehr/">Fabian Lehr</a>, 
+            <a href="../../../authors/philipp/">Bruno Philipp</a>, 
+            <a href="../../../authors/stimpfle/">August Martin Stimpfle</a>, 
+            <a href="../../../authors/taskin/">Kaan Taskin</a> and 
+            <a href="../../../authors/yamada/">Akihisa Yamada</a>
+        </div>
+        <span class="date">May 21</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
           <h5><a class="title" href="../../../entries/Mission_Time_LTL_Language_Partition.html">Language Partitioning for Mission-time Linear Temporal Logic</a></h5>
           <br>
           by
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/automata-and-formal-languages/index.xml
--- web/topics/computer-science/automata-and-formal-languages/index.xml
+++ web/topics/computer-science/automata-and-formal-languages/index.xml
@@ -4,6 +4,21 @@
     <title>Computer science/Automata and formal languages</title>
     <link>https://isa-afp.org/topics/computer-science/automata-and-formal-languages/</link>
     <description>AFP entries in Automata and formal languages</description><item>
+  <title>Munta: A Verified Model Checker for Timed Automata</title>
+  <link>https://isa-afp.org/entries/Munta_Model_Checker.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Munta: A Verified Model Checker for Timed Automata in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+    <category>Tools</category>
+</item>
+<item>
+  <title>Context-Free Grammars and Languages</title>
+  <link>https://isa-afp.org/entries/Context_Free_Grammar.html</link>
+  <pubDate>Wed, 21 May 2025 00:00:00 +0000</pubDate>
+  <description>Context-Free Grammars and Languages in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+</item>
+<item>
   <title>Language Partitioning for Mission-time Linear Temporal Logic</title>
   <link>https://isa-afp.org/entries/Mission_Time_LTL_Language_Partition.html</link>
   <pubDate>Mon, 03 Mar 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/concurrency/process-calculi/index.html
--- web/topics/computer-science/concurrency/process-calculi/index.html
+++ web/topics/computer-science/concurrency/process-calculi/index.html
@@ -89,6 +89,28 @@
     <h2>Subject Classification</h2>
       <p>ACM: <a href="https://dl.acm.org/topic/ccs2012/10003752.10003753.10003761.10003764">Theory of computation~Process calculi</a></p>
       <p>AMS: <a href="https://mathscinet.ams.org/mathscinet/msc/msc2020.html?t=68Q85">Computer science / Theory of computing / Models and methods for concurrent and distributed computing (process algebras, bisimulation, transition nets, etc.)</a></p>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../../authors/ballenghien/">Benoît Ballenghien</a> and 
+            <a href="../../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/concurrency/process-calculi/index.xml
--- web/topics/computer-science/concurrency/process-calculi/index.xml
+++ web/topics/computer-science/concurrency/process-calculi/index.xml
@@ -4,6 +4,24 @@
     <title>Computer science/Concurrency/Process calculi</title>
     <link>https://isa-afp.org/topics/computer-science/concurrency/process-calculi/</link>
     <description>AFP entries in Process calculi</description><item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>CSP Semantics over Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/HOL-CSP_RS.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>CSP Semantics over Restriction Spaces in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
   <title>Linear Resources and Process Compositions</title>
   <link>https://isa-afp.org/entries/ProcessComposition.html</link>
   <pubDate>Mon, 25 Nov 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/programming-languages/logics/index.html
--- web/topics/computer-science/programming-languages/logics/index.html
+++ web/topics/computer-science/programming-languages/logics/index.html
@@ -89,6 +89,16 @@
     <h2>Subject Classification</h2>
       <p>ACM: <a href="https://dl.acm.org/topic/ccs2012/10003752.10003790.10003806">Theory of computation~Programming logic</a></p>
       <p>AMS: <a href="https://mathscinet.ams.org/mathscinet/msc/msc2020.html?t=68Q60">Computer science / Theory of computing / Specification and verification (program logics, model checking, etc.)</a></p>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../../entries/Shallow_Expressions.html">Shallow Expressions</a></h5>
+          <br>
+          by
+            <a href="../../../../authors/fosters/">Simon Foster</a>
+        </div>
+        <span class="date">May 16</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/programming-languages/logics/index.xml
--- web/topics/computer-science/programming-languages/logics/index.xml
+++ web/topics/computer-science/programming-languages/logics/index.xml
@@ -4,6 +4,14 @@
     <title>Computer science/Programming languages/Logics</title>
     <link>https://isa-afp.org/topics/computer-science/programming-languages/logics/</link>
     <description>AFP entries in Logics</description><item>
+  <title>Shallow Expressions</title>
+  <link>https://isa-afp.org/entries/Shallow_Expressions.html</link>
+  <pubDate>Fri, 16 May 2025 00:00:00 +0000</pubDate>
+  <description>Shallow Expressions in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Programming languages/Logics</category>
+</item>
+<item>
   <title>AutoCorres2</title>
   <link>https://isa-afp.org/entries/AutoCorres2.html</link>
   <pubDate>Wed, 17 Apr 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/semantics-and-reasoning/index.html
--- web/topics/computer-science/semantics-and-reasoning/index.html
+++ web/topics/computer-science/semantics-and-reasoning/index.html
@@ -92,6 +92,47 @@
     <h2 class="head">2025</h2>
       <article class="entry">
         <div class="item-text">
+          <h5><a class="title" href="../../../entries/Shallow_Expressions.html">Shallow Expressions</a></h5>
+          <br>
+          by
+            <a href="../../../authors/fosters/">Simon Foster</a>
+        </div>
+        <span class="date">May 16</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/HOL-CSP_RS.html">CSP Semantics over Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
           <h5><a class="title" href="../../../entries/Differential_Privacy.html">Differential Privacy</a></h5>
           <br>
           by
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/computer-science/semantics-and-reasoning/index.xml
--- web/topics/computer-science/semantics-and-reasoning/index.xml
+++ web/topics/computer-science/semantics-and-reasoning/index.xml
@@ -4,6 +4,41 @@
     <title>Computer science/Semantics and reasoning</title>
     <link>https://isa-afp.org/topics/computer-science/semantics-and-reasoning/</link>
     <description>AFP entries in Semantics and reasoning</description><item>
+  <title>Shallow Expressions</title>
+  <link>https://isa-afp.org/entries/Shallow_Expressions.html</link>
+  <pubDate>Fri, 16 May 2025 00:00:00 +0000</pubDate>
+  <description>Shallow Expressions in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Programming languages/Logics</category>
+</item>
+<item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>CSP Semantics over Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/HOL-CSP_RS.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>CSP Semantics over Restriction Spaces in the AFP</description>
+    <category>Computer science/Semantics and reasoning</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
   <title>Differential Privacy</title>
   <link>https://isa-afp.org/entries/Differential_Privacy.html</link>
   <pubDate>Tue, 07 Jan 2025 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/index.html
--- web/topics/index.html
+++ web/topics/index.html
@@ -146,7 +146,7 @@
           </ul>
           <li>
             <h3>
-              <a href="../topics/computer-science/automata-and-formal-languages/">Automata and formal languages (62)
+              <a href="../topics/computer-science/automata-and-formal-languages/">Automata and formal languages (64)
               </a>
             </h3>
           </li>
@@ -160,7 +160,7 @@
           </li>
           <ul>
               <li>
-                <a href="../topics/computer-science/concurrency/process-calculi/">Process calculi (17)
+                <a href="../topics/computer-science/concurrency/process-calculi/">Process calculi (19)
                 </a>
               </li>
           </ul>
@@ -232,7 +232,7 @@
                 </a>
               </li>
               <li>
-                <a href="../topics/computer-science/programming-languages/logics/">Logics (35)
+                <a href="../topics/computer-science/programming-languages/logics/">Logics (36)
                 </a>
               </li>
               <li>
@@ -262,7 +262,7 @@
           </ul>
           <li>
             <h3>
-              <a href="../topics/computer-science/semantics-and-reasoning/">Semantics and reasoning (28)
+              <a href="../topics/computer-science/semantics-and-reasoning/">Semantics and reasoning (32)
               </a>
             </h3>
           </li>
@@ -364,7 +364,7 @@
       <ul>
           <li>
             <h3>
-              <a href="../topics/mathematics/algebra/">Algebra (100)
+              <a href="../topics/mathematics/algebra/">Algebra (101)
               </a>
             </h3>
           </li>
@@ -372,7 +372,7 @@
           </ul>
           <li>
             <h3>
-              <a href="../topics/mathematics/analysis/">Analysis (59)
+              <a href="../topics/mathematics/analysis/">Analysis (62)
               </a>
             </h3>
           </li>
@@ -436,7 +436,7 @@
           </ul>
           <li>
             <h3>
-              <a href="../topics/mathematics/number-theory/">Number theory (55)
+              <a href="../topics/mathematics/number-theory/">Number theory (58)
               </a>
             </h3>
           </li>
@@ -472,14 +472,14 @@
           </ul>
           <li>
             <h3>
-              <a href="../topics/mathematics/topology/">Topology (10)
+              <a href="../topics/mathematics/topology/">Topology (14)
               </a>
             </h3>
           </li>
           <ul>
           </ul>
       </ul>
-      <h2><a href="../topics/tools/">Tools (27)</a></h2>
+      <h2><a href="../topics/tools/">Tools (29)</a></h2>
       <ul>
       </ul>
       </div>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/algebra/index.html
--- web/topics/mathematics/algebra/index.html
+++ web/topics/mathematics/algebra/index.html
@@ -88,6 +88,18 @@
       <div>
     <h2>Subject Classification</h2>
       <p>AMS: <a href="https://mathscinet.ams.org/mathscinet/msc/msc2020.html?t=08-XX">General algebraic systems</a></p>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/algebra/index.xml
--- web/topics/mathematics/algebra/index.xml
+++ web/topics/mathematics/algebra/index.xml
@@ -4,6 +4,16 @@
     <title>Mathematics/Algebra</title>
     <link>https://isa-afp.org/topics/mathematics/algebra/</link>
     <description>AFP entries in Algebra</description><item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
   <title>A Generalization of the Cauchy–Davenport theorem</title>
   <link>https://isa-afp.org/entries/Generalized_Cauchy_Davenport.html</link>
   <pubDate>Sat, 23 Nov 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/analysis/index.html
--- web/topics/mathematics/analysis/index.html
+++ web/topics/mathematics/analysis/index.html
@@ -88,6 +88,40 @@
       <div>
     <h2>Subject Classification</h2>
       <p>ACM: <a href="https://dl.acm.org/topic/ccs2012/10002950.10003714">Mathematics of computing~Mathematical analysis</a></p>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/analysis/index.xml
--- web/topics/mathematics/analysis/index.xml
+++ web/topics/mathematics/analysis/index.xml
@@ -4,6 +4,35 @@
     <title>Mathematics/Analysis</title>
     <link>https://isa-afp.org/topics/mathematics/analysis/</link>
     <description>AFP entries in Analysis</description><item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
   <title>The Tensor Product on Hilbert Spaces</title>
   <link>https://isa-afp.org/entries/Hilbert_Space_Tensor_Product.html</link>
   <pubDate>Tue, 10 Sep 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/number-theory/index.html
--- web/topics/mathematics/number-theory/index.html
+++ web/topics/mathematics/number-theory/index.html
@@ -91,6 +91,27 @@
     <h2 class="head">2025</h2>
       <article class="entry">
         <div class="item-text">
+          <h5><a class="title" href="../../../entries/Elliptic_Functions.html">Complex Lattices, Elliptic Functions, and the Modular Group</a></h5>
+          <br>
+          by
+            <a href="../../../authors/eberl/">Manuel Eberl</a>, 
+            <a href="../../../authors/bordg/">Anthony Bordg</a>, 
+            <a href="../../../authors/li/">Wenda Li</a> and 
+            <a href="../../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Farey_Sequences.html">Farey Sequences and Ford Circles</a></h5>
+          <br>
+          by
+            <a href="../../../authors/paulson/">Lawrence C. Paulson</a>
+        </div>
+        <span class="date">May 15</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
           <h5><a class="title" href="../../../entries/Pentagonal_Number_Theorem.html">The Partition Function and the Pentagonal Number Theorem</a></h5>
           <br>
           by
@@ -110,6 +131,15 @@
         </div>
         <span class="date">Apr 23</span>
       </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Sophie_Germain.html">Sophie Germain’s Theorem</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>
+        </div>
+        <span class="date">Apr 09</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
@@ -219,7 +249,7 @@
           <h5><a class="title" href="../../../entries/Three_Squares.html">Three Squares Theorem</a></h5>
           <br>
           by
-            <a href="../../../authors/danilkin/">Anton Danilkin</a> and 
+            <a href="../../../authors/danilkin/">Anna Danilkin</a> and 
             <a href="../../../authors/chevalier/">Loïc Chevalier</a>
         </div>
         <span class="date">May 03</span>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/number-theory/index.xml
--- web/topics/mathematics/number-theory/index.xml
+++ web/topics/mathematics/number-theory/index.xml
@@ -4,6 +4,20 @@
     <title>Mathematics/Number theory</title>
     <link>https://isa-afp.org/topics/mathematics/number-theory/</link>
     <description>AFP entries in Number theory</description><item>
+  <title>Complex Lattices, Elliptic Functions, and the Modular Group</title>
+  <link>https://isa-afp.org/entries/Elliptic_Functions.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Complex Lattices, Elliptic Functions, and the Modular Group in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
+  <title>Farey Sequences and Ford Circles</title>
+  <link>https://isa-afp.org/entries/Farey_Sequences.html</link>
+  <pubDate>Thu, 15 May 2025 00:00:00 +0000</pubDate>
+  <description>Farey Sequences and Ford Circles in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>The Partition Function and the Pentagonal Number Theorem</title>
   <link>https://isa-afp.org/entries/Pentagonal_Number_Theorem.html</link>
   <pubDate>Wed, 23 Apr 2025 00:00:00 +0000</pubDate>
@@ -18,6 +32,13 @@
     <category>Mathematics/Number theory</category>
 </item>
 <item>
+  <title>Sophie Germain’s Theorem</title>
+  <link>https://isa-afp.org/entries/Sophie_Germain.html</link>
+  <pubDate>Wed, 09 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Sophie Germain’s Theorem in the AFP</description>
+    <category>Mathematics/Number theory</category>
+</item>
+<item>
   <title>Theta Functions</title>
   <link>https://isa-afp.org/entries/Theta_Functions.html</link>
   <pubDate>Mon, 02 Dec 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/topology/index.html
--- web/topics/mathematics/topology/index.html
+++ web/topics/mathematics/topology/index.html
@@ -89,6 +89,51 @@
     <h2>Subject Classification</h2>
       <p>ACM: <a href="https://dl.acm.org/topic/ccs2012/10002950.10003741.10003742">Mathematics of computing~Topology</a></p>
       <p>AMS: <a href="https://mathscinet.ams.org/mathscinet/msc/msc2020.html?t=54-XX">General topology</a></p>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Ultrametric.html">Ultrametric Structure for Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces-Examples.html">Examples of Restriction Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../../entries/Elementary_Ultrametric_Spaces.html">Definition and Elementary Properties of Ultrametric Spaces</a></h5>
+          <br>
+          by
+            <a href="../../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">Apr 10</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/mathematics/topology/index.xml
--- web/topics/mathematics/topology/index.xml
+++ web/topics/mathematics/topology/index.xml
@@ -4,6 +4,42 @@
     <title>Mathematics/Topology</title>
     <link>https://isa-afp.org/topics/mathematics/topology/</link>
     <description>AFP entries in Topology</description><item>
+  <title>Ultrametric Structure for Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Ultrametric.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Ultrametric Structure for Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Examples of Restriction Spaces</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces-Examples.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Examples of Restriction Spaces in the AFP</description>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Mathematics/Algebra</category>
+    <category>Computer science/Concurrency/Process calculi</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
+  <title>Definition and Elementary Properties of Ultrametric Spaces</title>
+  <link>https://isa-afp.org/entries/Elementary_Ultrametric_Spaces.html</link>
+  <pubDate>Thu, 10 Apr 2025 00:00:00 +0000</pubDate>
+  <description>Definition and Elementary Properties of Ultrametric Spaces in the AFP</description>
+    <category>Mathematics/Topology</category>
+</item>
+<item>
   <title>The Riesz Representation Theorem</title>
   <link>https://isa-afp.org/entries/Riesz_Representation.html</link>
   <pubDate>Tue, 04 Jun 2024 00:00:00 +0000</pubDate>
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/tools/index.html
--- web/topics/tools/index.html
+++ web/topics/tools/index.html
@@ -86,6 +86,27 @@
   </div>
 </header>
       <div>
+    <h2 class="head">2025</h2>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../entries/Munta_Model_Checker.html">Munta: A Verified Model Checker for Timed Automata</a></h5>
+          <br>
+          by
+            <a href="../../authors/wimmer/">Simon Wimmer</a>
+        </div>
+        <span class="date">May 22</span>
+      </article>
+      <article class="entry">
+        <div class="item-text">
+          <h5><a class="title" href="../../entries/Restriction_Spaces.html">Restriction Spaces: a Fixed-Point Theory</a></h5>
+          <br>
+          by
+            <a href="../../authors/ballenghien/">Benoît Ballenghien</a>, 
+            <a href="../../authors/puyobro/">Benjamin Puyobro</a> and 
+            <a href="../../authors/wolff/">Burkhart Wolff</a>
+        </div>
+        <span class="date">May 07</span>
+      </article>
     <h2 class="head">2024</h2>
       <article class="entry">
         <div class="item-text">
diff -r 848f2c00e57f -r 29846bc89de2 web/topics/tools/index.xml
--- web/topics/tools/index.xml
+++ web/topics/tools/index.xml
@@ -4,6 +4,24 @@
     <title>Tools</title>
     <link>https://isa-afp.org/topics/tools/</link>
     <description>AFP entries in Tools</description><item>
+  <title>Munta: A Verified Model Checker for Timed Automata</title>
+  <link>https://isa-afp.org/entries/Munta_Model_Checker.html</link>
+  <pubDate>Thu, 22 May 2025 00:00:00 +0000</pubDate>
+  <description>Munta: A Verified Model Checker for Timed Automata in the AFP</description>
+    <category>Computer science/Automata and formal languages</category>
+    <category>Tools</category>
+</item>
+<item>
+  <title>Restriction Spaces: a Fixed-Point Theory</title>
+  <link>https://isa-afp.org/entries/Restriction_Spaces.html</link>
+  <pubDate>Wed, 07 May 2025 00:00:00 +0000</pubDate>
+  <description>Restriction Spaces: a Fixed-Point Theory in the AFP</description>
+    <category>Tools</category>
+    <category>Mathematics/Analysis</category>
+    <category>Mathematics/Topology</category>
+    <category>Computer science/Semantics and reasoning</category>
+</item>
+<item>
   <title>Without Loss of Generality</title>
   <link>https://isa-afp.org/entries/Wlog.html</link>
   <pubDate>Fri, 30 Aug 2024 00:00:00 +0000</pubDate>