Theory Safety_Logic

(*<*)
theory Safety_Logic
imports
  Aczel_Sequences
  Galois
  Heyting
  Lifted_Predicates
begin

(*>*)
section‹ Safety logic\label{sec:safety_logic} ›

text‹

Following citet"AbadiPlotkin:1991" and "AbadiPlotkin:1993" and "AbadiLamport:1995"
(see also citet‹\S5.5› in "AbadiMerz:1996"), we work in the complete lattice of
stuttering-closed safety properties (i.e., stuttering-closed downsets)
and use this for logical purposes. We avoid many syntactic issues
via a shallow embedding into HOL.

›


subsection‹ Stuttering\label{sec:safety_logic-stuttering} ›

text‹

We define ‹stuttering equivalence› ala citet"Lamport:1994". This allows any
agent to repeat any state at any time. We define a normalisation function (♮)› on
typ('a, 's, 'v) trace.t and extract the (matroidal) closure over sets of
these from the Galois connection localegalois.image_vimage.

›

setup Sign.mandatory_path "trace"

primrec natural' :: "'s  ('a × 's) list  ('a × 's) list" where
  "natural' s [] = []"
| "natural' s (x # xs) = (if snd x = s then natural' s xs else x # natural' (snd x) xs)"

setup Sign.mandatory_path "final'"

lemma natural'[simp]:
  shows "trace.final' s (trace.natural' s xs) = trace.final' s xs"
by (induct xs arbitrary: s) simp_all

lemma natural'_cong:
  assumes "s = s'"
  assumes "trace.natural' s xs = trace.natural' s xs'"
  shows "trace.final' s xs = trace.final' s' xs'"
by (metis assms trace.final'.natural')

setup Sign.parent_path

setup Sign.mandatory_path "natural'"

lemma natural':
  shows "trace.natural' s (trace.natural' s xs) = trace.natural' s xs"
by (induct xs arbitrary: s) simp_all

lemma length:
  shows "length (trace.natural' s xs)  length xs"
by (induct xs arbitrary: s) (simp_all add: le_SucI)

lemma subseq:
  shows "subseq (trace.natural' s xs) xs"
by (induct xs arbitrary: s) auto

lemma remdups_adj:
  shows "s # map snd (trace.natural' s xs) = remdups_adj (s # map snd xs)"
by (induct xs arbitrary: s) simp_all

lemma append:
  shows "trace.natural' s (xs @ ys) = trace.natural' s xs @ trace.natural' (trace.final' s xs) ys"
by (induct xs arbitrary: s) simp_all

lemma eq_Nil_conv:
  shows "trace.natural' s xs = []  snd ` set xs  {s}"
    and "[] = trace.natural' s xs  snd ` set xs  {s}"
by (induct xs arbitrary: s) simp_all

lemma eq_Cons_conv:
  shows "trace.natural' s xs = y # ys
     (xs' ys'. xs = xs' @ y # ys'  snd ` set xs'  {s}  snd y  s  trace.natural' (snd y) ys' = ys)" (is "?lhs  ?rhs")
    and "y # ys = trace.natural' s xs
     (xs' ys'. xs = xs' @ y # ys'  snd ` set xs'  {s}  snd y  s  trace.natural' (snd y) ys' = ys)" (is ?thesis1)
proof -
  show "?lhs  ?rhs"
  proof(rule iffI)
    show "?lhs  ?rhs"
    proof(induct xs)
      case (Cons x xs) show ?case
      proof(cases "s = snd x")
        case True with Cons
        obtain xs' ys'
         where "xs = xs' @ y # ys'" and "snd ` set xs'  {s}" and "snd y  s"
           and "trace.natural' (snd y) ys' = ys"
          by auto
        with True show ?thesis by (simp add: exI[where x="x # xs'"])
      qed (use Cons.prems in force)
    qed simp
    show "?rhs  ?lhs"
      by (auto simp: trace.natural'.append trace.natural'.eq_Nil_conv)
  qed
  then show ?thesis1
    by (rule eq_commute_conv)
qed

lemma eq_append_conv:
  shows "trace.natural' s xs = ys @ zs
      (ys' zs'. xs = ys' @ zs'  trace.natural' s ys' = ys  trace.natural' (trace.final' s ys) zs' = zs)" (is "?lhs = ?rhs")
    and "ys @ zs = trace.natural' s xs
      (ys' zs'. xs = ys' @ zs'  trace.natural' s ys' = ys  trace.natural' (trace.final' s ys) zs' = zs)" (is ?thesis1)
proof -
  show "?lhs  ?rhs"
  proof(rule iffI)
    show "?lhs  ?rhs"
    proof(induct ys arbitrary: s xs)
      case (Cons y ys s xs)
      from Cons.prems
      obtain ys' zs'
       where "xs = ys' @ y # zs'" and "snd ` set ys'  {s}"
         and "snd y  s" and "trace.natural' (snd y) zs' = ys @ zs"
        by (clarsimp simp: trace.natural'.eq_Cons_conv)
      with Cons.hyps[where s="snd y" and xs="zs'"] show ?case
        by (clarsimp simp: trace.natural'.eq_Cons_conv) (metis append.assoc append_Cons)
    qed fastforce
    show "?rhs  ?lhs"
      by (auto simp: trace.natural'.append)
  qed
  then show ?thesis1
    by (rule eq_commute_conv)
qed

lemma replicate:
  shows "trace.natural' s (replicate i as) = (if snd as = s  i = 0 then [] else [as])"
by (auto simp: gr0_conv_Suc trace.natural'.eq_Nil_conv)

lemma map_natural':
  shows "trace.natural' (sf s) (map (map_prod af sf) (trace.natural' s xs))
       = trace.natural' (sf s) (map (map_prod af sf) xs)"
by (induct xs arbitrary: s; simp; metis)

lemma map_inj_on_sf:
  assumes "inj_on sf (insert s (snd ` set xs))"
  shows "trace.natural' (sf s) (map (map_prod af sf) xs) = map (map_prod af sf) (trace.natural' s xs)"
using assms
proof(induct xs arbitrary: s)
  case (Cons x xs s)
  from Cons.prems have "sf (snd x)  sf s" if "snd x  s"
    by (meson image_eqI inj_onD insert_iff list.set_intros(1) that)
  with Cons.prems show ?case
    by (auto intro: Cons.hyps)
qed simp

lemma amap_noop:
  assumes "trace.natural' s xs = map (map_prod af id) zs"
  shows "trace.natural' s zs = zs"
using assms by (induct xs arbitrary: s zs) (auto split: if_split_asm)

lemma take:
  shows "jlength xs. take i (trace.natural' s xs) = trace.natural' s (take j xs)"
proof(induct xs arbitrary: s i)
  case (Cons x xs s i) then show ?case by (cases i; fastforce)
qed simp

lemma idle_prefix:
  assumes "snd ` set xs  {s}"
  shows "trace.natural' s (xs @ ys) = trace.natural' s ys"
using assms by (simp add: trace.natural'.append trace.natural'.eq_Nil_conv)

lemma prefixE:
  assumes "trace.natural' s ys = trace.natural' s (xs @ xsrest)"
  obtains xs' xs'rest where "trace.natural' s xs = trace.natural' s xs'" and "ys = xs' @ xs'rest"
by (metis assms trace.natural'.eq_append_conv(2))

lemma aset_conv:
  shows "a  trace.aset (trace.T s (trace.natural' s xs) v)
     (s' s''. (a, s', s'')  set (trace.transitions' s xs)  s'  s'')"
by (induct xs arbitrary: s) (auto simp: trace.aset.simps)

setup Sign.parent_path

definition natural :: "('a, 's, 'v) trace.t  ('a, 's, 'v) trace.t" ("") where
  "σ = trace.T (trace.init σ) (trace.natural' (trace.init σ) (trace.rest σ)) (trace.term σ)"

setup Sign.mandatory_path "natural"

lemma sel[simp]:
  shows "trace.init (σ) = trace.init σ"
    and "trace.rest (σ) = trace.natural' (trace.init σ) (trace.rest σ)"
    and "trace.term (σ) = trace.term σ"
by (simp_all add: trace.natural_def)

lemma simps:
  shows "(trace.T s [] v) = trace.T s [] v"
    and "(trace.T s ((a, s) # xs) v) = (trace.T s xs v)"
    and "(trace.T s (trace.natural' s xs) v) = (trace.T s xs v)"
by (simp_all add: trace.natural_def trace.natural'.natural')

lemma idempotent[simp]:
  shows "(σ) = σ"
by (simp add: trace.natural_def trace.natural'.natural')

lemma idle:
  assumes "snd ` set xs  {s}"
  shows "(trace.T s xs v) = trace.T s [] v"
by (simp add: assms trace.natural_def trace.natural'.eq_Nil_conv)

lemma trace_conv:
  shows "(trace.T s xs v) = σ  trace.init σ = s  trace.natural' s xs = trace.natural' s (trace.rest σ)  trace.term σ = v"
    and "σ = (trace.T s xs v)  trace.init σ = s  trace.natural' s xs = trace.natural' s (trace.rest σ)  trace.term σ = v"
by (cases σ; fastforce simp: trace.natural_def)+

lemma map_natural:
  shows "(trace.map af sf vf (σ)) = (trace.map af sf vf σ)"
by (simp add: trace.natural_def trace.natural'.map_natural')

lemma continue:
  shows "(σ @-S xsv) = σ @-S (trace.natural' (trace.final σ) (fst xsv), snd xsv)"
by (simp add: trace.natural_def trace.natural'.append split: option.split)

lemma replicate:
  shows "(trace.T s (replicate i as) v)
        = (trace.T s (if snd as = s  i = 0 then [] else [as]) v)"
by (simp add: trace.natural_def trace.natural'.replicate)

lemma monotone:
  shows "mono "
using trace.natural.continue by (fastforce intro: monoI simp: trace.less_eq_t_def)

lemmas strengthen[strg] = st_monotone[OF trace.natural.monotone]
lemmas mono = monotoneD[OF trace.natural.monotone]
lemmas mono2mono[cont_intro, partial_function_mono]
  = monotone2monotone[OF trace.natural.monotone, simplified, of orda P for orda P]

lemma less_eqE:
  assumes "t  u"
  assumes "u' = u"
  obtains t' where "t = t'" and "t'  u'"
using assms
by atomize_elim
   (fastforce simp: trace.natural_def trace.split_Ex trace.less_eq_None(2)[unfolded prefix_def]
             elim!: trace.less_eqE prefixE trace.natural'.prefixE)

lemma less_eq_natural:
  assumes "σ1  σ2"
  shows "σ1 = σ1"
using assms
by (cases σ1)
   (auto simp: trace.natural_def prefix_def trace.natural'.eq_append_conv trace.natural'.natural'
        elim!: trace.less_eqE)

lemma map_le:
  assumes "σ1  σ2"
  shows "(trace.map af sf vf σ1)  (trace.map af sf vf σ2)"
using trace.natural.mono[OF trace.map.mono[OF assms], simplified trace.natural.map_natural] .

lemma map_inj_on_sf:
  assumes "inj_on sf (trace.sset σ)"
  shows "(trace.map af sf vf σ) = trace.map af sf vf (σ)"
using assms by (cases σ) (simp add: trace.natural_def trace.natural'.map_inj_on_sf trace.sset.simps)

lemma take:
  shows "j. (trace.take i σ) = trace.take j (σ)"
by (meson trace.natural.mono trace.less_eq_take_def)

lemma take_natural:
  shows "(trace.take i (σ)) = trace.take i (σ)"
using trace.natural.less_eq_natural by blast

lemma takeE:
  shows "σ1 = (trace.take i σ2); j. σ1 = trace.take j (σ2)  thesis  thesis"
    and "(trace.take i σ2) = σ1; j. σ1 = trace.take j (σ2)  thesis  thesis"
using trace.natural.take by blast+

setup Sign.parent_path

setup Sign.mandatory_path "aset"

lemma natural_conv:
  shows "a  trace.aset (σ)  (s s'. (a, s, s')  trace.steps σ)"
by (simp add: trace.natural_def trace.steps'_def trace.natural'.aset_conv)

setup Sign.parent_path

setup Sign.mandatory_path "sset"

lemma natural'[simp]:
  shows "trace.sset (trace.T s0 (trace.natural' s0 xs) v) = trace.sset (trace.T s0 xs v)"
by (induct xs arbitrary: s0) (simp_all add: trace.sset.simps)

lemma natural[simp]:
  shows "trace.sset (σ) = trace.sset σ"
by (simp add: trace.natural_def)

setup Sign.parent_path

setup Sign.mandatory_path "vset"

lemma natural[simp]:
  shows "trace.vset (σ) = trace.vset σ"
by (cases σ) (simp add: trace.natural_def trace.t.simps(8))

setup Sign.parent_path

setup Sign.mandatory_path "take"

lemma natural:
  shows "jSuc (length (trace.rest σ)). trace.take i (σ) = (trace.take j σ)"
using trace.natural'.take[where i=i and s= "trace.init σ" and xs="trace.rest σ"]
by (auto simp: trace.natural_def trace.take_def not_le) (use Suc_n_not_le_n in blast)

lemma naturalE:
  shows "σ1 = trace.take i (σ2); j. j  Suc (length (trace.rest σ2)); σ1 = (trace.take j σ2)  thesis  thesis"
    and "trace.take i (σ2) = σ1; j. j  Suc (length (trace.rest σ2)); (trace.take j σ2) = σ1  thesis  thesis"
using trace.take.natural[of σ2 i] by force+

setup Sign.parent_path

lemma steps'_alt_def:
  shows "trace.steps' s xs = set (trace.transitions' s (trace.natural' s xs))"
by (induct xs arbitrary: s) auto

setup Sign.mandatory_path "steps'"

lemma natural':
  shows "trace.steps' s (trace.natural' s xs) = trace.steps' s xs"
unfolding trace.steps'_def by (induct xs arbitrary: s) auto

lemma asetD:
  assumes "trace.steps σ  r"
  shows "a. a  trace.aset (σ)  a  fst ` r"
using assms by (force simp: trace.aset.natural_conv)

lemma range_initE:
  assumes "trace.steps' s0 xs  range af × range sf × range sf"
  assumes "(a, s, s')  trace.steps' s0 xs"
  obtains s0' where "s0 = sf s0'"
using assms by (induct xs arbitrary: s s0) (auto simp: trace.steps'_alt_def split: if_split_asm)

lemma map_range_conv:
  shows "trace.steps' (sf s) xs  range af × range sf × range sf
      (xs'. trace.natural' (sf s) xs = map (map_prod af sf) xs')" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (induct xs arbitrary: s) (auto 0 3 simp: Cons_eq_map_conv)
  show "?rhs  ?lhs"
    by (force simp: trace.steps'_alt_def trace.transitions'.map map_prod_conv)
qed

lemma step_conv:
  shows "trace.steps' s xs = {x}
      fst (snd x) = s  fst (snd x)  snd (snd x)
        (ys zs. snd ` set ys  {s}  snd ` set zs  {snd (snd x)}
                 xs = ys @ [(fst x, snd (snd x))] @ zs)" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (fastforce dest!: arg_cong[where f="set"]
                   simp: trace.steps'_alt_def set_singleton_conv set_replicate_conv_if
                         trace.transitions'.eq_Nil_conv trace.transitions'.eq_Cons_conv
                         trace.natural'.eq_Nil_conv trace.natural'.eq_Cons_conv
                  split: if_split_asm)
  show "?rhs  ?lhs"
    by (clarsimp simp: trace.steps'.append)
qed

setup Sign.parent_path

interpretation stuttering: galois.image_vimage_idempotent ""
by (simp add: galois.image_vimage_idempotent.intro)

abbreviation stuttering_equiv_syn :: "('a, 's, 'v) trace.t  ('a, 's, 'v) trace.t  bool" (infix "S" 50) where
  "σ1 S σ2  trace.stuttering.equivalent σ1 σ2"

setup Sign.mandatory_path "stuttering"

setup Sign.mandatory_path "cl"

setup Sign.mandatory_path "downwards"

lemma cl:
  shows "trace.stuttering.cl (downwards.cl P) = downwards.cl (trace.stuttering.cl P)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (clarsimp simp: trace.stuttering.cl_alt_def downwards.cl_def trace.less_eq_t_def)
       (metis trace.final'.natural' trace.natural.continue trace.natural.sel(1,2))
next
  show "?rhs  ?lhs"
    by (clarsimp elim!: downwards.clE trace.stuttering.clE)
       (erule (1) trace.natural.less_eqE; fastforce)
qed

lemma closed:
  assumes "P  downwards.closed"
  shows "trace.stuttering.cl P  downwards.closed"
by (metis assms downwards.closedI downwards.closed_conv trace.stuttering.cl.downwards.cl)

setup Sign.parent_path

setup Sign.mandatory_path "closed"

lemma downwards_imp: ― ‹ citet‹p13› in "AbadiPlotkin:1993"
  assumes "P  trace.stuttering.closed"
  assumes "Q  trace.stuttering.closed"
  shows "downwards.imp P Q   trace.stuttering.closed"
by (simp add: assms downwards.closed_imp downwards.heyting_imp downwards.imp_mp
              trace.stuttering.cl.downwards.closed trace.stuttering.closed_clI
              trace.stuttering.exchange_closed_inter trace.stuttering.least)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "equiv"

lemma simps:
  shows "snd ` set xs  {s}  trace.T s (xs @ ys) v S trace.T s ys v"
    and "snd ` set ys  {trace.final' s xs}  trace.T s (xs @ ys) v S trace.T s xs v"
    and "snd ` set xs  {snd x}  trace.T s (x # xs @ ys) v S trace.T s (x # ys) v"
by (fastforce simp: trace.natural_def trace.natural'.append trace.natural'.eq_Nil_conv)+

lemma append_cong:
  assumes "s = s'"
  assumes "trace.natural' s xs = trace.natural' s xs'"
  assumes "trace.natural' (trace.final' s xs) ys = trace.natural' (trace.final' s xs) ys'"
  assumes "v = v'"
  shows "trace.T s (xs @ ys) v S trace.T s' (xs' @ ys') v'"
using assms by (simp add: trace.natural_def trace.natural'.append cong: trace.final'.natural'_cong)

lemma E:
  assumes "trace.T s xs v S trace.T s' xs' v'"
  obtains "trace.natural' s xs = trace.natural' s' xs'" and "s = s'" and "v = v'"
using assms by (fastforce simp: trace.natural_def)

lemma append_conv:
  shows "trace.T s (xs @ ys) v S σ
     (xs' ys'. σ = trace.T s (xs' @ ys') v  trace.natural' s xs = trace.natural' s xs'
                  trace.natural' (trace.final' s xs) ys = trace.natural' (trace.final' s xs) ys')" (is ?thesis1)
    and "σ S trace.T s (xs @ ys) v
     (xs' ys'. σ = trace.T s (xs' @ ys') v  trace.natural' s xs = trace.natural' s xs'
                  trace.natural' (trace.final' s xs) ys = trace.natural' (trace.final' s xs) ys')" (is ?thesis2)
proof -
  show ?thesis1
    by (cases σ)
       (fastforce simp: trace.natural'.append trace.natural'.eq_append_conv
                  elim: trace.stuttering.equiv.E
                 intro: trace.stuttering.equiv.append_cong)
  then show ?thesis2
    by (rule eq_commute_conv)
qed

lemma map:
  assumes "σ1 S σ2"
  shows "trace.map af sf vf σ1 S trace.map af sf vf σ2"
by (metis assms trace.natural.map_natural)

lemma steps:
  assumes "σ1 S σ2"
  shows "trace.steps σ1 = trace.steps σ2"
using assms by (force simp: trace.steps'_alt_def trace.natural_def)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.parent_path


subsection‹ The ‹('a, 's, 'v) spec› lattice \label{sec:safety_logic-logic} ›

text‹

Our workhorse lattice consists of all sets of traces that are downwards and stuttering closed. This
combined closure is neither matroidal nor antimatroidal (\S\ref{sec:closures-matroids}).

We define the lattice as a type and instantiate the relevant type classes. In the following
read P ≤ Q› (P ⊆ Q› in the powerset model) as ``Q follows from P'' or ``P entails Q''.

›

setup Sign.mandatory_path "raw"

setup Sign.mandatory_path "spec"

definition cl :: "('a, 's, 'v) trace.t set  ('a, 's, 'v) trace.t set" where
  "cl P = downwards.cl (trace.stuttering.cl P)"

setup Sign.parent_path

interpretation spec: closure_powerset_distributive raw.spec.cl
unfolding raw.spec.cl_def
by (simp add: closure_powerset_distributive_comp downwards.closure_powerset_distributive_axioms
              trace.stuttering.closure_powerset_distributive_axioms trace.stuttering.cl.downwards.cl)

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "cl"

lemma empty[simp]:
  shows "raw.spec.cl {} = {}"
by (simp add: raw.spec.cl_def downwards.cl_empty trace.stuttering.cl_bot)

setup Sign.parent_path

setup Sign.mandatory_path "closed"

lemma I:
  assumes "P  downwards.closed"
  assumes "P  trace.stuttering.closed"
  shows "P  raw.spec.closed"
by (metis assms raw.spec.cl_def downwards.closed_conv raw.spec.closed trace.stuttering.closed_conv)

lemma empty[intro]:
  shows "{}  raw.spec.closed"
using raw.spec.cl.empty by blast

lemma downwards_closed:
  assumes "P  raw.spec.closed"
  shows "P  downwards.closed"
by (metis assms downwards.closed raw.spec.cl_def raw.spec.closed_conv)

lemma stuttering_closed:
  assumes "P  raw.spec.closed"
  shows "P  trace.stuttering.closed"
using assms raw.spec.cl_def raw.spec.closed_conv by fast

lemma downwards_imp:
  assumes "P  raw.spec.closed"
  assumes "Q  raw.spec.closed"
  shows "downwards.imp P Q  raw.spec.closed"
by (meson assms downwards.closed_imp raw.spec.closed.I raw.spec.closed.stuttering_closed
          trace.stuttering.cl.closed.downwards_imp)

lemma heyting_downwards_imp:
  assumes "P  raw.spec.closed"
  shows "P  downwards.imp Q R  P  Q  R"
by (simp add: assms downwards.heyting_imp raw.spec.closed.downwards_closed)

lemma takeE:
  assumes "σ  P"
  assumes "P  raw.spec.closed"
  shows "trace.take i σ  P"
by (meson assms downwards.closed_in raw.spec.closed.downwards_closed trace.less_eq_take)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.parent_path

typedef ('a, 's, 'v) spec = "raw.spec.closed :: ('a, 's, 'v) trace.t set set"
morphisms unMkS MkS
by blast

setup_lifting type_definition_spec

instantiation spec :: (type, type, type) complete_distrib_lattice
begin

lift_definition bot_spec :: "('a, 's, 'v) spec" is empty ..
lift_definition top_spec :: "('a, 's, 'v) spec" is UNIV ..
lift_definition sup_spec :: "('a, 's, 'v) spec  ('a, 's, 'v) spec  ('a, 's, 'v) spec" is sup ..
lift_definition inf_spec :: "('a, 's, 'v) spec  ('a, 's, 'v) spec  ('a, 's, 'v) spec" is inf ..
lift_definition less_eq_spec :: "('a, 's, 'v) spec  ('a, 's, 'v) spec  bool" is less_eq .
lift_definition less_spec :: "('a, 's, 'v) spec  ('a, 's, 'v) spec  bool" is less .
lift_definition Inf_spec :: "('a, 's, 'v) spec set  ('a, 's, 'v) spec" is Inf ..
lift_definition Sup_spec :: "('a, 's, 'v) spec set  ('a, 's, 'v) spec" is "λX. Sup X  raw.spec.cl {}" ..

instance
by (standard; transfer; auto simp: raw.spec.closed_strict_complete_distrib_lattice_axiomI[OF raw.spec.cl.empty])

end

declare
  SUPE[where 'a="('a, 's, 'v) spec", intro!]
  SupE[where 'a="('a, 's, 'v) spec", intro!]
  Sup_le_iff[where 'a="('a, 's, 'v) spec", simp]
  SupI[where 'a="('a, 's, 'v) spec", intro]
  SUPI[where 'a="('a, 's, 'v) spec", intro]
  rev_SUPI[where 'a="('a, 's, 'v) spec", intro?]
  INFE[where 'a="('a, 's, 'v) spec", intro]

text‹

Observations about this type:
  it is not a BNF (datatype) as it uses the powerset
  it fails to be T0 or sober due to the lack of limit points (completeness) in typ('a, 's, 'v) trace.t
   also stuttering closure precludes T0
  the classcomplete_distrib_lattice instance shows that arbitrary/infinitary Sup›s and Inf›s distribute
   in other words: safety properties are closed under arbitrary intersections and unions
   in other words: Alexandrov
  conclude: the lack of limit points makes this model easier to work in and adds expressivity
   see \S\ref{sec:safety_closure} for further discussion

›

setup Sign.mandatory_path "spec"

lemmas antisym = antisym[where 'a="('a, 's, 'v) spec"]
lemmas eq_iff = order.eq_iff[where 'a="('a, 's, 'v) spec"]

setup Sign.parent_path


subsection‹ Irreducible elements\label{sec:safety_logic-singleton} ›

text‹

The irreducible elements of typ('a, 's, 'v) trace.t are the closures of singletons.

›

setup Sign.mandatory_path "raw"

definition singleton :: "('a, 's, 'v) trace.t  ('a, 's, 'v) trace.t set" where
  "singleton σ = raw.spec.cl {σ}"

lemma singleton_le_conv:
  shows "raw.singleton σ1  raw.singleton σ2  σ1   σ2" (is "?lhs  ?rhs")
proof(rule iffI)
  assume ?lhs
  then have "σ  downwards.cl { σ2}" if "σ   σ1" for σ
    using that trace.natural.mono
    by (force simp: raw.singleton_def raw.spec.cl_def
             intro: downwards.clI[where y="σ1"]
             elim!: downwards.clE trace.stuttering.clE
             dest!: subsetD[where c=σ]
              dest: trace.natural.less_eq_natural)
  then show ?rhs
    by (fastforce simp flip: downwards.order_embedding[where x="σ1"]
                       elim: downwards.clE trace.stuttering.clE)
next
  show "?rhs  ?lhs"
    by (clarsimp simp: raw.singleton_def raw.spec.cl_def)
       (metis downwards.order_embedding trace.natural.idempotent trace.stuttering.cl.downwards.cl
              trace.stuttering.cl_mono trace.stuttering.equiv_cl_singleton)
qed

setup Sign.parent_path

setup Sign.mandatory_path "spec"

lift_definition singleton :: "('a, 's, 'v) trace.t  ('a, 's, 'v) spec" ("_") is raw.singleton
by (simp add: raw.singleton_def)

abbreviation singleton_trace_syn :: "'s  ('a × 's) list  'v option  ('a, 's, 'v) spec" ("_, _, _") where
  "s, xs, v  trace.T s xs v"

setup Sign.mandatory_path "singleton"

lemma Sup_prime:
  shows "Sup_prime σ"
by (clarsimp simp: Sup_prime_on_def)
   (transfer; auto simp: raw.singleton_def elim!: Sup_prime_onE[OF raw.spec.Sup_prime_on_singleton])

lemma nchotomy:
  shows "Xraw.spec.closed. x = (spec.singleton ` X)"
by transfer
   (use raw.spec.closed_conv in auto simp: raw.singleton_def
                                       simp flip: raw.spec.distributive[simplified])

lemmas exhaust = bexE[OF spec.singleton.nchotomy]

lemma collapse[simp]:
  shows "(spec.singleton ` {σ. σ  P}) = P"
by (rule spec.singleton.exhaust[of P]; blast intro: antisym)

lemmas not_bot = Sup_prime_not_bot[OF spec.singleton.Sup_prime] ―‹ Non-triviality ›

setup Sign.parent_path

lemma singleton_le_ext_conv:
  shows "P  Q  (σ. σ  P  σ  Q)" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?rhs  ?lhs"
    by (rule spec.singleton.exhaust[where x=P]; rule spec.singleton.exhaust[where x=Q]; blast)
qed fastforce

lemmas singleton_le_conv = raw.singleton_le_conv[transferred]
lemmas singleton_le_extI = iffD2[OF spec.singleton_le_ext_conv, rule_format]

lemma singleton_eq_conv[simp]:
  shows "σ = σ'  σ S σ'"
by (auto simp: spec.eq_iff spec.singleton_le_conv)

lemma singleton_cong:
  assumes "σ S σ'"
  shows "σ = σ'"
using assms by simp

setup Sign.mandatory_path "singleton"

named_theorems le_conv ‹ simplification rules for ‹⦉σ⦊ ≤ const …› ›

lemmas antisym = antisym[OF spec.singleton_le_extI spec.singleton_le_extI]

lemmas top = spec.singleton.collapse[of , simplified, symmetric]

lemma monotone:
  shows "mono spec.singleton"
by (simp add: monoI trace.natural.mono spec.singleton_le_conv)

lemmas strengthen[strg] = st_monotone[OF spec.singleton.monotone]
lemmas mono = monoD[OF spec.singleton.monotone]
lemmas mono2mono[cont_intro, partial_function_mono]
  = monotone2monotone[OF spec.singleton.monotone, simplified]

lemma simps[simp]:
  shows "σ = σ"
    and "s, xs, v  s, trace.natural' s xs, v"
    and "snd ` set xs  {s}  s, xs @ ys, v = s, ys, v"
    and "snd ` set ys  {trace.final' s xs}  s, xs @ ys, v = s, xs, v"
    and "snd ` set xs  {snd x}  s, x # xs @ ys, v = s, x # ys, v"
    and "s, (a, s) # xs, v = s, xs, v"
by (simp_all add: antisym spec.singleton_le_conv trace.stuttering.equiv.simps trace.natural.simps)

lemma Cons: ―‹ self-applies, not usable by simp›
  assumes "snd ` set as  {s'}"
  shows "s, (a, s') # as, v = s, [(a, s')], v"
by (simp add: assms spec.singleton.simps(4)[where xs="[(a, s')]" and ys="as", simplified])

lemmas Sup_irreducible = iffD1[OF heyting.Sup_prime_Sup_irreducible_iff spec.singleton.Sup_prime]
lemmas sup_irreducible = Sup_irreducible_on_imp_sup_irreducible_on[OF spec.singleton.Sup_irreducible, simplified]
lemmas Sup_leE[elim] = Sup_prime_onE[OF spec.singleton.Sup_prime, simplified]
lemmas sup_le_conv[simp] = sup_irreducible_le_conv[OF spec.singleton.sup_irreducible]
lemmas Sup_le_conv[simp] = Sup_prime_on_conv[OF spec.singleton.Sup_prime, simplified]
lemmas compact_point = Sup_prime_is_compact[OF spec.singleton.Sup_prime]
lemmas compact[cont_intro] = compact_points_are_ccpo_compact[OF spec.singleton.compact_point]

lemma Inf:
  shows "(spec.singleton ` X) = (spec.singleton ` {σ. σ1X. σ  σ1})"
by (fastforce simp: le_INF_iff spec.singleton_le_conv
              dest: spec.singleton.mono
             intro: spec.singleton.antisym)

lemmas inf = spec.singleton.Inf[where X="{σ1, σ2}", simplified] for σ1 σ2

lemma less_eq_Some[simp]:
  shows "s, xs, Some v  σ
      trace.term σ = Some v  trace.init σ = s  trace.natural' s (trace.rest σ) = trace.natural' s xs"
by (auto simp: spec.singleton_le_conv trace.natural_def)

lemma less_eq_None:
  shows [iff]: "s, xs, None  s, xs, v'"
by (auto simp: spec.singleton_le_conv trace.natural_def trace.less_eq_None)

lemma map_cong:
  assumes "a. a  trace.aset (σ')  af a = af' a"
  assumes "x. x  trace.sset (σ')  sf x = sf' x"
  assumes "v. v  trace.vset (σ')  vf v = vf' v"
  assumes "σ = σ'"
  shows "trace.map af sf vf σ = trace.map af' sf' vf' σ'"
proof -
  from assms have "trace.map af sf vf ( σ) S trace.map af' sf' vf' ( σ')"
    by (simp del: trace.sset.natural trace.vset.natural cong: trace.t.map_cong)
  then show ?thesis
    by (simp add: trace.natural.map_natural)
qed

lemma map_le:
  assumes "σ  σ'"
  shows "trace.map af sf vf σ  trace.map af sf vf σ'"
using assms by (simp add: spec.singleton_le_conv trace.natural.map_le)

lemma takeI:
  assumes "σ  P"
  shows "trace.take i σ  P"
by (meson assms order.trans spec.singleton.mono trace.less_eq_take)

setup Sign.parent_path

lemmas assms_cong = order.assms_cong[where 'a="('a, 's, 'v) spec"]
lemmas concl_cong = order.concl_cong[where 'a="('a, 's, 'v) spec"]

declare spec.singleton.transfer[transfer_rule del]

setup Sign.parent_path


subsection‹ Maps\label{sec:safety_logic-maps} ›

text‹

Lift consttrace.map to the typ('a, 's, 'v) spec lattice via image and inverse image.

Note that the image may yield a set that is not stuttering closed
(i.e., we need to close the obvious model-level definition of
spec.map› under stuttering) as arbitrary
sf› may introduce stuttering not present in
P›. In contrast the inverse image preserves stuttering. These
issues are elided here through the use of
constspec.singleton.

›

setup Sign.mandatory_path "spec"

definition map :: "('a  'b)  ('s  't)  ('v  'w)  ('a, 's, 'v) spec  ('b, 't, 'w) spec" where
  "map af sf vf P = (spec.singleton ` trace.map af sf vf ` {σ. σ  P})"

definition invmap :: "('a  'b)  ('s  't)  ('v  'w)  ('b, 't, 'w) spec  ('a, 's, 'v) spec" where
  "invmap af sf vf P = (spec.singleton ` trace.map af sf vf -` {σ. σ  P})"

abbreviation amap ::"('a  'b)  ('a, 's, 'v) spec  ('b, 's, 'v) spec" where
  "amap af  spec.map af id id"
abbreviation ainvmap ::"('a  'b)  ('b, 's, 'v) spec  ('a, 's, 'v) spec" where
  "ainvmap af  spec.invmap af id id"
abbreviation smap ::"('s  't)  ('a, 's, 'v) spec  ('a, 't, 'v) spec" where
  "smap sf  spec.map id sf id"
abbreviation sinvmap ::"('s  't)  ('a, 't, 'v) spec  ('a, 's, 'v) spec" where
  "sinvmap sf  spec.invmap id sf id"
abbreviation vmap ::"('v  'w)  ('a, 's, 'v) spec  ('a, 's, 'w) spec" where ―‹ aka liftM›
  "vmap vf  spec.map id id vf"
abbreviation vinvmap ::"('v  'w)  ('a, 's, 'w) spec  ('a, 's, 'v) spec" where
  "vinvmap vf  spec.invmap id id vf"

interpretation map_invmap: galois.complete_lattice_distributive_class
  "spec.map af sf vf"
  "spec.invmap af sf vf" for af sf vf
proof standard
  show "spec.map af sf vf P  Q  P  spec.invmap af sf vf Q" (is "?lhs  ?rhs") for P Q
  proof(rule iffI)
    show "?lhs  ?rhs"
      by (fastforce simp: spec.map_def spec.invmap_def intro: spec.singleton_le_extI)
    show "?rhs  ?lhs"
      by (fastforce simp: spec.map_def spec.invmap_def
                    dest: order.trans[of _ P] spec.singleton.map_le[where af=af and sf=sf and vf=vf])
  qed
  show "spec.invmap af sf vf (X)  (spec.invmap af sf vf ` X)" for X
    by (fastforce simp: spec.invmap_def)
qed

setup Sign.mandatory_path "singleton"

lemma map_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.map af sf vf P  (σ'. σ'  P  σ  trace.map af sf vf σ')"
by (simp add: spec.map_def)

lemma invmap_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.invmap af sf vf P  trace.map af sf vf σ  P"
by (simp add: spec.invmap_def) (meson order.refl order.trans spec.singleton.map_le)

setup Sign.parent_path

setup Sign.mandatory_path "map"

lemmas bot = spec.map_invmap.lower_bot

lemmas monotone = spec.map_invmap.monotone_lower
lemmas mono = monotoneD[OF spec.map.monotone]

lemmas Sup = spec.map_invmap.lower_Sup
lemmas sup = spec.map_invmap.lower_sup

lemmas Inf_le = spec.map_invmap.lower_Inf_le ―‹ Converse does not hold ›
lemmas inf_le = spec.map_invmap.lower_inf_le ―‹ Converse does not hold ›

lemmas invmap_le = spec.map_invmap.lower_upper_contractive

lemma singleton:
  shows "spec.map af sf vf σ = trace.map af sf vf σ"
by (auto simp: spec.map_def spec.eq_iff spec.singleton_le_conv intro: trace.natural.map_le)

lemma top:
  assumes "surj af"
  assumes "surj sf"
  assumes "surj vf"
  shows "spec.map af sf vf  = "
by (rule antisym)
   (auto simp: assms spec.singleton.top spec.map.Sup spec.map.singleton surj_f_inv_f
        intro: exI[where x="trace.map (inv af) (inv sf) (inv vf) σ" for σ])

lemma id:
  shows "spec.map id id id P = P"
    and "spec.map (λx. x) (λx. x) (λx. x) P = P"
by (simp_all add: spec.map_def flip: id_def)

lemma comp:
  shows "spec.map af sf vf  spec.map ag sg vg = spec.map (af  ag) (sf  sg) (vf  vg)" (is "?lhs = ?rhs")
    and "spec.map af sf vf (spec.map ag sg vg P) = spec.map (λa. af (ag a)) (λs. sf (sg s)) (λv. vf (vg v)) P" (is ?thesis1)
proof -
  have "?lhs P = ?rhs P" for P
    by (rule spec.singleton.exhaust[where x=P])
       (simp add: spec.map.Sup spec.map.singleton image_image comp_def)
  then show "?lhs = ?rhs" and ?thesis1 by (simp_all add: comp_def)
qed

lemmas map = spec.map.comp

lemma inf_distr:
  shows "spec.map af sf vf P  Q = spec.map af sf vf (P  spec.invmap af sf vf Q)" (is "?lhs = ?rhs")
    and "Q  spec.map af sf vf P = spec.map af sf vf (spec.invmap af sf vf Q  P)" (is ?thesis1)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym)
    obtain X where Q: "Q =  (spec.singleton ` X)" using spec.singleton.nchotomy[of Q] by blast
    then have *: "trace.take j ( σQ)  ?rhs"
           if "σP  P"
          and "σQ  X"
          and "  (trace.take i (trace.map af sf vf σP)) = trace.take j ( σQ)"
          for σP σQ i j
      using that
      by (auto simp: spec.singleton.le_conv spec.singleton_le_conv heyting.inf_SUP_distrib
                     spec.map_def spec.singleton.inf trace.less_eq_take_def
                     trace.take.map spec.singleton.takeI trace.take.take trace.natural.take_natural
             intro!: exI[where x="trace.take i σP"] exI[where x=j])
    with Q show "?lhs  ?rhs"
      by (subst spec.map_def)
         (fastforce simp: heyting.inf_SUP_distrib spec.singleton.inf trace.less_eq_take_def
                    elim: trace.take.naturalE(2))
    show "?rhs  ?lhs"
      by (simp add: le_infI1 spec.map_invmap.galois spec.map_invmap.upper_lower_expansive)
  qed
  then show ?thesis1
    by (simp add: inf.commute)
qed

setup Sign.parent_path

setup Sign.mandatory_path "smap"

lemma comp:
  shows "spec.smap sf  spec.smap sg = spec.smap (sf  sg)"
    and "spec.smap sf (spec.smap sg P) = spec.smap (λs. sf (sg s)) P"
by (simp_all add: comp_def spec.map.comp id_def)

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemmas bot = spec.map_invmap.upper_bot
lemmas top = spec.map_invmap.upper_top

lemmas monotone = spec.map_invmap.monotone_upper
lemmas mono = monotoneD[OF spec.invmap.monotone]

lemmas Sup = spec.map_invmap.upper_Sup
lemmas sup = spec.map_invmap.upper_sup

lemmas Inf = spec.map_invmap.upper_Inf
lemmas inf = spec.map_invmap.upper_inf

lemma singleton:
  shows "spec.invmap af sf vf σ = (spec.singleton ` {σ'. trace.map af sf vf σ'  σ})"
by (simp add: spec.invmap_def)

lemma id:
  shows "spec.invmap id id id P = P"
    and "spec.invmap (λx. x) (λx. x) (λx. x) P = P"
unfolding id_def[symmetric] by (metis spec.map.id(1) spec.map_invmap.lower_upper_lower(2))+

lemma comp:
  shows "spec.invmap af sf vf (spec.invmap ag sg vg P) = spec.invmap (λx. ag (af x)) (λs. sg (sf s)) (λv. vg (vf v)) P"  (is "?lhs P = ?rhs P")
    and "spec.invmap af sf vf  spec.invmap ag sg vg = spec.invmap (ag  af) (sg  sf) (vg  vf)" (is ?thesis1)
proof -
  show "?lhs P = ?rhs P" for P
    by (auto intro: spec.singleton.antisym spec.singleton_le_extI simp: spec.singleton.le_conv)
  then show ?thesis1
    by (simp add: fun_eq_iff comp_def)
qed

lemmas invmap = spec.invmap.comp

(* converse seems unlikely *)
lemma invmap_inf_distr_le:
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
  shows "spec.invmap af sf vf P  Q  spec.invmap af sf vf (P  spec.map af sf vf Q)"
    and "Q  spec.invmap af sf vf P  spec.invmap af sf vf (spec.map af sf vf Q  P)"
by (metis order.refl inf_mono spec.map_invmap.upper_inf spec.map_invmap.upper_lower_expansive)+

setup Sign.parent_path

setup Sign.mandatory_path "amap"

lemma invmap_le: ―‹ af = id› in spec.invmap›
  shows "spec.amap af (spec.invmap id sf vf P)  spec.invmap id sf vf (spec.amap af P)"
proof -
  have "spec.invmap id sf vf P  spec.invmap af sf vf (spec.amap af P)" (is "?lhs  ?rhs")
  proof(rule spec.singleton_le_extI)
    show "σ  ?rhs" if "σ  ?lhs" for σ
      using that by (simp add: spec.singleton.le_conv exI[where x="trace.map id sf vf σ"] flip: id_def)
  qed
  then show ?thesis
    by (simp add: spec.map_invmap.galois spec.invmap.comp flip: id_def)
qed

lemma surj_invmap: ―‹ af = id› in spec.invmap›
  fixes P :: "('a, 't, 'w) spec"
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
  assumes "surj af"
  shows "spec.amap af (spec.invmap id sf vf P) = spec.invmap id sf vf (spec.amap af P)" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.amap.invmap_le spec.singleton_le_extI])
  have 1: "σ3. σ2 = trace.map af id id σ3  σ1 S σ3"
    if "trace.map af id id σ1 S σ2"
   for σ1 :: "('a, 't, 'w) trace.t" and σ2
  proof -
    have **: "ys'. ys = map (map_prod af id) ys'  trace.natural' s xs = trace.natural' s ys'"
      if "trace.natural' s (map (map_prod af id) xs) = trace.natural' s ys"
     for s :: "'t" and xs ys
      using that
    proof(induct ys arbitrary: s xs)
      case Nil then show ?case
        by (fastforce simp: trace.natural'.eq_Nil_conv)
    next
      case (Cons y ys s xs) show ?case
      proof(cases "snd y = s")
        case True with Cons.prems show ?thesis
          by (fastforce dest: Cons.hyps
                        simp: iffD1[OF surj_iff surj af]
                   simp flip: id_def
                       intro: exI[where x="map_prod (inv af) id y # ys'" for ys'])
      next
        case False with Cons.prems show ?thesis
          by (force dest!: Cons.hyps
                     simp: trace.natural'.eq_Cons_conv trace.natural'.idle_prefix
                           map_eq_append_conv snd_image_map_prod
                simp flip: id_def
                    intro: exI[where x="(a, s) # xs" for a s xs])
      qed
    qed
    from that show ?thesis
      by (cases σ2) (clarsimp simp: ** trace.natural_def trace.split_Ex)
  qed
  have 2: "zs. xs = map (map_prod af id) zs  map (map_prod id sf) zs = ys"
    if xs_ys: "map (map_prod id sf) xs = map (map_prod af id) ys"
   for xs ys
  proof -
    have "zs. xs' = map (map_prod af id) zs  map (map_prod id sf) zs = ys'"
      if "length xs' = length ys'"
     and "prefix xs' xs"
     and "prefix ys' ys"
     and "map (map_prod id sf) xs = map (map_prod af id) ys" for xs' ys'
      using that
    proof(induct xs' ys' rule: rev_induct2)
      case (snoc x xs y ys) then show ?case
        by (cases x; cases y)
           (force simp: prefix_def simp flip: id_def intro: exI[where x="zs @ [(fst y, snd x)]" for zs])
    qed simp
    from this[OF map_eq_imp_length_eq[OF xs_ys] prefix_order.refl prefix_order.refl xs_ys]
    show ?thesis .
  qed
  fix σ
  assume "σ  ?rhs"
  then obtain σP where σP: "σP  P" "trace.map id sf vf σ  trace.map af id id σP"
    by (clarsimp simp: spec.singleton.le_conv)
  then obtain i σP' where σP': "trace.map id sf vf σ = trace.map af id id σP'" "trace.take i σP S σP'"
    by (fastforce simp: spec.singleton_le_conv trace.less_eq_take_def trace.take.map
                  dest: 1[OF sym]
                 elim!: trace.take.naturalE)
  then obtain zs where zs: "trace.rest σ = map (map_prod af id) zs" "map (map_prod id sf) zs = trace.rest σP'"
    by (cases σ, cases σP') (clarsimp dest!: 2)
  from σP  P σP'(1) σP'(2)[symmetric] zs show "σ  ?lhs"
    by (cases σ, cases σP')
       (clarsimp intro!: exI[where x="trace.T (trace.init σ) zs (trace.term σ)"]
                  elim!: order.trans[rotated]
                   simp: spec.singleton.le_conv spec.singleton_le_conv trace.natural.mono)
qed

setup Sign.parent_path

setup Sign.parent_path


subsection‹ The idle process\label{sec:safety_logic-idle} ›

text‹

As observed by citet"AbadiPlotkin:1991", many laws
require the processes involved to accept all initial states (see, for
instance, \S\ref{sec:safety_logic-monad_laws}). We call the minimal
such process spec.idle›. It is also the lower bound on
specification by transition relation (\S\ref{sec:safety_logic-trans_rel}).

›

setup Sign.mandatory_path "spec"

definition idle :: "('a, 's, 'v) spec" where
  "idle = (s. s, [], None)"

named_theorems idle_le ‹ rules for ‹spec.idle ≤ const …› ›

setup Sign.mandatory_path "singleton"

lemma idle_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.idle  trace.steps σ = {}  trace.term σ = None"
by (auto simp: spec.idle_def spec.singleton_le_conv trace.natural.simps trace.natural'.eq_Nil_conv
               trace.less_eq_None)

setup Sign.parent_path

setup Sign.mandatory_path "idle"

lemma minimal_le:
  shows "s, [], None  spec.idle"
by (simp add: spec.singleton.idle_le_conv)

lemma map_le[spec.idle_le]:
  assumes "spec.idle  P"
  assumes "surj sf"
  shows "spec.idle  spec.map af sf vf P"
by (strengthen ord_to_strengthen(2)[OF assms(1)])
   (use surj sf in auto simp: spec.idle_def spec.map.Sup spec.map.singleton image_image
                                 spec.singleton_le_conv trace.natural.simps trace.less_eq_None)

lemma invmap_le:
  assumes "spec.idle  P"
  shows "spec.idle  spec.invmap af sf vf P"
by (strengthen ord_to_strengthen(2)[OF assms])
   (auto simp: spec.idle_def spec.map.Sup spec.map.singleton simp flip: spec.map_invmap.galois)

setup Sign.parent_path

setup Sign.mandatory_path "map_invmap"

lemma cl_alt_def:
  shows "spec.map_invmap.cl _ _ _ af sf vf P
       = {σ |σ σ'. σ'  P  trace.map af sf vf σ  trace.map af sf vf σ'}" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (rule spec.singleton.exhaust[of P])
       (fastforce simp: spec.map_invmap.cl_def
                        spec.map.Sup spec.map.singleton spec.invmap.Sup spec.invmap.singleton
                 intro: spec.singleton.mono)
  show "?rhs  ?lhs"
    by (clarsimp simp: spec.map_invmap.cl_def simp flip: spec.map.singleton)
       (simp add: order.trans[OF _ spec.map.mono] flip: spec.map_invmap.galois)
qed

lemma cl_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.map_invmap.cl _ _ _ af sf vf P  trace.map af sf vf σ  spec.map af sf vf P"
by (simp add: spec.map_invmap.cl_def spec.singleton.invmap_le_conv)

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Actions\label{sec:safety_logic-actions} ›

text‹

Our primitive actions are arbitrary relations on the state, labelled by the agent performing the
state transition and a value to return.

For refinement purposes we need idle ≤ action a F›; see \S\ref{sec:refinement-action}.

›

setup Sign.mandatory_path "spec"

definition action :: "('v × 'a × 's × 's) set  ('a, 's, 'v) spec" where
  "action F = ((v, a, s, s')F. s, [(a, s')], Some v)  spec.idle"

definition guard :: "('s  bool)  ('a, 's, unit) spec" where
  "guard g = spec.action ({()} × UNIV × Diag g)"

definition return :: "'v  ('a, 's, 'v) spec" where
  "return v = spec.action ({v} × UNIV × Id)"

abbreviation (input) read :: "('s  'v option)  ('a, 's, 'v) spec" where
  "read f  spec.action {(v, a, s, s) |a s v. f s = Some v}"

abbreviation (input) "write" :: "'a  ('s  's)  ('a, 's, unit) spec" where
  "write a f  spec.action {((), a, s, f s) |s. True}"

lemma action_le[case_names idle step]:
  assumes "spec.idle  P"
  assumes "v a s s'. (v, a, s, s')  F  s, [(a, s')], Some v  P"
  shows "spec.action F  P"
by (simp add: assms spec.action_def split_def)

setup Sign.mandatory_path "idle"

lemma action_le[spec.idle_le]:
  shows "spec.idle  spec.action F"
by (simp add: spec.action_def)

lemma guard_le[spec.idle_le]:
  shows "spec.idle  spec.guard g"
by (simp add: spec.guard_def spec.idle_le)

lemma return_le[spec.idle_le]:
  shows "spec.idle  spec.return v"
by (simp add: spec.return_def spec.idle_le)

setup Sign.parent_path

setup Sign.mandatory_path "map"

lemma action_le:
  fixes F :: "('v × 'a × 's × 's) set"
  shows "spec.map af sf vf (spec.action F)  spec.action (map_prod vf (map_prod af (map_prod sf sf)) ` F)"
by (force simp: spec.action_def spec.idle_def spec.map.Sup spec.map.sup spec.map.singleton SUP_le_iff)

lemma action:
  fixes F :: "('v × 'a × 's × 's) set"
  shows "spec.map af sf vf (spec.action F)  spec.idle
       = spec.action (map_prod vf (map_prod af (map_prod sf sf)) ` F)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: spec.idle_le spec.map.action_le)
  show "?rhs  ?lhs"
    by (force simp: spec.action_def spec.idle_def spec.map.sup spec.map.singleton spec.map.Sup)
qed

lemma surj_sf_action:
  assumes "surj sf"
  shows "spec.map af sf vf (spec.action F) = spec.action (map_prod vf (map_prod af (map_prod sf sf)) ` F)"
by (simp add: assms sup.absorb1 spec.idle_le flip: spec.map.action)

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma empty:
  shows "spec.action {} = spec.idle"
by (simp add: spec.action_def)

lemma idleI:
  assumes "snd ` set xs  {s}"
  shows "s, xs, None  spec.action F"
using assms by (simp add: spec.action_def spec.singleton.le_conv)

lemma stepI:
  assumes "(v, a, s, s')  F"
  assumes "v''. w = Some v''  v'' = v"
  shows "s, [(a, s')], w  spec.action F"
using assms by (cases w; force simp: spec.action_def spec.singleton.mono trace.less_eq_None)

lemma stutterI:
  assumes "(v, a, s, s)  F"
  shows "s, [], Some v  spec.action F"
by (fastforce simp: spec.singleton_le_conv trace.natural.simps
             intro: order.trans[OF _ spec.action.stepI[OF assms]])

lemma stutter_stepI:
  assumes "(v, a, s, s)  F"
  shows "s, [(b, s)], Some v  spec.action F"
by (fastforce simp: spec.singleton_le_conv trace.natural.simps
             intro: order.trans[OF _ spec.action.stepI[OF assms]])

lemma stutter_stepsI:
  assumes "(v, a, s, s)  F"
  assumes "snd ` set xs  {s}"
  shows "s, xs, Some v  spec.action F"
by (simp add: assms trace.natural'.eq_Nil_conv order.trans[OF _ spec.action.stutterI[OF assms(1)]])

lemma monotone:
  shows "mono spec.action"
by (force simp: spec.action_def intro: monoI)

lemmas strengthen[strg] = st_monotone[OF spec.action.monotone]
lemmas mono = monotoneD[OF spec.action.monotone]
lemmas mono2mono[cont_intro, partial_function_mono]
  = monotone2monotone[OF spec.action.monotone, simplified]

lemma Sup:
  shows "spec.action (X) = (FX. spec.action F)  spec.idle"
by (force simp: spec.eq_iff spec.action_def)

lemma
  shows SUP: "spec.action (xX. F x) = (xX. spec.action (F x))  spec.idle"
    and SUP_not_empty: "X  {}  spec.action (xX. F x) = (xX. spec.action (F x))"
by (auto simp: spec.action.Sup image_image sup.absorb1 SUPI spec.idle_le simp flip: ex_in_conv)

lemma sup:
  shows "spec.action (F  G) = spec.action F  spec.action G"
using spec.action.Sup[where X="{F, G}"] by (simp add: sup_absorb1 le_supI1 spec.idle_le)

(* the converse is complex *)
lemma Inf_le:
  shows "spec.action (Fs)  (spec.action ` Fs)"
by (simp add: spec.action_def ac_simps SUP_le_iff SUP_upper le_INF_iff le_supI2)

lemma inf_le:
  shows "spec.action (F  G)  spec.action F  spec.action G"
using spec.action.Inf_le[where Fs="{F, G}"] by simp

lemma stutter_agents_le:
  assumes "A  {}; r  {}  B  {}"
  assumes "r  Id"
  shows "spec.action ({v} × A × r)  spec.action ({v} × B × r)"
using assms
by (subst spec.action_def) (fastforce simp: spec.idle_le intro!: spec.action.stutter_stepI)

lemma read_agents:
  assumes "A  {}"
  assumes "B  {}"
  assumes "r  Id"
  shows "spec.action ({v} × A × r) = spec.action ({v} × B × r)"
by (rule antisym[OF spec.action.stutter_agents_le spec.action.stutter_agents_le]; rule assms)

lemma invmap_le: ―‹ A typical refinement ›
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
  shows "spec.action (map_prod vf (map_prod af (map_prod sf sf)) -` F)  spec.invmap af sf vf (spec.action F)"
by (meson order.trans image_vimage_subset spec.action.mono spec.map.action_le spec.map_invmap.galois)

setup Sign.parent_path

setup Sign.mandatory_path "singleton"

lemma action_le_conv:
  shows "σ  spec.action F
     (trace.steps σ = {}  case_option True (λv. a. (v, a, trace.init σ, trace.init σ)  F) (trace.term σ))
      (xF. trace.steps σ = {snd x}  case_option True ((=) (fst x)) (trace.term σ))" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    unfolding spec.action_def spec.singleton.sup_le_conv
  proof(induct rule: disjE[consumes 1, case_names step idle])
    case step
    then obtain v a s s' where *: "σ  (trace.T s [(a, s')] (Some v))" and F: "(v, a, s, s')  F"
      by (clarsimp simp: spec.singleton_le_conv)
    from * show ?case
    proof(induct rule: trace.less_eqE)
      case prefix with F show ?case
        by (clarsimp simp add: trace.natural'.eq_Nil_conv prefix_Cons split: if_splits)
           (force simp: trace.steps'_alt_def)
    next
      case (maximal v) with F show ?case
        by (clarsimp simp: trace.natural.trace_conv trace.natural'.eq_Nil_conv
                           trace.natural'.eq_Cons_conv trace.steps'.append split: if_splits;
            force)
    qed
  qed (simp add: spec.singleton.le_conv)
  show "?rhs  ?lhs"
    by (cases σ)
       (auto simp: trace.steps'.step_conv
            intro: spec.action.idleI spec.action.stutter_stepsI spec.action.stepI
            elim!: order.trans[OF eq_refl[OF spec.singleton.Cons]]
            split: option.split_asm)
qed

lemma action_Some_leE:
  assumes "σ  spec.action F"
  assumes "trace.term σ = Some v"
  obtains x
    where "x  F"
      and "trace.init σ = fst (snd (snd x))"
      and "trace.final σ = snd (snd (snd x))"
      and "trace.steps σ  {snd x}"
      and "v = fst x"
using assms by (auto simp: spec.singleton.action_le_conv trace.steps'.step_conv trace.steps'.append)

lemma action_not_idle_leE:
 assumes "σ  spec.action F"
 assumes "σ  trace.T (trace.init σ) [] None"
 obtains x
 where "x  F"
   and "trace.init σ = fst (snd (snd x))"
   and "trace.final σ = snd (snd (snd x))"
   and "trace.steps σ  {snd x}"
   and "case_option True ((=) (fst x)) (trace.term σ)"
 using assms
by (cases σ)
   (auto 0 0 simp: spec.singleton.action_le_conv trace.natural.idle option.case_eq_if
                   trace.steps'.step_conv trace.steps'.append)

lemma action_not_idle_le_splitE:
  assumes "σ  spec.action F"
  assumes "σ  trace.T (trace.init σ) [] None"
  obtains (return) v a
          where "(v, a, trace.init σ, trace.init σ)  F"
            and "trace.steps σ = {}"
            and "trace.term σ = Some v"
        | (step) v a ys zs
          where "(v, a, trace.init σ, trace.final σ)  F"
            and "trace.init σ  trace.final σ"
            and "snd ` set ys  {trace.init σ}"
            and "snd ` set zs  {trace.final σ}"
            and "trace.rest σ = ys @ [(a, trace.final σ)] @ zs"
            and "case_option True ((=) v) (trace.term σ)"
using assms
by (cases σ)
   (auto 0 0 simp: spec.singleton.action_le_conv trace.natural.idle option.case_eq_if
                    trace.steps'.step_conv trace.steps'.append
             cong: if_cong)

lemma guard_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.guard g  trace.steps σ = {}  (case_option True g (trace.init σ) (trace.term σ))"
by (fastforce simp: spec.guard_def spec.singleton.action_le_conv trace.steps'.step_conv
             split: option.split)

lemma return_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.return v
      trace.steps σ = {}  (case_option True ((=) v) (trace.term σ))"
by (fastforce simp: spec.return_def spec.singleton.action_le_conv trace.steps'.step_conv
             split: option.split)

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma mono_stronger:
  assumes "v a s s'. (v, a, s, s')  F; s  s'  (v, a, s, s')  F'"
  assumes "v a s. (v, a, s, s)  F  a'. (v, a', s, s)  F'"
  shows "spec.action F  spec.action F'"
proof (induct rule: spec.action_le[OF spec.idle.action_le, case_names step])
  case (step v a s s') then show ?case
    by (cases "s = s'")
       (auto dest: assms intro: spec.action.stutterI spec.action.stepI)
qed

lemma cong:
  assumes "v a s s'. s  s'  (v, a, s, s')  F  (v, a, s, s')  F'"
  assumes "v a s. (v, a, s, s)  F  a'. (v, a', s, s)  F'"
  assumes "v a s. (v, a, s, s)  F'  a'. (v, a', s, s)  F"
  shows "spec.action F = spec.action F'"
using assms by (blast intro!: spec.antisym spec.action.mono_stronger)

lemma le_actionD:
  assumes "spec.action F  spec.action F'"
  shows "(v, a, s, s')  F; s  s'  (v, a, s, s')  F'"
    and "(v, a, s, s)  F  a'. (v, a', s, s)  F'"
proof -
  fix v a s s'
  show "(v, a, s, s')  F'" if "(v, a, s, s')  F" and "s  s'"
    using iffD1[OF spec.singleton_le_ext_conv assms] that
    by (fastforce simp: spec.singleton.action_le_conv
                  dest: spec[where x="trace.T s [(a, s')] (Some v)"])
  show "a'. (v, a', s, s)  F'" if "(v, a, s, s)  F"
    using iffD1[OF spec.singleton_le_ext_conv assms] that
    by (fastforce simp: spec.singleton.action_le_conv
                  dest: spec[where x="trace.T s [] (Some v)"])
qed

lemma eq_action_conv:
  shows "spec.action F = spec.action F'
     (v a s s'. s  s'  (v, a, s, s')  F  (v, a, s, s')  F')
       (v a s. (v, a, s, s)  F  (a'. (v, a', s, s)  F'))
       (v a s. (v, a, s, s)  F'  (a'. (v, a', s, s)  F))"
by (rule iffI, metis order.refl spec.action.le_actionD, blast intro!: spec.action.cong)

setup Sign.parent_path

lemma return_alt_def:
  assumes "A  {}"
  shows "spec.return v = spec.action ({v} × A × Id)"
unfolding spec.return_def using assms by (blast intro: spec.action.cong)

setup Sign.mandatory_path "return"

lemma cong:
  assumes "v a s s'. (v, a, s, s')  F  s' = s"
  assumes "v s. v  fst ` F  a. (v, a, s, s)  F"
  shows "spec.action F = (spec.return ` fst ` F)  spec.idle"
by (simp add: spec.return_def image_image flip: spec.action.SUP)
   (rule spec.action.cong; auto intro: rev_bexI dest: assms(1) intro: assms(2))

lemma action_le:
  assumes "Id  snd ` snd ` F"
  shows "spec.return ()  spec.action F"
unfolding spec.return_def
proof(induct rule: spec.action_le)
  case (step v a s s') with subsetD[OF assms, where c="(s, s)"] show ?case
    by (force intro: spec.action.stutterI)
qed (simp add: spec.idle_le)

setup Sign.parent_path

setup Sign.mandatory_path "guard"

lemma alt_def:
  assumes "A  {}"
  shows "spec.guard g = spec.action ({()} × A × Diag g)"
unfolding spec.guard_def using assms by (fastforce simp: intro: spec.action.cong)

lemma bot:
  shows "spec.guard  = spec.idle"
    and "spec.guard False = spec.idle"
by (simp_all add: spec.guard_def spec.action.empty)

lemma top:
  shows "spec.guard  = spec.return ()"
    and "spec.guard True = spec.return ()"
by (simp_all add: spec.guard_def spec.return_def flip: Id_def)

lemma monotone:
  shows "mono spec.guard"
proof(rule monotoneI)
  show "spec.guard g  spec.guard g'" if "g  g'" for g g' :: "'s pred"
    unfolding spec.guard_def by (strengthen ord_to_strengthen(1)[OF g  g']) simp
qed

lemmas strengthen[strg] = st_monotone[OF spec.guard.monotone]
lemmas mono = monotoneD[OF spec.guard.monotone]
lemmas mono2mono[cont_intro, partial_function_mono] = monotone2monotone[OF spec.guard.monotone, simplified]

lemma Sup:
  shows "spec.guard (X) = (spec.guard ` X)  spec.idle"
by (auto simp: spec.guard_def Diag_Sup
    simp flip: spec.action.Sup[where X="(λx. {()} × UNIV × Diag x) ` X", simplified image_image]
        intro: arg_cong[where f=spec.action])

lemma sup:
  shows "spec.guard (g  h) = spec.guard g  spec.guard h"
by (simp add: spec.guard.Sup[where X="{g, h}" for g h, simplified]
              ac_simps sup_absorb2 le_supI2 spec.idle_le)

lemma return_le:
  shows "spec.guard g  spec.return ()"
by (simp add: spec.guard_def spec.return_def Sigma_mono spec.action.mono)

lemma guard_less: ―‹ Non-triviality ›
  assumes "g < g'"
  shows "spec.guard g < spec.guard g'"
proof(rule le_neq_trans)
  show "spec.guard g  spec.guard g'"
    by (strengthen ord_to_strengthen(1)[OF order_less_imp_le[OF assms]]) simp
  from assms obtain s where "g' s" "¬g s" by (metis leD predicate1I)
  from ¬g s have "¬s, [], Some ()  spec.guard g"
    by (clarsimp simp: spec.guard_def spec.action_def
                       spec.singleton_le_conv spec.singleton.le_conv trace.natural.simps)
  moreover
  from g' s have "s, [], Some ()  spec.guard g'"
    by (simp add: spec.guard_def spec.action.stutterI)
  ultimately show "spec.guard g  spec.guard g'" by metis
qed

lemma cong:
  assumes "v a s s'. (v, a, s, s')  F  s' = s"
  shows "spec.action F = spec.guard (λs. s  fst ` snd ` snd ` F)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (force simp: spec.guard_def intro: spec.action.mono dest: assms)
  show "?rhs  ?lhs"
    unfolding spec.guard_def
    by (rule spec.action_le;
        clarsimp simp: spec.idle_le; blast intro: spec.action.stutterI dest: assms)
qed

lemma action_le:
  assumes "Diag g  snd ` snd ` F"
  shows "spec.guard g  spec.action F"
unfolding spec.guard_def
proof(induct rule: spec.action_le)
  case (step v a s s') with subsetD[OF assms, where c="(s, s)"] show ?case
    by (force intro: spec.action.stutterI)
qed (simp add: spec.idle_le)

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Operations on return values\label{sec:safety_logic-return_values} ›

text‹

For various purposes, including defining a history-respecting
sequential composition (bind, see \S\ref{sec:safety_logic-bind}), we
use a Galois pair of operations that saturate or eradicate return
values.

›

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "term"

definition none :: "('a, 's, 'v) spec  ('a, 's, 'w) spec" where
  "none P = {s, xs, None |s xs v. s, xs, v  P}"

definition all :: "('a, 's, 'v) spec  ('a, 's, 'w) spec" where
  "all P = {s, xs, v |s xs v. s, xs, None  P}"

setup Sign.parent_path

interpretation "term": galois.complete_lattice_distributive_class spec.term.none spec.term.all
proof standard
  show "spec.term.none P  Q  P  spec.term.all Q" (is "?lhs  ?rhs")
   for P :: "('a, 'b, 'c) spec"
   and  Q :: "('a, 'b, 'f) spec"
  proof(rule iffI)
    show "?lhs  ?rhs"
      by (fastforce simp: spec.term.none_def spec.term.all_def trace.split_all
                   intro: spec.singleton_le_extI)
    show "?rhs  ?lhs"
      by (fastforce simp: spec.term.none_def spec.term.all_def
                          spec.singleton_le_conv trace.natural_def trace.less_eq_None
                    elim: trace.less_eqE order.trans[rotated]
                    dest: order.trans[of _ P])
  qed
  show "spec.term.all ( X)   (spec.term.all ` X)" for X :: "('a, 'b, 'f) spec set"
    by (auto 0 5 simp: spec.term.all_def)
qed

setup Sign.mandatory_path "singleton.term"

lemma none_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.term.none P  trace.term σ = None  trace.init σ, trace.rest σ, None  P" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (fastforce simp: spec.term.none_def trace.natural_def spec.singleton_le_conv trace.less_eq_None
                 intro: order.trans[rotated])
  show "?rhs  ?lhs"
    by (cases σ) (fastforce simp: spec.term.none_def)
qed

lemma all_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.term.all P  (w. trace.init σ, trace.rest σ, w  P)" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (cases σ) (fastforce simp: spec.term.none_def simp flip: spec.term.galois)
  show "?rhs  ?lhs"
    by (cases σ) (fastforce simp: spec.term.all_def intro: order.trans[rotated])
qed

setup Sign.parent_path

setup Sign.mandatory_path "term"

setup Sign.mandatory_path "none"

lemma singleton:
  shows "spec.term.none σ = trace.init σ, trace.rest σ, None"
by (force simp: spec.eq_iff spec.term.galois spec.singleton.le_conv spec.singleton.mono trace.less_eq_None)

lemmas bot[simp] = spec.term.lower_bot

lemmas monotone = spec.term.monotone_lower
lemmas mono = monotoneD[OF spec.term.none.monotone]

lemmas Sup = spec.term.lower_Sup
lemmas sup = spec.term.lower_sup

lemmas Inf_le = spec.term.lower_Inf_le

lemma Inf_not_empty:
  assumes "X  {}"
  shows "spec.term.none (X) = (xX. spec.term.none x)"
by (rule antisym[OF spec.term.lower_Inf_le])
   (use assms in auto intro: spec.singleton_le_extI simp: spec.singleton.term.none_le_conv le_Inf_iff)

lemma inf:
  shows "spec.term.none (P  Q) = spec.term.none P  spec.term.none Q"
    and "spec.term.none (Q  P) = spec.term.none Q  spec.term.none P"
using spec.term.none.Inf_not_empty[where X="{P, Q}"] by (simp_all add: ac_simps)

lemma inf_unit:
  fixes P Q :: "(_, _, unit) spec"
  shows "spec.term.none (P  Q) = spec.term.none P  Q" (is "?thesis1 P Q")
    and "spec.term.none (P  Q) = P  spec.term.none Q" (is ?thesis2)
proof -
  show *: "?thesis1 P Q" for P Q
    by (rule spec.singleton.antisym; metis le_inf_iff spec.singleton.term.none_le_conv trace.t.collapse)
  from *[where P=Q and Q=P] show ?thesis2
    by (simp add: ac_simps)
qed

lemma idempotent[simp]:
  shows "spec.term.none (spec.term.none P) = spec.term.none P"
by (rule spec.singleton.exhaust[of P])
   (simp add: spec.term.none.Sup spec.term.none.singleton image_image)

lemma contractive[iff]:
  shows "spec.term.none P  P"
by (rule spec.singleton_le_extI) (simp add: spec.singleton.le_conv trace.split_all)

lemma map_gen:
  fixes vf :: "'v  'w"
  fixes vf' :: "'a  'b" ―‹ arbitrary type ›
  shows "spec.term.none (spec.map af sf vf P) = spec.map af sf vf' (spec.term.none P)" (is "?lhs = ?rhs")
by (fastforce simp: spec.map_def spec.eq_iff image_image trace.split_all trace.split_Ex
                    spec.term.none.Sup spec.term.none.singleton spec.singleton.term.none_le_conv
              elim: order.trans[rotated])

lemmas map = spec.term.none.map_gen[where vf'=id] ―‹ simp›-friendly ›

lemma invmap_gen:
  fixes vf :: "'v  'w"
  fixes vf' :: "'a  'b" ―‹ arbitrary type ›
  shows "spec.term.none (spec.invmap af sf vf P) = spec.invmap af sf vf' (spec.term.none P)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: spec.map_invmap.lower_upper_contractive spec.term.none.mono
            flip: spec.map_invmap.galois spec.term.none.map_gen[where vf=vf])
  show "?rhs  ?lhs"
    by (rule spec.singleton_le_extI)
       (clarsimp simp: spec.singleton.invmap_le_conv spec.singleton.term.none_le_conv)
qed

lemmas invmap = spec.term.none.invmap_gen[where vf'=id] ―‹ simp›-friendly ›

lemma idle:
  shows "spec.term.none spec.idle = spec.idle"
by (simp add: spec.idle_def spec.term.none.Sup spec.term.none.singleton image_image)

lemma return:
  shows "spec.term.none (spec.return v) = spec.idle"
by (auto simp: spec.eq_iff spec.return_def spec.action_def spec.term.none.idle spec.singleton.idle_le_conv
               spec.term.none.sup spec.term.none.Sup spec.term.none.singleton)

lemma guard:
  shows "spec.term.none (spec.guard g) = spec.idle"
by (rule antisym[OF spec.term.none.mono[OF spec.guard.return_le, simplified spec.term.none.return]
                    spec.term.none.mono[OF spec.idle.guard_le, simplified spec.term.none.idle]])

setup Sign.parent_path

lemma none_all_le:
  shows "spec.term.none P  spec.term.all P"
using spec.term.galois by fastforce

lemma none_all[simp]:
  shows "spec.term.none (spec.term.all P) = spec.term.none P"
by (metis spec.eq_iff spec.term.lower_upper_contractive
          spec.term.none.idempotent spec.term.none.mono spec.term.none_all_le)

lemma all_none[simp]:
  shows "spec.term.all (spec.term.none P) = spec.term.all P"
by (metis spec.eq_iff spec.term.galois spec.term.none_all)

setup Sign.mandatory_path "all"

lemmas bot[simp] = spec.term.upper_bot

lemmas top = spec.term.upper_top

lemmas monotone = spec.term.monotone_upper
lemmas mono = monotoneD[OF spec.term.all.monotone]

lemma expansive:
  shows "P  spec.term.all P"
using spec.term.galois by blast

lemmas Sup = spec.term.upper_Sup
lemmas sup = spec.term.upper_sup

lemmas Inf = spec.term.upper_Inf
lemmas inf = spec.term.upper_inf

lemmas singleton = spec.term.all_def[where P="σ"] for σ

lemma monomorphic:
  shows "spec.term.cl _ = spec.term.all"
unfolding spec.term.cl_def by simp

lemma closed_conv:
  assumes "P  spec.term.closed _"
  shows "P = spec.term.all P"
using assms spec.term.closed_conv by (auto simp: spec.term.all.monomorphic)

lemma closed[iff]:
  shows "spec.term.all P  spec.term.closed _"
using spec.term.closed_upper by (auto simp: spec.term.all.monomorphic)

lemma idempotent[simp]:
  shows "spec.term.all (spec.term.all P) = spec.term.all P"
by (metis antisym spec.term.galois spec.term.lower_upper_contractive spec.term.none.idempotent)

lemma map: ―‹ vf = id› on the RHS ›
  fixes vf :: "'v  'w"
  shows "spec.term.all (spec.map af sf vf P) = spec.map af sf id (spec.term.all P)" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI])
  fix σ
  assume "σ  ?lhs"
  then obtain σ' i w
    where "σ'  P"
      and "trace.T (trace.init σ) (trace.rest σ) w S trace.map af sf vf (trace.take i σ')"
    using that by (fastforce elim!: trace.less_eq_takeE trace.take.naturalE
                              simp: trace.take.map spec.singleton.le_conv spec.singleton_le_conv)
  then show "σ  ?rhs"
    by (simp add: spec.singleton.le_conv spec.singleton_le_conv)
       (fastforce intro: exI[where x="trace.T (trace.init σ') (trace.rest (trace.take i σ')) (trace.term σ)"]
                         exI[where x=None]
                   elim: order.trans[rotated]
                   simp: trace.natural_def spec.singleton.mono trace.less_eq_None take_is_prefix)
next
  show "?rhs  ?lhs"
    by (simp add: spec.term.none.map flip: spec.term.galois)
       (simp flip: spec.term.none.map[where vf=vf])
qed

lemma invmap: ―‹ vf = id› on the RHS ›
  fixes vf :: "'v  'w"
  shows "spec.term.all (spec.invmap af sf vf P) = spec.invmap af sf id (spec.term.all P)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: order.trans[OF spec.term.none.contractive spec.map_invmap.lower_upper_contractive]
            flip: spec.map_invmap.galois spec.term.galois spec.term.all.map[where vf=vf])
  show "?rhs  ?lhs"
    by (simp add: spec.term.none.invmap flip: spec.term.galois)
       (simp flip: spec.term.none.invmap_gen[where vf=vf])
qed

lemma vmap_unit_absorb:
  shows "spec.vmap () (spec.term.all P) = spec.term.all P" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (simp add: spec.term.none.map spec.map.id flip: spec.term.galois)
  show "σ  ?lhs" if "σ  ?rhs" for σ
    using that
    by (clarsimp simp: spec.singleton.le_conv
               intro!: exI[where x="trace.map id id undefined σ"])
       (metis (mono_tags) order.refl fun_unit_id trace.t.map_ident)
qed

lemma vmap_unit:
  shows "spec.vmap () (spec.term.all P) = spec.term.all (spec.vmap () P)"
by (simp add: spec.map.id spec.term.all.map spec.term.all.vmap_unit_absorb)

lemma idle:
  shows "spec.term.all spec.idle = (v. spec.return v)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (clarsimp simp: spec.term.all_def spec.singleton.le_conv option.case_eq_if)
  show "?rhs  ?lhs"
    by (simp add: spec.term.none.Sup spec.term.none.return flip: spec.term.galois)
qed

lemma action:
  fixes F :: "('v × 'a × 's × 's) set"
  shows "spec.term.all (spec.action F) = spec.action (UNIV × snd ` F)  (v. spec.return v)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (clarsimp simp: spec.term.all_def spec.singleton.le_conv spec.singleton.action_le_conv
                split: option.split)
       meson
  show "?rhs  ?lhs"
    by (force simp: spec.action_def spec.idle_le spec.term.none.idle spec.term.none.return
                    spec.term.none.Sup spec.term.none.sup spec.term.none.singleton
         simp flip: spec.term.galois)
qed

lemma return:
  shows "spec.term.all (spec.return v) = (v. spec.return v)"
by (auto simp: spec.return_def spec.term.all.action
    simp flip: spec.action.SUP_not_empty spec.action.sup
        intro: arg_cong[where f="spec.action"])

lemma guard:
  shows "spec.term.all (spec.guard g) = (v. spec.return v)"
by (simp add: spec.eq_iff spec.idle.guard_le spec.term.none.Sup spec.term.none.return
              spec.term.all.mono[OF spec.guard.return_le, unfolded spec.term.all.return]
        flip: spec.term.galois)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "idle.term"

lemma none_le_conv[spec.idle_le]:
  shows "spec.idle  spec.term.none P  spec.idle  P"
by (metis spec.term.all.monomorphic spec.term.cl_def spec.term.galois spec.term.none.idle)

lemma all_le_conv[spec.idle_le]:
  shows "spec.idle  spec.term.all P  spec.idle  P"
by (simp add: spec.term.none.idle flip: spec.term.galois)

setup Sign.parent_path

setup Sign.mandatory_path "term.closed"

lemma return_unit:
  shows "spec.return ()  spec.term.closed _"
by (rule spec.term.closed_clI) (simp add: spec.term.all.return spec.term.all.monomorphic)

lemma none_inf:
  fixes P :: "('a, 's, 'v) spec"
  fixes Q :: "('a, 's, 'w) spec"
  assumes "P  spec.term.closed _"
  shows "P  spec.term.none Q = spec.term.none (spec.term.none P  Q)" (is "?lhs = ?rhs")
    and "spec.term.none Q  P = spec.term.none (Q  spec.term.none P)" (is ?thesis1)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym[OF spec.singleton_le_extI])
    show "σ  ?rhs" if "σ  ?lhs" for σ
      using that by (cases σ) (simp add: spec.singleton.le_conv)
    show "?rhs  ?lhs"
      by (auto simp: spec.term.galois intro: le_infI1 le_infI2 spec.term.none_all_le spec.term.all.expansive)
  qed
  then show ?thesis1
    by (simp add: ac_simps)
qed

lemma none_inf_monomorphic:
  fixes P :: "('a, 's, 'v) spec"
  fixes Q :: "('a, 's, 'v) spec"
  assumes "P  spec.term.closed _"
  shows "P  spec.term.none Q = spec.term.none (P  Q)" (is ?thesis1)
    and "spec.term.none Q  P = spec.term.none (Q  P)" (is ?thesis2)
by (simp_all add: spec.term.closed.none_inf[OF assms, simplified] spec.term.none.inf)

lemma singleton_le_extI:
  assumes "Q  spec.term.closed _"
  assumes "s xs. s, xs, None  P  s, xs, None  Q"
  shows "P  Q"
by (subst spec.term.closed_conv[OF assms(1)], rule spec.singleton_le_extI)
   (auto simp: trace.split_all spec.term.none.singleton spec.term.all.monomorphic
    simp flip: spec.term.galois
        intro: assms(2)
         elim: order.trans[rotated])

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Bind\label{sec:safety_logic-bind} ›

text‹

We define monadic bind› in terms of bi-strict continue›. The latter supports left and right residuals
(see, amongst many others, citet"HoareHe:1987" and "HoareHeSanders:1987" and "Pratt:1990"), whereas
bind› encodes the non-retractability of observable actions, i.e., spec.term.none f ≤ f ⤜ g›, which
defeats a general right residual.

It is tempting to write this in a more direct style (using constcase_option) but the set comprehension
syntax is not friendly to strengthen/monotonicity facts.

›

setup Sign.mandatory_path "spec"

definition continue :: "('a, 's, 'v) spec  ('v  ('a, 's, 'w) spec)  ('a, 's, 'w) spec" where
  "continue f g =
    {trace.init σf, trace.rest σf @ trace.rest σg, trace.term σg
        |σf σg v. σf  f  trace.init σg = trace.final σf  trace.term σf = Some v  σg  g v}"

definition bind :: "('a, 's, 'v) spec  ('v  ('a, 's, 'w) spec)  ('a, 's, 'w) spec" where
  "bind f g = spec.term.none f  spec.continue f g"

adhoc_overloading
  Monad_Syntax.bind spec.bind

setup Sign.mandatory_path "singleton"

lemma continue_le_conv:
  shows "σ  spec.continue f g
     (xs ys v w. trace.init σ, xs, Some v  f
                    trace.final' (trace.init σ) xs, ys, w  g v
                    σ  trace.T (trace.init σ) (xs @ ys) w)" (is "?lhs  ?rhs")
proof(rule iffI)
  assume ?lhs
  then obtain s xs ys v w
        where σ: "σ  (trace.T s (xs @ ys) w)"
          and f: "s, xs, Some v  f"
          and g: "trace.final' s xs, ys, w  g v"
    by (clarsimp simp: spec.continue_def trace.split_all spec.singleton_le_conv)
  from σ show ?rhs
  proof(cases rule: trace.less_eqE)
    case prefix
    from prefix(3)[simplified, simplified trace.natural'.append] show ?thesis
    proof(cases rule: prefix_append_not_NilE)
      case incomplete
      then obtain zs where zs: "trace.natural' s xs = trace.natural' (trace.init σ) (trace.rest σ) @ zs"
        by (rule prefixE)
      from f prefix(2) zs
      have "trace.init σ, trace.rest σ @ zs, Some v  f"
        by (clarsimp elim!: order.trans[rotated])
           (metis trace.natural'.append trace.final'.natural' trace.natural'.natural')
      moreover
      from g prefix(2) zs
      have "trace.final' (trace.init σ) (trace.rest σ @ zs), ys, None  g v"
        by (clarsimp elim!: order.trans[rotated])
           (metis spec.singleton.less_eq_None trace.final'.natural' trace.final'.simps(3))
      moreover note trace.term (σ) = None
      ultimately show ?thesis
        by (fastforce simp: trace.less_eq_None)
    next
      case (continue us)
      from continue(1)
      obtain ys' zs'
        where "trace.rest σ = ys' @ zs'"
          and "trace.natural' (trace.init σ) ys' = trace.natural' s xs"
          and "trace.natural' (trace.final' (trace.init σ) (trace.natural' s xs)) zs' = us"
        by (clarsimp simp: trace.natural'.eq_append_conv)
      with f g prefix(1,2) continue(2-) show ?thesis
        by - (rule exI[where x="ys'"];
              force simp: spec.singleton_le_conv trace.less_eq_None trace.natural_def
                    cong: trace.final'.natural'_cong
                   elim!: order.trans[rotated]
                   intro: exI[where x="None"])
    qed
  next
    case (maximal x) with f g show ?thesis
      by (fastforce simp: trace.stuttering.equiv.append_conv
                    cong: trace.final'.natural'_cong
                   elim!: order.trans[rotated])
  qed
next
  show "?rhs  ?lhs"
    using spec.singleton.mono by (auto 10 0 simp: spec.continue_def trace.split_Ex)
qed

setup Sign.parent_path

setup Sign.mandatory_path "continue"

lemma mono:
  assumes "f  f'"
  assumes "v. g v  g' v"
  shows "spec.continue f g  spec.continue f' g'"
unfolding spec.continue_def
apply (strengthen ord_to_strengthen(1)[OF assms(1)])
apply (strengthen ord_to_strengthen(1)[OF assms(2)])
apply (rule order.refl)
done

lemma strengthen[strg]:
  assumes "st_ord F f f'"
  assumes "x. st_ord F (g x) (g' x)"
  shows "st_ord F (spec.continue f g) (spec.continue f' g')"
using assms by (cases F; simp add: spec.continue.mono)

lemma mono2mono[cont_intro, partial_function_mono]:
  assumes "monotone orda (≤) f"
  assumes "x. monotone orda (≤) (λy. g y x)"
  shows "monotone orda (≤) (λx. spec.continue (f x) (g x))"
using assms by (simp add: monotone_def spec.continue.mono)

definition resL :: "('v  ('a, 's, 'w) spec)  ('a, 's, 'w) spec  ('a, 's, 'v) spec" where
  "resL g P = {f. spec.continue f g  P}"

definition resR :: "('a, 's, 'v) spec  ('a, 's, 'w) spec  ('v  ('a, 's, 'w) spec)" where
  "resR f P = {g. spec.continue f g  P}"

interpretation L: galois.complete_lattice_class "λf. spec.continue f g" "spec.continue.resL g" for g
proof
  show "spec.continue f g  P  f  spec.continue.resL g P" (is "?lhs  ?rhs") for f P
  proof(rule iffI)
    assume ?rhs
    then have "spec.continue f g  spec.continue (spec.continue.resL g P) g"
      by (simp add: spec.continue.mono)
    also have "  P"
      by (auto simp: spec.continue.resL_def spec.continue_def)
    finally show ?lhs .
  qed (simp add: spec.continue.resL_def Sup_upper)
qed

interpretation R: galois.complete_lattice_class "λg. spec.continue f g" "spec.continue.resR f"
  for f :: "('a, 's, 'v) spec"
proof
  show "spec.continue f g  P  g  spec.continue.resR f P" (is "?lhs  ?rhs")
   for g :: "'v  ('a, 's, 'w) spec"
   and P :: "('a, 's, 'w) spec"
  proof(rule iffI)
    assume ?rhs
    then have "spec.continue f g  spec.continue f (spec.continue.resR f P)"
      by (simp add: le_fun_def spec.continue.mono)
    also have "  P"
      by (auto simp: spec.continue.resR_def spec.continue_def)
    finally show ?lhs .
  qed (simp add: spec.continue.resR_def Sup_upper)
qed

setup Sign.parent_path

setup Sign.mandatory_path "singleton"

lemma bind_le_conv:
  shows "σ  spec.bind f g  σ  spec.term.none f  σ  spec.continue f g"
by (simp add: spec.bind_def)

lemma bind_le[consumes 1]:
  assumes "σ  f  g"
  obtains
    (incomplete) "σ  spec.term.none f"
  | (continue) σf σg vf
    where "σf  f" and "trace.final σf = trace.init σg" and "trace.term σf = Some vf"
      and "σg  g vf" and "σg  trace.T (trace.init σg) [] None"
      and "σ = trace.T (trace.init σf) (trace.rest σf @ trace.rest σg) (trace.term σg)"
using assms[unfolded spec.singleton.bind_le_conv]
proof(atomize_elim, induct rule: stronger_disjE[consumes 1, case_names incomplete continue])
  case continue
  from σ  spec.continue f g obtain xs ys v w
    where f: "trace.init σ, xs, Some v  f"
      and g: "trace.final' (trace.init σ) xs, ys, w  g v"
      and σ: "σ  trace.T (trace.init σ) (xs @ ys) w"
    by (clarsimp simp: spec.singleton.continue_le_conv)
  with ¬ σ  spec.term.none f obtain ys'
    where "trace.final' (trace.init σ) xs, ys', trace.term σ  g v"
      and "trace.rest σ = xs @ ys'"
      and "trace.natural' (trace.final' (trace.init σ) xs) ys' = []  (y. trace.term σ = Some y)"
    by (atomize_elim, cases σ)
       (auto elim!: trace.less_eqE prefix_append_not_NilE
              elim: order.trans[OF spec.singleton.mono, rotated]
              dest: spec.singleton.mono[OF iffD2[OF trace.less_eq_None(2)[where s="trace.init σ" and σ="trace.T (trace.init σ) xs (Some v)"], simplified]]
                    order.trans[OF spec.singleton.less_eq_None]
              simp: trace.less_eq_None spec.singleton.le_conv trace.natural'.eq_Nil_conv)+
  with f show ?case
    by (cases σ) (force simp: trace.natural_def)
qed blast

setup Sign.parent_path

lemma bind_le[case_names incomplete continue]:
  assumes "spec.term.none f  P"
  assumes "σf σg v. σf  f; trace.init σg = trace.final σf; trace.term σf = Some v; σg  g v;
                      σg  trace.T (trace.init σg) [] None
                  trace.init σf, trace.rest σf @ trace.rest σg, trace.term σg  P"
  shows "f  g  P"
by (rule spec.singleton_le_extI) (use assms in fastforce elim: spec.singleton.bind_le)

setup Sign.mandatory_path "bind"

definition resL :: "('v  ('a, 's, 'w) spec)  ('a, 's, 'w) spec  ('a, 's, 'v) spec" where
  "resL g P = {f. f  g  P}"

lemma incompleteI:
  assumes "s, xs, None  f"
  shows "s, xs, None  f  g"
using assms by (auto simp: spec.bind_def spec.singleton.term.none_le_conv)

lemma continueI:
  assumes f: "s, xs, Some v  f"
  assumes g: "trace.final' s xs, ys, w  g v"
  shows "s, xs @ ys, w  f  g"
using assms by (force simp: spec.bind_def spec.continue_def intro!: disjI2)

lemma singletonL:
  shows "σ  g
      = spec.term.none σ
         {trace.init σ, trace.rest σ @ trace.rest σg, trace.term σg |σg.
               trace.final σ = trace.init σg  (v. trace.term σ = Some v  σg  g v)}" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.bind_le)
    case (continue σf σg v) then show ?case
      by (cases σf; cases σg)
         (simp add: trace.split_Ex;
          metis order.refl spec.singleton.simps(1) trace.final'.natural' trace.stuttering.equiv.append_cong)
  qed force
  show "?rhs  ?lhs"
    by (cases σ)
       (force simp: spec.term.none.singleton spec.singleton.bind_le_conv spec.singleton.continue_le_conv)
qed

lemma mono:
  assumes "f  f'"
  assumes "v. g v  g' v"
  shows "spec.bind f g  spec.bind f' g'"
unfolding spec.bind_def
apply (strengthen ord_to_strengthen(1)[OF assms(1)])
apply (strengthen ord_to_strengthen(1)[OF assms(2)])
apply (rule order.refl)
done

lemma strengthen[strg]:
  assumes "st_ord F f f'"
  assumes "x. st_ord F (g x) (g' x)"
  shows "st_ord F (spec.bind f g) (spec.bind f' g')"
using assms by (cases F; simp add: spec.bind.mono)

lemma mono2mono[cont_intro, partial_function_mono]:
  assumes "monotone orda (≤) f"
  assumes "x. monotone orda (≤) (λy. g y x)"
  shows "monotone orda (≤) (λx. spec.bind (f x) (g x))"
using assms by (simp add: monotone_def spec.bind.mono)

interpretation L: galois.complete_lattice_class "λf. f  g" "spec.bind.resL g" for g
proof
  show "f  g  P  f  spec.bind.resL g P" (is "?lhs  ?rhs") for f P
  proof(rule iffI)
    assume ?rhs
    then have "f  g  spec.bind.resL g P  g"
      by (simp add: spec.bind.mono)
    also have "  P"
      by (simp add: spec.bind.resL_def spec.bind_def spec.term.none.Sup spec.continue.L.lower_Sup)
    finally show ?lhs .
  qed (simp add: spec.bind.resL_def Sup_upper)
qed

lemmas SUPL = spec.bind.L.lower_SUP
lemmas SupL = spec.bind.L.lower_Sup
lemmas supL = spec.bind.L.lower_sup[of f1 f2 g] for f1 f2 g

lemmas INFL_le = spec.bind.L.lower_INF_le
lemmas InfL_le = spec.bind.L.lower_Inf_le
lemmas infL_le = spec.bind.L.lower_inf_le[of f1 f2 g] for f1 f2 g

lemma SUPR:
  shows "spec.bind f (λv. xX. g x v) = (xX. f  g x)  (f  )" (is "?thesis1") ―‹ constSup over typ('a, 's, 'v) spec
    and "spec.bind f (xX. g x) = (xX. f  g x)  (f  )" (is ?thesis2) ―‹ constSup over functions ›
proof -
  show ?thesis1
    by (cases "X = {}")
       (simp_all add: spec.bind_def spec.continue.R.lower_bot sup_SUP ac_simps
                      spec.continue.R.lower_SUP[where f=g and X=X, unfolded Sup_fun_def image_image]
                flip: bot_fun_def)
  then show ?thesis2
    by (simp add: Sup_fun_def image_image)
qed

lemma SUPR_not_empty:
  assumes "X  {}"
  shows "spec.bind f (λv. xX. g x v) = (xX. f  g x)"
using assms by (clarsimp simp: spec.bind.SUPR spec.bind.mono sup.absorb1 SUPI simp flip: ex_in_conv)

lemmas supR = spec.bind.SUPR_not_empty[where g=id and X="{g1, g2}" for g1 g2, simplified]

lemma InfR_le:
  shows "spec.bind f (λv. xX. g x v)  (xX. f  g x)"
by (meson INF_lower order.refl le_INF_iff spec.bind.mono)

lemma infR_le:
  shows "spec.bind f (g1  g2)  (f  g1)  (f  g2)"
    and "spec.bind f (λv. g1 v  g2 v)  (f  g1)  (f  g2)"
by (simp_all add: spec.bind.mono)

lemma Inf_le:
  shows "spec.bind (xX. f x) (λv. (xX. g x v))  (xX. spec.bind (f x) (g x))"
by (auto simp: le_INF_iff intro: spec.bind.mono)

lemma inf_le:
  shows "spec.bind (f1  f2) (λv. g1 v  g2 v)  spec.bind f1 g1  spec.bind f2 g2"
by (simp add: spec.bind.mono)

lemma mcont2mcont[cont_intro]:
  assumes "mcont luba orda Sup (≤) f"
  assumes "v. mcont luba orda Sup (≤) (λx. g x v)"
  shows "mcont luba orda Sup (≤) (λx. spec.bind (f x) (g x))"
proof(rule ccpo.mcont2mcont'[OF complete_lattice_ccpo _ _ assms(1)])
  show "mcont Sup (≤) Sup (≤) (λf. bind f (g x))" for x
    by (intro mcontI contI monotoneI) (simp_all add: spec.bind.mono flip: spec.bind.SUPL)
  show "mcont luba orda Sup (≤) (λx. bind f (g x))" for f
    by (intro mcontI monotoneI contI)
       (simp_all add: mcont_monoD[OF assms(2)] spec.bind.mono
                flip: spec.bind.SUPR_not_empty contD[OF mcont_cont[OF assms(2)]])
qed

lemmas botL[simp] = spec.bind.L.lower_bot

lemma botR:
  shows "f   = spec.term.none f"
by (simp add: spec.bind_def spec.continue.R.lower_bot)

lemma eq_bot_conv:
  shows "spec.bind f g =   f = "
by (fastforce simp: spec.continue.L.lower_bot spec.bind_def spec.term.galois simp flip: bot.extremum_unique)

lemma idleL[simp]:
  shows "spec.idle  g = spec.idle"
by (simp add: spec.idle_def spec.bind.SupL image_image spec.bind.singletonL spec.term.none.singleton)

lemma idleR:
  shows "f  spec.idle = f  " (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (fastforce simp: spec.bind.botR trace.split_all spec.singleton.le_conv
                intro!: spec.bind_le
                 intro: spec.bind.incompleteI order.trans[rotated])
  show "?rhs  ?lhs"
    by (simp add: spec.bind.mono)
qed

lemmas ifL = if_distrib[where f="λf. spec.bind f g" for g]

setup Sign.parent_path

setup Sign.mandatory_path "idle"

lemma bind_le_conv[spec.idle_le]:
  shows "spec.idle  f  g  spec.idle  f" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (fastforce simp: spec.idle_def spec.singleton.mono trace.less_eq_None spec.singleton.bind_le_conv
                        spec.singleton.term.none_le_conv spec.singleton.continue_le_conv
                  elim: order.trans[rotated])
  show "?rhs  ?lhs"
    by (simp add: spec.bind_def spec.idle.term.none_le_conv le_supI1)
qed

setup Sign.parent_path

setup Sign.mandatory_path "term"

setup Sign.mandatory_path "none"

lemma bindL_le[iff]:
  shows "spec.term.none f  f  g"
by (simp add: spec.bind_def)

lemma bind:
  shows "spec.term.none (f  g) = f  (λv. spec.term.none (g v))"
by (rule spec.singleton.antisym)
   (auto elim: spec.singleton.bind_le
         simp: trace.split_all spec.bind.incompleteI spec.bind.continueI spec.singleton.term.none_le_conv)

setup Sign.parent_path

setup Sign.mandatory_path "all"

lemma bind:
  shows "spec.term.all (f  g) = spec.term.all f  (f  (λv. spec.term.all (g v)))" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI])
  show "σ  ?rhs" if "σ  ?lhs" for σ
    using that
    by (cases σ)
       (fastforce simp: trace.split_all spec.singleton.le_conv
                 elim!: spec.singleton.bind_le
                 intro: spec.bind.continueI)
  show "?rhs  ?lhs"
    by (simp add: spec.term.none.sup spec.term.none.bind spec.bind.mono flip: spec.term.galois)
qed

setup Sign.parent_path

setup Sign.parent_path

paragraph‹ The monad laws for constspec.bind. \label{sec:safety_logic-monad_laws} ›

setup Sign.mandatory_path "bind"

lemma bind:
  fixes f :: "(_, _, _) spec"
  shows "f  g  h = f  (λv. g v  h)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.bind_le)
    case incomplete show ?case
      by (simp add: spec.bind.mono spec.term.none.bind)
  next
    case (continue σfg σh v) then show ?case
      by (cases σh)
         (fastforce elim: spec.singleton.bind_le spec.bind.continueI
                    simp: spec.singleton.le_conv trace.split_all)
  qed
  show "?rhs  ?lhs"
  proof(induct rule: spec.bind_le)
    case incomplete show ?case
      by (strengthen ord_to_strengthen(2)[OF spec.term.none.bindL_le])
         (simp add: spec.term.none.bind)
  next
    case (continue σf σgh v)
    note * = continue.hyps(1-3)
    from σgh  g v  h show ?case
    proof(cases rule: spec.singleton.bind_le)
      case incomplete with * show ?thesis
        by (cases σf)
           (clarsimp simp: spec.singleton.le_conv spec.bind.incompleteI spec.bind.continueI)
    next
      case (continue σg σh vg) with * show ?thesis
        by (cases σf; cases σg)
           (simp flip: append_assoc; fastforce intro!: spec.bind.continueI)
    qed
  qed
qed

lemmas assoc = spec.bind.bind

lemma returnL_le:
  shows "g v  spec.return v  g" (is "?lhs  ?rhs")
proof(rule spec.singleton_le_extI)
  show "σ  ?rhs" if "σ  ?lhs" for σ
    by (rule spec.bind.continueI[where xs="[]" and s="trace.init σ" and ys="trace.rest σ" and w="trace.term σ" and v=v, simplified])
       (simp_all add: spec.return_def spec.action.stutterI that)
qed

lemma returnL:
  assumes "spec.idle  g v"
  shows "spec.return v  g = g v"
by (rule antisym[OF spec.bind_le spec.bind.returnL_le])
   (simp_all add: assms spec.term.none.return spec.singleton.return_le_conv trace.split_all)

lemma returnR[simp]:
  shows "f  spec.return = f" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (auto intro: spec.bind_le
              simp: trace.split_all spec.singleton.return_le_conv order.trans[OF spec.singleton.less_eq_None(1)]
             split: option.split_asm)
  show "σ  ?lhs" if "σ  ?rhs" for σ
    using that
    by (cases "σ"; cases "trace.term σ";
        clarsimp simp: spec.bind.incompleteI spec.bind.continueI[where ys="[]", simplified] spec.singleton.le_conv)
qed

lemma return: ―‹ Does not require spec.idle ≤ g v›
  fixes f :: "('a, 's, 'v) spec"
  fixes g :: "'v  'x  ('a, 's, 'w) spec"
  shows "f  (λv. spec.return x  g v) = f  (λv. g v x)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.bind_le)
    case (continue σf σrg v)
    from σrg  spec.return x  g v show ?case
    proof(induct rule: spec.singleton.bind_le[case_names incomplete continue2])
      case incomplete with σf  f trace.init σrg = trace.final σf show ?thesis
        by (cases σf)
           (auto simp: spec.term.none.return spec.singleton.le_conv
                intro: spec.bind.incompleteI order.trans[rotated])
    next
      case (continue2 σr σg vr) with continue show ?case
        by (cases σf) (simp add: trace.split_all spec.singleton.le_conv spec.bind.continueI)
    qed
  qed simp
  show "?rhs  ?lhs"
    by (simp add: spec.bind.mono spec.bind.returnL_le)
qed

setup Sign.mandatory_path "term"

lemma noneL[simp]:
  shows "spec.term.none f  g = spec.term.none f"
by (simp add: spec.bind.bind flip: spec.bind.botR bot_fun_def)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "map"

lemma bind_le: ―‹ Converse does not hold: it may be that no final states of f› satisfy g›
  fixes f :: "('a, 's, 'v) spec"
  fixes g :: "'v  ('a, 's, 'w) spec"
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'w  'x"
  shows "spec.map af sf vf (f  g)  spec.map af sf id f  (λv. spec.map af sf vf (g v))"
by (subst (1) spec.map_def)
   (force simp: spec.singleton.le_conv trace.split_all trace.final'.map
         intro: spec.bind.incompleteI spec.bind.continueI
          elim: spec.singleton.bind_le)

lemma bind_inj_sf:
  fixes f :: "('a, 's, 'x) spec"
  fixes g :: "'x  ('a, 's, 'v) spec"
  assumes "inj sf"
  shows "spec.map af sf vf (f  g) = spec.map af sf id f  (λv. spec.map af sf vf (g v))" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.map.bind_le])
  show "?rhs  ?lhs"
  proof(induct rule: spec.bind_le)
    case incomplete show ?case
      by (metis spec.map.mono spec.term.none.bindL_le spec.term.none.map_gen)
  next
    case (continue σf σg v)
    from continue(1,4) obtain σf' σg'
      where *: "σf'  f" "σf  trace.map af sf id σf'"
               "σg'  g v" "σg  trace.map af sf vf σg'"
      by (clarsimp simp: spec.singleton.le_conv)
    with continue(2,3)
    have "sf (trace.init σg') = sf (trace.final σf')"
      by (cases σf; cases σg; cases σf'; cases σg'; clarsimp)
         (clarsimp simp: spec.singleton_le_conv simp flip: trace.final'.map[where af=af and sf=sf];
          erule trace.less_eqE; simp add: trace.natural.trace_conv; metis trace.final'.natural')
    with continue(2,3) * show ?case
      by (cases σf; cases σg; cases σf'; cases σg')
         (fastforce dest: inj_onD[OF assms, simplified]
                    elim: trace.less_eqE spec.bind.continueI
                    simp: spec.singleton.le_conv trace.final'.map trace.less_eq_None
                          spec.singleton_le_conv trace.natural_def trace.natural'.append)
  qed
qed

setup Sign.parent_path

setup Sign.mandatory_path "vmap"

lemma eq_return: ―‹ generalizes @{thm [source] "spec.bind.returnR"}
  shows "spec.vmap vf P = P  spec.return  vf" (is ?thesis1)
    and "spec.vmap vf P = P  (λv. spec.return (vf v))" (is "?lhs = ?rhs") ―‹ useful for flip/symmetric ›
proof -
  show "?lhs = ?rhs"
  proof(rule antisym)
    show "?lhs  ?rhs"
      by (rule spec.singleton.exhaust[of P])
         (fastforce simp: trace.split_all spec.singleton.le_conv spec.map.Sup spec.map.singleton map_option_case
                   intro: spec.bind.incompleteI spec.bind.continueI[where ys="[]", simplified]
                   split: option.split)
next
    show "?rhs  ?lhs"
      by (rule spec.bind_le)
         (force simp: trace.split_all spec.singleton.le_conv trace.less_eq_None trace.natural.mono
                      spec.term.galois spec.term.all.expansive spec.term.all.map spec.map.id
               split: option.split_asm)+
  qed
  then show ?thesis1
    by (simp add: comp_def)
qed

lemma unitL: ―‹ monomorphise ignored return values ›
  shows "f  g = spec.vmap () f  g"
by (simp add: spec.vmap.eq_return comp_def spec.bind.bind spec.bind.return)

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemma bind:
  fixes f :: "('b, 't, 'v) spec"
  fixes g :: "'v  ('b, 't, 'x) spec"
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'w  'x"
  shows "spec.invmap af sf vf (f  g) = spec.invmap af sf id f  (λv. spec.invmap af sf vf (g v))" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI])
  fix σ assume "σ  ?lhs"
  then have "trace.map af sf vf σ  f  g" by (simp add: spec.singleton.le_conv)
  then show "σ  ?rhs"
  proof(induct rule: spec.singleton.bind_le)
    case incomplete then show ?case
      by (cases σ) (clarsimp simp: spec.singleton.le_conv spec.bind.incompleteI)
  next
    case (continue σf σg vf) then show ?case
      by (cases σ; cases σf; cases σg)
         (clarsimp simp: spec.bind.continueI map_eq_append_conv spec.singleton.le_conv trace.final'.map)
  qed
next
  show "?rhs  ?lhs"
    by (simp add: order.trans[OF spec.map.bind_le] spec.bind.mono spec.map_invmap.lower_upper_contractive
            flip: spec.map_invmap.galois)
qed

lemma split_vinvmap:
  shows "spec.invmap af sf vf P = spec.invmap af sf id P  (λv. v'vf -` {v}. spec.return v')" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI])
  show "σ  ?rhs" if "σ  ?lhs" for σ
    using that
    by (cases σ; cases "trace.term σ")
       (auto simp: spec.singleton.le_conv
            intro: spec.bind.incompleteI spec.bind.continueI[where ys="[]", simplified])
  show "?rhs  ?lhs"
  proof(induct rule: spec.bind_le)
    case (continue σf σg v) then show ?case
      by (cases σf; cases "trace.term σg")
         (auto simp: spec.singleton.le_conv split: option.split_asm elim: order.trans[rotated])
  qed (simp add: spec.term.none.invmap_gen[where vf'=vf] spec.invmap.mono)
qed

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma return_const:
  assumes "V  {}"
  assumes "W  {}"
  shows "spec.action (V × F) = spec.action (W × F)  (vV. spec.return v)" (is "?lhs = ?rhs")
proof(rule antisym)
  from W  {} show "?lhs  ?rhs"
    by - (rule spec.action_le;
          fastforce intro: spec.bind.continueI[where xs="[x]" and v="SOME w. w  W" for x, simplified]
                           spec.action.stepI
                     simp: some_in_eq spec.singleton.le_conv spec.singleton.action_le_conv
                           spec.idle.action_le spec.idle.bind_le_conv)
  from V  {} show "?rhs  ?lhs"
    by - (rule spec.bind_le,
           fastforce simp: spec.term.galois spec.term.all.action intro: le_supI1 spec.action.mono,
           auto 0 3 simp: spec.singleton.le_conv spec.singleton.action_le_conv
               simp flip: trace.steps'.empty_conv
                simp del: trace.steps'.simps split: option.splits)
qed

setup Sign.parent_path

setup Sign.mandatory_path "term.closed"

lemma bind_all_return:
  assumes "f  spec.term.closed _"
  shows "f  (range spec.return) = spec.term.all f" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (subst (2) spec.term.closed_conv[OF assms])
       (simp add: spec.term.none.bind spec.term.none.Sup image_image spec.term.none.return
                  spec.bind.botR spec.bind.idleR
                  spec.term.all.monomorphic
            flip: spec.term.galois)
next
  fix σ
  assume "σ  ?rhs"
  then obtain v where "trace.init σ, trace.rest σ, Some v  f"
    by (subst (asm) spec.term.closed_conv[OF assms])
       (force simp: spec.singleton.le_conv spec.term.all.monomorphic)
  then show "σ  ?lhs"
    by (cases σ; cases "trace.term σ")
       (auto simp: spec.singleton.le_conv spec.bind.continueI[where ys="[]", simplified]
            split: option.split)
qed

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Kleene star \label{sec:safety_logic-kleene} ›

text‹

We instantiate the generic Kleene locale with monomorphic spec.return ()›. The polymorphic
(⨆v. spec.return v)› fails the comp_unitR› axiom (ε ≤ x ⟹ x ∙ ε = x"›).

›

setup Sign.mandatory_path "spec"

interpretation kleene: weak_kleene "spec.return ()" "λx y. spec.bind x y"
by standard (simp_all add: spec.bind.bind spec.bind.supL spec.bind.supR
                           spec.bind.returnL order.trans[OF spec.idle.return_le])

setup Sign.mandatory_path "idle.kleene"

lemmas star_le[spec.idle_le] = order.trans[OF spec.idle.return_le spec.kleene.epsilon_star_le]

lemmas rev_star_le[spec.idle_le] = spec.idle.kleene.star_le[unfolded spec.kleene.star_rev_star]

setup Sign.parent_path

setup Sign.mandatory_path "return.kleene"

lemmas star_le = spec.kleene.epsilon_star_le

lemmas rev_star_le = spec.return.kleene.star_le[unfolded spec.kleene.star_rev_star]

setup Sign.parent_path

setup Sign.mandatory_path "kleene"

lemma star_idle:
  shows "spec.kleene.star spec.idle = spec.return ()"
by (subst spec.kleene.star.simps) (simp add: sup.absorb2 spec.idle.return_le)

lemmas rev_star_idle = spec.kleene.star_idle[unfolded spec.kleene.star_rev_star]

setup Sign.parent_path

setup Sign.mandatory_path "term.all.kleene"

lemma star_closed_le:
  fixes P :: "(_, _, unit) spec"
  assumes "P  spec.term.closed _"
  shows "spec.term.all (spec.kleene.star P)  spec.kleene.star P" (is "_  ?rhs")
proof(induct rule: spec.kleene.star.fixp_induct[where P="λR. spec.term.all (R P)  ?rhs", case_names adm bot step])
  case (step R) show ?case
    by (auto simp: spec.term.all.sup spec.term.all.bind spec.kleene.expansive_star spec.term.all.return
        simp flip: spec.term.all.closed_conv[OF assms]
            intro: spec.kleene.epsilon_star_le
                   order.trans[OF spec.bind.mono[OF order.refl step] spec.kleene.fold_starL])
qed simp_all

setup Sign.parent_path

setup Sign.mandatory_path "term.closed.kleene"

lemma star:
  assumes "P  spec.term.closed _"
  shows "spec.kleene.star P  spec.term.closed _"
by (rule spec.term.closed_clI)
   (simp add: spec.term.all.kleene.star_closed_le[OF assms] spec.term.all.monomorphic)

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Transition relations \label{sec:safety_logic-trans_rel} ›

text‹

Using constspec.kleene.star we can specify the transitions each agent is allowed to perform.
These constraints ((⊓) spec.rel r›) distribute through all program constructs (for suitable r›).

Observations:
  the Galois connection between spec.rel› and spec.steps› is much easier to show in the powerset model
   see citet‹Footnote~2› in "vanStaden:2015"
  most useful facts about spec.steps› depend on the model

›

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "rel"

definition act :: "('a, 's) steps  ('a, 's, unit) spec" where ―‹ lift above constspec.return to ease some proofs ›
  "act r = spec.action ({()} × (r  UNIV × Id))"

abbreviation monomorphic :: "('a, 's) steps  ('a, 's, unit) spec" where
  "monomorphic r  spec.kleene.star (spec.rel.act r)"

lemma act_alt_def:
  shows "spec.rel.act r = spec.action ({()} × r)  spec.return ()"
by (simp add: spec.rel.act_def spec.return_def Sigma_Un_distrib2 flip: spec.action.sup)

setup Sign.parent_path

definition rel :: "('a, 's) steps  ('a, 's, 'v) spec" where
  "rel r = spec.term.all (spec.rel.monomorphic r)"

definition steps :: "('a, 's, 'v) spec  ('a, 's) steps" where
  "steps P = {r. P  spec.rel r}"

setup Sign.mandatory_path "rel.act"

lemma monotone:
  shows "mono spec.rel.act"
proof(rule monotoneI)
  show "spec.rel.act r  spec.rel.act r'" if "r  r'" for r r' :: "('a, 's) steps"
    using that unfolding spec.rel.act_def by (strengthen ord_to_strengthen(1)[OF r  r']) simp
qed

lemmas strengthen[strg] = st_monotone[OF spec.rel.act.monotone]
lemmas mono = monotoneD[OF spec.rel.act.monotone]

lemma empty:
  shows "spec.rel.act {} = spec.return ()"
by (simp add: spec.rel.act_def spec.return_def spec.action.empty)

lemma UNIV:
  shows "spec.rel.act UNIV = spec.action ({()} × UNIV)"
by (simp add: spec.rel.act_def)

lemma sup:
  shows "spec.rel.act (r  s) = spec.rel.act r  spec.rel.act s"
by (fastforce simp: spec.rel.act_def simp flip: spec.action.sup intro: arg_cong[where f=spec.action])

lemma stutter:
  shows "spec.rel.act (UNIV × Id) = spec.return ()"
by (simp add: spec.rel.act_def spec.return_def)

setup Sign.parent_path

setup Sign.mandatory_path "term"

setup Sign.mandatory_path "all"

setup Sign.mandatory_path "rel"

lemma act_mono:
  shows "spec.term.all (spec.rel.act r) = spec.rel.act r"
by (simp add: spec.rel.act_alt_def spec.term.all.sup spec.term.all.action spec.term.all.return UNIV_unit)

setup Sign.parent_path

lemma rel:
  shows "spec.term.all (spec.rel r) = spec.rel r"
by (simp add: spec.rel_def)

setup Sign.parent_path

setup Sign.mandatory_path "closed"

setup Sign.mandatory_path "rel"

lemma act:
  shows "spec.rel.act r  spec.term.closed _"
by (metis spec.term.all.rel.act_mono spec.term.all.closed)

setup Sign.parent_path

lemma rel:
  shows "spec.rel r  spec.term.closed _"
by (metis spec.term.all.closed spec.term.all.rel)

setup Sign.parent_path

setup Sign.mandatory_path "none"

lemma inf_none_rel: ―‹ polymorphic constants ›
  shows "spec.term.none (spec.rel r :: ('a, 's, 'w) spec)  spec.term.none P
       = spec.rel r  (spec.term.none P :: ('a, 's, 'v) spec)" (is ?thesis1)
    and "spec.term.none P  spec.term.none (spec.rel r :: ('a, 's, 'w) spec)
       = spec.term.none P  (spec.rel r :: ('a, 's, 'v) spec)" (is ?thesis2)
proof -
  show ?thesis1
    by (metis spec.term.closed.rel spec.term.closed.none_inf(1)
              spec.term.none.idempotent spec.term.none.inf(2) spec.term.none_all spec.term.all.rel)
  then show ?thesis2
    by (simp add: ac_simps)
qed

lemma inf_rel:
  shows "spec.term.none P  spec.rel r = spec.term.none (P  spec.rel r)" (is ?thesis1)
    and "spec.rel r  spec.term.none P = spec.term.none (spec.rel r  P)" (is ?thesis2)
by (simp_all add: ac_simps spec.term.none.inf(2) spec.term.none.inf_none_rel(2))

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "return"

setup Sign.mandatory_path "rel"

lemma act_le:
  shows "spec.return ()  spec.rel.act r"
by (simp add: spec.rel.act.mono flip: spec.rel.act.empty)

setup Sign.parent_path

lemma rel_le:
  shows "spec.return v  spec.rel r"
by (simp add: spec.rel_def spec.term.none.return spec.idle.kleene.star_le flip: spec.term.galois)

lemma Sup_rel_le:
  shows "range spec.return  spec.rel r"
by (simp add: spec.return.rel_le)

setup Sign.parent_path

setup Sign.mandatory_path "idle"

setup Sign.mandatory_path "rel"

lemmas act_le[spec.idle_le] = order.trans[OF spec.idle.return_le spec.return.rel.act_le]

setup Sign.parent_path

lemmas rel_le[spec.idle_le] = order.trans[OF spec.idle.return_le spec.return.rel_le]

setup Sign.parent_path

setup Sign.mandatory_path "singleton"

setup Sign.mandatory_path "rel"

setup Sign.mandatory_path "act"

lemma le_conv[spec.singleton.le_conv]:
  shows "σ  spec.rel.act r  trace.steps σ = {}  (xr. trace.steps σ = {x})"
by (auto simp: spec.rel.act_def spec.singleton.le_conv spec.singleton.action_le_conv trace.steps'.step_conv
        split: option.split)

setup Sign.parent_path

setup Sign.mandatory_path "monomorphic"

lemma le_steps:
  assumes "trace.steps σ  r"
  shows "σ  spec.rel.monomorphic r"
using assms
proof(induct "trace.rest σ" arbitrary: σ rule: rev_induct)
  case Nil then show ?case
    by (simp add: spec.singleton.rel.act.le_conv order.trans[OF _ spec.kleene.expansive_star])
next
  case (snoc x xs σ)
  from snoc(2,3)
  have *: "trace.init σ, xs, Some ()  spec.rel.monomorphic r"
    by (cases σ) (fastforce intro: snoc(1) simp: trace.steps'.append)
  have **: "trace.final' (trace.init σ) xs, [x], trace.term σ  spec.rel.act r"
  proof(cases "trace.final' (trace.init σ) xs = snd x")
    case True with snoc.prems snoc.hyps(2) show ?thesis
      by (simp add: spec.singleton.le_conv)
  next
    case False with snoc.prems snoc.hyps(2) show ?thesis
      by (cases σ) (clarsimp simp: spec.singleton.le_conv trace.steps'.append)
  qed
  show ?case
    by (rule order.trans[OF spec.bind.continueI[OF * **, simplified snoc.hyps(2) trace.t.collapse]
                            spec.kleene.fold_starR])
qed

setup Sign.parent_path

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "rel.act"

lemmas mono_le = spec.kleene.expansive_star

setup Sign.parent_path

setup Sign.mandatory_path "rel.monomorphic"

lemma alt_def:
  shows "spec.rel.monomorphic r = (spec.singleton ` {σ. trace.steps σ  r})" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.kleene.star.fixp_induct[case_names adm bot step])
    case (step R)
    have "spec.return ()  ?rhs"
      by (force intro: spec.singleton_le_extI simp: spec.singleton.le_conv
                 dest: trace.steps'.simps(5))
    moreover
    have "spec.rel.act r  ?rhs  ?rhs"
    proof(induct rule: spec.bind_le)
      case incomplete show ?case
        by (rule spec.singleton_le_extI)
           (clarsimp simp: spec.singleton.le_conv;
            metis order.refl empty_subsetI insert_subsetI trace.steps'.empty_conv(1))
    next
      case (continue σf σg v) then show ?case
        by (fastforce intro!: exI[where x="trace.T (trace.init σf) (trace.rest σf @ trace.rest σg) (trace.term σg)"]
                       simp: trace.steps'.append spec.singleton.rel.act.le_conv
                       dest: trace.steps.mono[OF iffD1[OF spec.singleton_le_conv], simplified,
                                                             simplified trace.steps'.natural'])
    qed
    ultimately show ?case
      by - (strengthen ord_to_strengthen(1)[OF step]; simp)
  qed simp_all
  show "?rhs  ?lhs"
    by (simp add: spec.singleton.rel.monomorphic.le_steps)
qed

setup Sign.parent_path

setup Sign.mandatory_path "singleton"

setup Sign.mandatory_path "rel"

lemma monomorphic_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.rel.monomorphic r  trace.steps σ  r"
by (fastforce simp: spec.rel.monomorphic.alt_def spec.singleton_le_conv trace.steps'.natural'
              dest: trace.steps.mono)

setup Sign.parent_path

lemma rel_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.rel r  trace.steps σ  r"
by (cases σ) (auto simp add: spec.rel_def spec.singleton.le_conv)

setup Sign.parent_path

interpretation rel: galois.complete_lattice_class spec.steps spec.rel
proof(rule galois.upper_preserves_InfI)
  show "mono spec.rel"
    by (simp add: monoD monotoneI spec.kleene.monotone_star spec.rel.act.mono spec.rel_def)
  show "(xX. spec.rel x)  spec.rel (X)" for X :: "('a, 'b) steps set"
    by (fastforce intro: spec.singleton_le_extI simp: le_INF_iff spec.singleton.le_conv)
qed (simp add: spec.steps_def)

lemma rel_alt_def:
  shows "spec.rel r = (spec.singleton ` {σ. trace.steps σ  r})"
by (simp flip: spec.singleton.rel_le_conv)

setup Sign.mandatory_path "vmap"

lemma unit_rel:
  shows "spec.vmap () (spec.rel r) = spec.rel r"
by (simp add: spec.rel_def spec.term.all.vmap_unit_absorb)

setup Sign.parent_path

setup Sign.mandatory_path "rel"

lemma monomorphic_conv: ―‹ if the return type is typunit
  shows "spec.rel r = spec.rel.monomorphic r"
by (simp add: spec.rel_def
        flip: spec.term.all.closed_conv[OF spec.term.closed.kleene.star[OF spec.term.closed.rel.act]])

lemma monomorphic_act_le: ―‹ typunit return type ›
  shows "spec.rel.act r  spec.rel r"
by (simp add: spec.rel.monomorphic_conv spec.rel.act.mono_le)

lemma empty:
  shows "spec.rel {} = (v. spec.return v)"
by (simp add: spec.rel_def spec.kleene.star_epsilon spec.rel.act.empty spec.term.all.return)

lemmas UNIV = spec.rel.upper_top
lemmas top = spec.rel.UNIV

lemmas INF = spec.rel.upper_INF
lemmas Inf = spec.rel.upper_Inf
lemmas inf = spec.rel.upper_inf

lemmas Sup_le = spec.rel.Sup_upper_le
lemmas sup_le = spec.rel.sup_upper_le ―‹ Converse does not hold: the RHS allows interleaving of r› and s› steps ›

lemma reflcl:
  shows "spec.rel (r  A × Id) = spec.rel r"
    and "spec.rel (A × Id  r) = spec.rel r"
by (simp_all add: spec.rel_def spec.rel.act_def ac_simps flip: Times_Un_distrib1)

lemma minus_Id:
  shows "spec.rel (r - A × Id) = spec.rel r"
by (metis Un_Diff_cancel spec.rel.reflcl(2))

lemma Id:
  shows "spec.rel (A × Id) = (v. spec.return v)"
by (subst spec.rel.minus_Id[where A=A, symmetric]) (simp add: spec.rel.empty)

lemmas monotone = spec.rel.monotone_upper
lemmas mono = monotoneD[OF spec.rel.monotone, of r r' for r r']

lemma mono_reflcl:
  assumes "r  s  UNIV × Id"
  shows "spec.rel r  spec.rel s"
by (metis assms spec.rel.mono spec.rel.reflcl(1))

lemma unfoldL:
  shows "spec.rel r = spec.rel.act r  spec.rel r" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (rule order.trans[OF spec.bind.returnL_le spec.bind.mono[OF spec.return.rel.act_le order.refl]])
  show "?rhs  ?lhs"
    by (subst (2) spec.rel_def, subst spec.kleene.star_unfoldL)
       (simp add: spec.term.all.sup spec.term.all.bind le_supI1 flip: spec.rel_def)
qed

lemma foldR: ―‹ arbitrary interstitial return type ›
  shows "spec.rel r  spec.rel.act r = spec.rel r" (is "?lhs = ?rhs")
proof -
  have "?lhs = spec.rel.monomorphic r  spec.rel.act r"
    by (subst spec.vmap.unitL) (simp add: spec.vmap.unit_rel spec.rel.monomorphic_conv)
  also have " = ?rhs"
  proof(rule antisym)
    show "spec.rel.monomorphic r  spec.rel.act r  ?rhs"
      by (simp add: spec.kleene.fold_starR spec.rel.monomorphic_conv)
    show "?rhs spec.rel.monomorphic r  spec.rel.act r"
      by (simp add: spec.rel.monomorphic_conv)
         (rule spec.bind.mono[OF order.refl spec.return.rel.act_le, where 'c=unit, simplified])
  qed
  finally show ?thesis .
qed

lemma wind_bind: ―‹ arbitrary interstitial return type ›
  shows "spec.rel r  spec.rel r = spec.rel r" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.bind_le)
    case incomplete show ?case
      by (simp add: spec.term.all.rel spec.term.galois)
  next
    case (continue σf σg v) then show ?case
      by (simp add: spec.singleton.rel_le_conv trace.steps'.append)
  qed
  show "?rhs  ?lhs"
    by (meson order.trans order.refl spec.bind.mono spec.bind.returnL_le spec.return.rel_le)
qed

lemma wind_bind_leading: ―‹ arbitrary interstitial return type ›
  assumes "r'  r"
  shows "spec.rel r'  spec.rel r = spec.rel r" (is "?lhs = ?rhs")
proof(rule antisym)
  from assms show "?lhs  ?rhs"
    by (metis order.refl spec.bind.mono spec.rel.mono spec.rel.wind_bind)
  show "?rhs  ?lhs"
    by (meson order.trans spec.eq_iff spec.bind.mono spec.bind.returnL_le spec.return.rel_le)
qed

lemma wind_bind_trailing: ―‹ arbitrary interstitial return type ›
  assumes "r'  r"
  shows "spec.rel r  spec.rel r' = spec.rel r" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  from assms show "?lhs  ?rhs"
    by (metis order_refl spec.bind.mono spec.rel.mono spec.rel.wind_bind)
  show "σ  ?lhs" if "σ  ?rhs" for σ
    using that
    by (cases σ)
       (force simp: spec.singleton.le_conv spec.singleton.bind_le_conv spec.singleton.continue_le_conv)
qed

text‹ Interstitial unit, for unfolding ›

lemmas unwind_bind = spec.rel.wind_bind[where 'd=unit, symmetric]
lemmas unwind_bind_leading = spec.rel.wind_bind_leading[where 'd=unit, symmetric]
lemmas unwind_bind_trailing = spec.rel.wind_bind_trailing[where 'd=unit, symmetric]

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemma rel:
  shows "spec.invmap af sf vf (spec.rel r) = spec.rel (map_prod af (map_prod sf sf) -` (r  UNIV × Id))"
by (fastforce intro: antisym spec.singleton_le_extI simp: spec.singleton.invmap_le_conv spec.singleton.rel_le_conv trace.steps'.map)

lemma range:
  shows "spec.invmap af sf vf P = spec.invmap af sf vf (P  spec.rel (range af × range sf × range sf))"
by (rule antisym[OF spec.singleton_le_extI])
    (auto simp: spec.singleton.le_conv trace.steps'.map spec.invmap.mono)

setup Sign.parent_path

setup Sign.mandatory_path "map"

lemma inf_rel:
  shows "spec.map af sf vf P  spec.rel r
       = spec.map af sf vf (P  spec.rel (map_prod af (map_prod sf sf) -` (r  UNIV × Id)))"
    and "spec.rel r  spec.map af sf vf P
       = spec.map af sf vf (spec.rel (map_prod af (map_prod sf sf) -` (r  UNIV × Id))  P)"
by (simp_all add: spec.invmap.rel spec.map.inf_distr ac_simps)

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma rel_le:
  fixes F :: "('v × 'a × 's × 's) set"
  fixes r :: "('a, 's) steps"
  assumes "v a s s'. (v, a, s, s')  F  (a, s, s')  r  s = s'"
  shows "spec.action F  spec.rel r"
unfolding spec.rel_def
by (strengthen ord_to_strengthen(2)[OF spec.kleene.expansive_star])
   (fastforce simp: spec.rel.act_def spec.term.all.action intro: le_supI1 spec.action.mono dest: assms)

setup Sign.parent_path

setup Sign.mandatory_path "kleene"

lemma star_le:
  assumes "S  spec.rel r"
  shows "spec.kleene.star S  spec.rel r"
by (strengthen ord_to_strengthen(1)[OF assms])
   (simp add: spec.rel_def spec.kleene.idempotent_star
        flip: spec.term.all.closed_conv[OF spec.term.closed.kleene.star[OF spec.term.closed.rel.act]])

setup Sign.parent_path

setup Sign.mandatory_path "bind"

lemma relL_le:
  shows "g x  spec.rel r  g"
by (rule order.trans[OF spec.bind.returnL_le spec.bind.mono[OF spec.return.rel_le order.refl]])

lemma relR_le:
  shows "f  f  spec.rel r"
by (rule order.trans[OF eq_refl[OF spec.bind.returnR[symmetric]]
                        spec.bind.mono[OF order.refl spec.return.rel_le]])

lemma inf_rel:
  shows "(f  g)  spec.rel r = (spec.rel r  f)  (λx. spec.rel r  g x)" (is ?thesis1)
    and "spec.rel r  (f  g) = (spec.rel r  f)  (λx. spec.rel r  g x)" (is "?lhs = ?rhs")
proof -
  show "?lhs = ?rhs"
  proof(rule antisym[OF spec.singleton_le_extI])
    fix σ assume lhs: "σ  ?lhs"
    then have "σ  f  g" by simp
    then show "σ  ?rhs"
    proof(cases rule: spec.singleton.bind_le)
      case incomplete with lhs show ?thesis
        by (cases σ) (simp add: spec.singleton.le_conv spec.bind.incompleteI)
    next
      case (continue σf σg vf) with lhs show ?thesis
        by (cases σf) (simp add: spec.singleton.le_conv trace.steps'.append spec.bind.continueI)
    qed
  next
    show "?rhs  ?lhs"
    proof(induct rule: spec.bind_le)
      case incomplete show ?case
        by (auto simp: spec.term.none.inf spec.term.galois spec.term.all.rel intro: le_infI1 le_infI2)
    next
      case (continue σf σg v) then show ?case
        by (cases σf; cases σg) (simp add: spec.singleton.rel_le_conv spec.bind.continueI trace.steps'.append)
    qed
  qed
  then show ?thesis1
    by (simp add: ac_simps)
qed

lemma inf_rel_distr_le:
  shows "(f  spec.rel r)  (λv. g1 v  g2)  (f  g1)  (spec.rel r  (λ_::unit. g2))"
by (rule spec.bind_le;
    force simp: trace.split_all spec.singleton.le_conv spec.term.galois spec.term.none.inf
                spec.term.all.bind spec.term.all.rel
         intro: le_infI1 le_infI2 spec.bind.continueI)

setup Sign.parent_path

setup Sign.mandatory_path "singleton"

lemma inf_rel:
  shows "σ  spec.rel r = (spec.singleton ` {σ'. σ'  σ  trace.steps σ'  r})" (is "?lhs = ?rhs")
    and "spec.rel r  σ = (spec.singleton ` {σ'. σ'  σ  trace.steps σ'  r})" (is ?thesis2)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym[OF spec.singleton_le_extI])
    show "σ'  ?rhs" if "σ'  ?lhs" for σ'
      using that
      by (fastforce simp: spec.singleton.le_conv spec.singleton_le_conv
                    elim: trace.natural.less_eqE[where u="σ" and u'=σ, simplified]
                    dest: trace.stuttering.equiv.steps)
    show "?rhs  ?lhs"
      by (clarsimp simp: spec.singleton.le_conv spec.singleton.mono)
  qed
  then show ?thesis2
    by (rule inf_commute_conv)
qed

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma inf_rel:
  fixes F :: "('v × 'a × 's × 's) set"
  fixes r :: "('a, 's) steps"
  assumes "a. refl (r `` {a})"
  shows "spec.action F  spec.rel r = spec.action (F  UNIV × r)" (is "?lhs = ?rhs")
    and "spec.rel r  spec.action F = spec.action (F  UNIV × r)" (is ?thesis1)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym[OF spec.singleton_le_extI])
    from reflD[OF assms] show "σ  ?rhs" if "σ  ?lhs" for σ
      using that
      by (auto 0 2 simp: spec.singleton.le_conv spec.singleton.action_le_conv
                    split: option.split_asm)
    show "?rhs  ?lhs"
      by (rule order.trans[OF spec.action.inf_le inf.mono[OF order.refl spec.action.rel_le]]) simp
  qed
  then show ?thesis1
    by (rule inf_commute_conv)
qed

lemma inf_rel_reflcl:
  shows "spec.action F  spec.rel r = spec.action (F  UNIV × (r  UNIV × Id))"
    and "spec.rel r  spec.action F = spec.action (F  UNIV × (r  UNIV × Id))"
by (simp_all add: refl_on_def spec.rel.reflcl ac_simps flip: spec.action.inf_rel)

setup Sign.parent_path

setup Sign.mandatory_path "return"

lemma inf_rel:
  shows "spec.rel r  spec.return v = spec.return v"
    and "spec.return v  spec.rel r = spec.return v"
by (simp_all add: spec.return_def ac_simps spec.action.inf_rel_reflcl
                  Sigma_Un_distrib2 Int_Un_distrib Times_Int_Times
            flip: Sigma_Un_distrib2)

setup Sign.parent_path

setup Sign.mandatory_path "kleene.star"

lemma inf_rel:
  shows "spec.kleene.star P  spec.rel r = spec.kleene.star (P  spec.rel r)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (induct rule: spec.kleene.star.fixp_induct)
       (simp_all add: ac_simps inf_sup_distrib1 spec.bind.inf_rel le_supI1 le_supI2 spec.bind.mono)
  show "?rhs  ?lhs"
    by (induct rule: spec.kleene.star.fixp_induct)
       (simp_all add: ac_simps le_supI2 inf_sup_distrib1
                      spec.bind.inf_rel spec.bind.mono spec.return.inf_rel)
qed

setup Sign.parent_path

setup Sign.mandatory_path "steps"

lemma simps[simp]:
  shows "(a, s, s)  spec.steps P"
by (simp add: spec.steps_def exI[where x="UNIV × -Id"]
              spec.rel.minus_Id[where r=UNIV and A=UNIV, simplified] spec.rel.UNIV)

lemma member_conv:
  shows "x  spec.steps P  (σ. σ  P  x  trace.steps σ)"
by (meson spec.rel.galois spec.singleton.rel_le_conv spec.singleton_le_ext_conv subset_Compl_singleton)

setup Sign.mandatory_path "term"

lemma none:
  shows "spec.steps (spec.term.none P) = spec.steps P"
by (metis order.eq_iff spec.rel.galois spec.term.all.rel spec.term.galois)

lemma all:
  shows "spec.steps (spec.term.all P) = spec.steps P"
by (metis spec.steps.term.none spec.term.none_all)

setup Sign.parent_path

lemmas bot = spec.rel.lower_bot

lemmas monotone = spec.rel.monotone_lower
lemmas mono = monotoneD[OF spec.steps.monotone]

lemmas Sup = spec.rel.lower_Sup
lemmas sup = spec.rel.lower_sup
lemmas Inf_le = spec.rel.lower_Inf_le
lemmas inf_le = spec.rel.lower_inf_le

lemma singleton:
  shows "spec.steps σ = trace.steps σ"
by (meson subset_antisym order.refl spec.rel.galois spec.singleton.rel_le_conv)

lemma idle:
  shows "spec.steps spec.idle = {}"
by (simp add: spec.steps_def spec.idle.rel_le)

lemma action:
  shows "spec.steps (spec.action F) = snd ` F - UNIV × Id"
by (force simp: spec.action_def split_def
                spec.steps.Sup spec.steps.sup spec.steps.singleton spec.steps.idle)

lemma return:
  shows "spec.steps (spec.return v) = {}"
by (simp add: spec.return_def spec.steps.action)

lemma bind_le: ―‹ see spec.steps.bind›
  shows "spec.steps (f  g)  spec.steps f  (v. spec.steps (g v))"
by (force simp: spec.steps.member_conv spec.singleton.le_conv trace.split_all trace.steps'.append
         elim!: spec.singleton.bind_le)

lemma kleene_star:
  shows "spec.steps (spec.kleene.star P) = spec.steps P" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
  proof(induct rule: spec.kleene.star.fixp_induct[case_names adm bot step])
    case (step S) then show ?case
      by (simp add: spec.steps.sup spec.steps.return order.trans[OF spec.steps.bind_le])
  qed (simp_all add: spec.steps.bot)
  show "?rhs  ?lhs"
    by (simp add: spec.steps.mono spec.kleene.expansive_star)
qed

lemma map:
  shows "spec.steps (spec.map af sf vf P)
       = map_prod af (map_prod sf sf) ` spec.steps P - UNIV × Id"
by (rule spec.singleton.exhaust[of P])
   (force simp: spec.map.Sup spec.map.singleton spec.steps.Sup spec.steps.singleton trace.steps'.map image_Union)

lemma invmap_le:
  shows "spec.steps (spec.invmap af sf vf P)
        map_prod af (map_prod sf sf) -` (spec.steps (P  spec.rel (range af × range sf × range sf))  UNIV × Id) - UNIV × Id"
by (simp add: spec.rel.galois spec.rel.minus_Id
              order.trans[OF _ spec.invmap.mono[OF spec.rel.upper_lower_expansive]]
        flip: vimage_Un spec.invmap.rel[where vf=vf] spec.invmap.range)

setup Sign.mandatory_path "rel"

lemma monomorphic:
  fixes r :: "('a, 's) steps"
  shows "spec.steps (spec.rel.monomorphic r) = r - UNIV × Id" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: spec.rel.galois spec.rel.minus_Id flip: spec.rel.monomorphic_conv)
  show "?rhs  ?lhs"
    by (force simp: spec.rel.monomorphic.alt_def spec.term.all.Sup spec.term.all.singleton
                    spec.steps.Sup spec.steps.singleton
              dest: spec[where x="trace.T s [(a, s')] None" for a s s']
             split: if_splits)
qed

setup Sign.parent_path

lemma rel:
  fixes r :: "('a, 's) steps"
  shows "spec.steps (spec.rel r) = r - UNIV × Id"
by (simp add: spec.rel_def spec.steps.term.all spec.steps.rel.monomorphic)

lemma top:
  shows "spec.steps  = UNIV × - Id"
using spec.steps.rel[where r=UNIV] by (simp add: spec.rel.UNIV)

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Sequential assertions \label{sec:safety_logic-pre_post} ›

text‹

We specify sequential behavior with preconditions and postconditions.

›


subsubsection‹ Preconditions \label{sec:safety_logic-pre_post-pre} ›

setup Sign.mandatory_path "spec"

definition pre :: "'s pred  ('a, 's, 'v) spec" where
  "pre P = (spec.singleton ` {σ. P (trace.init σ)})"

setup Sign.mandatory_path "singleton"

lemma pre_le_conv[spec.singleton.le_conv]:
  shows "σ  spec.pre P  P (trace.init σ)"
by (auto simp add: spec.pre_def spec.singleton_le_conv trace.natural_def elim: trace.less_eqE)

lemma inf_pre:
  shows "spec.pre P  σ = (if P (trace.init σ) then σ else )" (is ?thesis1)
    and "σ  spec.pre P = (if P (trace.init σ) then σ else )" (is ?thesis2)
proof -
  show ?thesis1
    by (cases σ; rule spec.singleton.antisym)
       (auto simp: spec.singleton.le_conv spec.singleton_le_conv spec.singleton.not_bot trace.natural.trace_conv
            split: if_split_asm
             elim: trace.less_eqE)
  then show ?thesis2
    by (rule inf_commute_conv)
qed

setup Sign.parent_path

setup Sign.mandatory_path "idle"

lemma pre_le_conv[spec.idle_le]:
  shows "spec.idle  (spec.pre P :: ('a, 's, 'v) spec)  P = " (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (rule ccontr)
       (simp add: fun_eq_iff spec.pre_def spec.idle_def trace.split_Ex
                  spec.singleton_le_conv trace.less_eq_None trace.natural.simps)
  show "?rhs  ?lhs"
    by (rule spec.singleton_le_extI) (simp add: spec.singleton.le_conv)
qed

setup Sign.parent_path

setup Sign.mandatory_path "term"

setup Sign.mandatory_path "all"

lemma pre:
  shows "spec.term.all (spec.pre P) = spec.pre P"
by (rule spec.singleton.antisym; simp add: spec.singleton.le_conv)

setup Sign.parent_path

setup Sign.mandatory_path "closed"

lemma pre:
  shows "spec.pre P  spec.term.closed _"
by (rule spec.term.closed_upper[of "spec.pre P", simplified spec.term.all.pre])

lemma none_inf_pre:
  fixes P :: "'s pred"
  fixes Q :: "('a, 's, 'v) spec"
  shows "spec.term.none (Q  spec.pre P) = (spec.term.none Q  spec.pre P :: ('a, 's, 'w) spec)" (is "?lhs = ?rhs")
    and "spec.term.none (spec.pre P  Q) = (spec.pre P  spec.term.none Q :: ('a, 's, 'w) spec)" (is ?thesis2)
proof -
  show "?lhs = ?rhs"
    apply (subst spec.term.none_all[symmetric])
    apply (subst spec.term.all.inf)
    apply (subst spec.term.closed.none_inf_monomorphic(2)[symmetric])
    apply (simp_all add: spec.term.all.pre spec.term.closed.pre)
    done
  then show ?thesis2
    by (simp add: ac_simps)
qed

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "pre"

lemma bot[iff]:
  shows "spec.pre False = "
    and "spec.pre  = "
by (simp_all add: spec.pre_def)

lemma top[iff]:
  shows "spec.pre True = "
    and "spec.pre  = "
by (simp_all add: spec.pre_def full_SetCompr_eq spec.singleton.top)

lemma top_conv:
  shows "spec.pre P = ( :: ('a, 's, 'v) spec)  P = "
by (auto intro: iffD1[OF spec.idle.pre_le_conv[where 'a='a and 's='s and 'v='v]])

lemma K:
  shows "spec.pre P = (if P then  else )"
by (simp add: spec.pre_def full_SetCompr_eq spec.singleton.top)

lemma monotone:
  shows "mono spec.pre"
by (fastforce simp: spec.pre_def intro: monoI)

lemmas strengthen[strg] = st_monotone[OF spec.pre.monotone]
lemmas mono = monotoneD[OF spec.pre.monotone]

lemma SUP:
  shows "spec.pre (xX. P x) = (xX. spec.pre (P x))"
by (auto simp: spec.pre_def spec.eq_iff intro: rev_SUPI)

lemma Sup:
  shows "spec.pre (X) = (xX. spec.pre x)"
by (metis image_ident image_image spec.pre.SUP)

lemma Bex:
  shows "spec.pre (λs. xX. P x s) = (xX. spec.pre (P x))"
by (simp add: Sup_fun_def flip: spec.pre.SUP)

lemma Ex:
  shows "spec.pre (λs. x. P x s) = (x. spec.pre (P x))"
by (simp add: Sup_fun_def flip: spec.pre.SUP)

lemma
  shows disj: "spec.pre (P  Q) = spec.pre P  spec.pre Q"
    and sup: "spec.pre (P  Q) = spec.pre P  spec.pre Q"
using spec.pre.Sup[where X="{P, Q}"] by (simp_all add: sup_fun_def)

lemma INF:
  shows "spec.pre (xX. P x) = (xX. spec.pre (P x))"
by (auto simp: spec.eq_iff spec.singleton.pre_le_conv le_INF_iff intro: spec.singleton_le_extI)

lemma Inf:
  shows "spec.pre (X) = (xX. spec.pre x)"
by (metis image_ident image_image spec.pre.INF)

lemma Ball:
  shows "spec.pre (λs. xX. P x s) = (xX. spec.pre (P x))"
by (simp add: Inf_fun_def flip: spec.pre.INF)

lemma All:
  shows "spec.pre (λs. x. P x s) = (x. spec.pre (P x))"
by (simp add: Inf_fun_def flip: spec.pre.INF)

lemma inf:
  shows conj: "spec.pre (P  Q) = spec.pre P  spec.pre Q"
    and "spec.pre (P  Q) = spec.pre P  spec.pre Q"
using spec.pre.Inf[where X="{P, Q}"] by (simp_all add: inf_fun_def)

lemma inf_action_le: ―‹ Converse does not hold ›
  shows "spec.pre P  spec.action F  spec.action (UNIV × UNIV × Collect P × UNIV  F)" (is "?lhs  ?rhs")
    and "spec.action F  spec.pre P  spec.action (F  UNIV × UNIV × Collect P × UNIV)" (is ?thesis2)
proof -
  show "?lhs  ?rhs"
  proof(rule spec.singleton_le_extI)
    show "σ  ?rhs" if "σ  ?lhs" for σ
      using that[simplified, unfolded spec.singleton.action_le_conv spec.singleton.le_conv]
      by (cases σ;
          safe; clarsimp simp: trace.steps'.step_conv spec.action.idleI spec.action.stutter_stepsI
                        split: option.split_asm;
          subst spec.singleton.Cons; blast intro: spec.action.stepI)
  qed
  then show ?thesis2
    by (simp add: ac_simps)
qed

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemma pre:
  shows "spec.invmap af sf vf (spec.pre P) = spec.pre (λs. P (sf s))"
by (rule spec.singleton.antisym) (simp_all add: spec.singleton.le_conv)

setup Sign.parent_path

setup Sign.mandatory_path "bind"

lemma inf_pre:
  shows "spec.pre P  (f  g) = (spec.pre P  f)  g" (is "?lhs = ?rhs")
    and "(f  g)  spec.pre P = (f  spec.pre P)  g" (is ?thesis1)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym)
    show "?lhs  ?rhs"
      by (fastforce simp: spec.bind_def spec.continue_def inf_sup_distrib1 inf_Sup spec.singleton.inf_pre
               simp flip: spec.term.closed.none_inf_pre spec.singleton.pre_le_conv)
    show "?rhs  ?lhs"
      by (fastforce simp: spec.bind_def spec.continue_def inf_sup_distrib1 spec.singleton.pre_le_conv
               simp flip: spec.term.closed.none_inf_pre)
  qed
  then show ?thesis1
    by (simp add: ac_simps)
qed

setup Sign.parent_path

setup Sign.mandatory_path "steps"

lemma pre:
  assumes "P s0"
  shows "spec.steps (spec.pre P :: ('a, 's, 'v) spec) = UNIV × - Id"
proof -
  have "(a, s, s')  spec.steps (spec.pre P)" if "s  s'" for a :: 'a and s s'
    using assms that
    by (simp add: spec.singleton.le_conv spec.steps.member_conv trace.steps'.Cons_eq_if
                  exI[where x="trace.T s0 [(undefined, s), (a, s')] None"])
  then show ?thesis
    by auto
qed

setup Sign.parent_path

setup Sign.parent_path


subsubsection‹ Postconditions\label{sec:safety_logic-pre_post-post} ›

text‹

Unlike constspec.pre spec.post› can be expressed in terms of other constants.

›

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "post"

definition act :: "('v  's pred)  ('v × 'a × 's × 's) set" where
  "act Q = {(v, a, s, s') |v a s s'. Q v s'}"

setup Sign.mandatory_path "act"

lemma simps[simp]:
  shows "spec.post.act False = {}"
    and "spec.post.act  = {}"
    and "spec.post.act  = {}"
    and "spec.post.act True = UNIV"
    and "spec.post.act  = UNIV"
    and "spec.post.act  = UNIV"
    and "spec.post.act (Q  Q') = spec.post.act Q  spec.post.act Q'"
    and "spec.post.act (X) = (xX. spec.post.act x)"
    and "spec.post.act (λv. xY. R x v) = (xY. spec.post.act (R x))"
by (auto 0 2 simp: spec.post.act_def)

lemma monotone:
  shows "mono spec.post.act"
proof(rule monotoneI)
  show "spec.post.act Q  spec.post.act Q'" if "Q  Q'" for Q Q' :: "'v  's pred"
    using that unfolding spec.post.act_def by blast
qed

lemmas strengthen[strg] = st_monotone[OF spec.post.act.monotone]
lemmas mono = monotoneD[OF spec.post.act.monotone]

setup Sign.parent_path

setup Sign.parent_path

definition post :: "('v  's pred)  ('a, 's, 'v) spec" where
  "post Q =   (λ_::unit. spec.action (spec.post.act Q))"

setup Sign.mandatory_path "singleton"

lemma post_le_conv[spec.singleton.le_conv]:
  fixes Q :: "'v  's pred"
  shows "σ  spec.post Q
     (case trace.term σ of None  True | Some v  Q v (trace.final σ))" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (fastforce simp: spec.post_def spec.post.act_def
                        spec.singleton.le_conv spec.singleton.action_le_conv trace.steps'.step_conv
                 split: option.split
                  elim: spec.singleton.bind_le)
  show "?rhs  ?lhs"
    by (cases σ)
       (simp add: spec.post_def spec.post.act_def spec.action.stutterI
                  spec.bind.incompleteI spec.bind.continueI[where ys="[]", simplified]
           split: option.splits)
qed

setup Sign.parent_path

setup Sign.mandatory_path "idle"

lemma post_le[spec.idle_le]:
  shows "spec.idle  spec.post Q"
by (rule spec.singleton_le_extI) (simp add: spec.singleton.le_conv)

setup Sign.parent_path

setup Sign.mandatory_path "term"

setup Sign.mandatory_path "none"

lemma post_le:
  shows "spec.term.none P  spec.post Q"
by (rule spec.singleton_le_extI) (simp add: spec.singleton.le_conv)

lemma post:
  shows "spec.term.none (spec.post Q :: ('a, 's, 'v) spec)
       = spec.term.none ( :: ('a, 's, unit) spec)"
by (metis spec.eq_iff spec.term.galois spec.term.none.post_le spec.term.none_all top_greatest)

setup Sign.parent_path

setup Sign.mandatory_path "all"

lemma post:
  shows "spec.term.all (spec.post Q) = "
by (metis spec.term.all_none spec.term.none.post spec.term.upper_top)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.mandatory_path "post"

lemma bot[iff]:
  shows "spec.post False = spec.term.none ( :: (_, _, unit) spec)"
    and "spec.post  = spec.term.none ( :: (_, _, unit) spec)"
    and "spec.post  = spec.term.none ( :: (_, _, unit) spec)"
by (simp_all add: spec.post_def spec.action.empty spec.bind.idleR spec.bind.botR)

lemma monotone:
  shows "mono spec.post"
by (simp add: spec.post_def monoI spec.action.mono spec.bind.mono spec.post.act.mono)

lemmas strengthen[strg] = st_monotone[OF spec.post.monotone]
lemmas mono = monotoneD[OF spec.post.monotone]

lemma SUP_not_empty:
  fixes X :: "'a set"
  fixes Q :: "'a  'v  's pred"
  assumes "X  {}"
  shows "spec.post (λv. xX. Q x v) = (xX. spec.post (Q x))"
by (simp add: assms spec.post_def flip: spec.bind.SUPR_not_empty spec.action.SUP_not_empty)

lemma disj:
  shows "spec.post (Q  Q') = spec.post Q  spec.post Q'"
    and "spec.post (λrv. Q rv  Q' rv) = spec.post Q  spec.post Q'"
    and "spec.post (λrv. Q rv  Q' rv) = spec.post Q  spec.post Q'"
using spec.post.SUP_not_empty[where X="UNIV" and Q="λx. if x then Q' else Q"]
by (simp_all add: UNIV_bool sup_fun_def)

lemma INF:
  shows "spec.post (xX. Q x) = (xX. spec.post (Q x))"
    and "spec.post (λv. xX. Q x v) = (xX. spec.post (Q x))"
    and "spec.post (λv s. xX. Q x v s) = (xX. spec.post (Q x))"
by (fastforce intro: antisym spec.singleton_le_extI simp: spec.singleton.le_conv le_Inf_iff split: option.split)+

lemma Inf:
  shows "spec.post (X) = (xX. spec.post x)"
by (metis image_ident image_image spec.post.INF(1))

lemma Ball:
  shows "spec.post (λv s. xX. Q x v s) = (xX. spec.post (Q x))"
by (simp add: Inf_fun_def flip: spec.post.INF)

lemma All:
  shows "spec.post (λv s. x. Q x v s) = (x. spec.post (Q x))"
by (simp add: Inf_fun_def flip: spec.post.INF)

lemma inf:
  shows "spec.post (Q  Q') = spec.post Q  spec.post Q'"
    and "spec.post (λrv. Q rv  Q' rv) = spec.post Q  spec.post Q'"
    and conj: "spec.post (λrv. Q rv  Q' rv) = spec.post Q  spec.post Q'"
by (simp_all add: inf_fun_def flip: spec.post.INF[where X="UNIV" and Q="λx. if x then Q' else Q", simplified UNIV_bool, simplified])

lemma top[iff]:
  shows "spec.post True = "
    and "spec.post  = "
    and "spec.post  = "
by (simp_all add: top_fun_def flip: spec.post.INF[where X="{}", simplified])

lemma top_conv:
  shows "spec.post Q = ( :: ('a, 's, 'v) spec)  Q = "
by (fastforce simp: spec.singleton.post_le_conv
              dest: arg_cong[where f="λx. σ. σ  x"] spec[where x="trace.T s [] (Some v)" for s v])

lemma K:
  shows "spec.post (λ_ _. Q) = (if Q then  else   (λ_::unit. ))"
by (auto  simp flip: spec.bind.botR bot_fun_def)

setup Sign.parent_path

lemma bind_post_pre:
  shows "f  spec.post Q  g = f  (λv. g v  spec.pre (Q v))" (is "?lhs = ?rhs")
    and "spec.post Q  f  g = f  (λv. spec.pre (Q v)  g v)" (is ?thesis1)
proof -
  show "?lhs = ?rhs"
  proof(rule spec.singleton.antisym)
    show "σ  ?rhs" if "σ  ?lhs" for σ
      using that
      by (induct rule: spec.singleton.bind_le)
         (cases σ; force simp: trace.split_all spec.singleton.le_conv
                        intro: spec.bind.incompleteI spec.bind.continueI)+
    show "σ  ?lhs" if "σ  ?rhs" for σ
      using that
      by (induct rule: spec.singleton.bind_le)
         (cases σ; force simp: trace.split_all spec.singleton.le_conv
                        intro: spec.bind.incompleteI spec.bind.continueI)+
qed
  then show ?thesis1
    by (simp add: ac_simps)
qed

setup Sign.mandatory_path "invmap"

lemma post:
  shows "spec.invmap af sf vf (spec.post Q) = spec.post (λv s. Q (vf v) (sf s))"
by (rule antisym[OF spec.singleton_le_extI spec.singleton_le_extI])
   (simp_all add: spec.singleton.invmap_le_conv spec.singleton.post_le_conv trace.final'.map split: option.split_asm)

setup Sign.parent_path

setup Sign.mandatory_path "action"

lemma post_le_conv:
  shows "spec.action F  spec.post Q  (v a s s'. (v, a, s, s')  F  Q v s')"
by (fastforce simp: spec.action_def split_def spec.singleton.le_conv spec.idle.post_le)

setup Sign.parent_path

setup Sign.mandatory_path "bind"

lemma post_le:
  assumes "v. g v  spec.post Q"
  shows "f  g  spec.post Q"
by (strengthen ord_to_strengthen(1)[OF assms])
   (simp add: spec.post_def spec.bind.mono flip: spec.bind.bind)

lemma inf_post:
  shows "(f  g)  spec.post Q = f  (λv. g v  spec.post Q)" (is "?lhs = ?rhs")
    and "spec.post Q  (f  g) = f  (λv. spec.post Q  g v)" (is ?thesis2)
proof -
  show "?lhs = ?rhs"
  proof(rule antisym[OF spec.singleton_le_extI])
    fix σ
    assume lhs: "σ  ?lhs"
    from lhs[simplified, THEN conjunct1] lhs[simplified, THEN conjunct2] show "σ  ?rhs"
    proof(induct rule: spec.singleton.bind_le)
      case (continue σf σg vf) then show ?case
        by (cases σf)
           (force intro: spec.bind.continueI simp: spec.singleton.le_conv split: option.split)
    qed (simp add: spec.singleton.bind_le_conv)
  qed (simp add: spec.bind.mono spec.bind.post_le)
  then show ?thesis2
    by (simp add: ac_simps)
qed

lemma mono_stronger:
  assumes f: "f  f'  spec.post Q"
  assumes g: "v. g v  spec.pre (Q v)  g' v"
  shows "spec.bind f g  spec.bind f' g'"
by (strengthen ord_to_strengthen(1)[OF f]) (simp add: g spec.bind.mono spec.bind_post_pre)

setup Sign.parent_path

setup Sign.parent_path


subsubsection‹ Strongest postconditions\label{sec:safety_logic-strongest_post} ›

setup Sign.mandatory_path "spec.post"

definition strongest :: "('a, 's, 'v) spec  'v  's pred" where
  "strongest P = {Q. P  spec.post Q}"

interpretation strongest: galois.complete_lattice_class "spec.post.strongest" "spec.post"
by (simp add: spec.post.strongest_def galois.upper_preserves_InfI spec.post.Inf spec.post.monotone)

lemma strongest_alt_def:
  shows "spec.post.strongest P = (λv s. σ. σ  P  trace.term σ = Some v  trace.final σ = s)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (rule spec.singleton.exhaust[of P])
       (fastforce simp: spec.post.strongest_def spec.singleton.le_conv
                  dest: spec[where x="λv s. σX. trace.term σ = Some v  trace.final σ = s" for X]
                 split: option.split)
  show "?rhs  ?lhs"
    by (fastforce simp: spec.post.strongest_def spec.singleton.le_conv dest: order.trans)
qed

setup Sign.mandatory_path "strongest"

lemma singleton:
  shows "spec.post.strongest σ
      = (λv s. case trace.term σ of None  False | Some v'  v' = v  trace.final σ = s)"
by (auto simp: spec.post.strongest_alt_def fun_eq_iff trace.split_all
         cong: trace.final'.natural'_cong
        split: option.split)

lemmas monotone = spec.post.strongest.monotone_lower
lemmas mono = monoD[OF spec.post.strongest.monotone]
lemmas Sup = spec.post.strongest.lower_Sup
lemmas sup = spec.post.strongest.lower_sup

lemma top[iff]:
  shows "spec.post.strongest  = "
by (simp add: spec.post.strongest_def spec.post.top_conv top.extremum_unique)

lemma action:
  shows "spec.post.strongest (spec.action F) = (λv s'. a s. (v, a, s, s')  F)"
by (simp add: spec.post.strongest_def spec.action.post_le_conv) fast

lemma return:
  shows "spec.post.strongest (spec.return v) = (λv' s. v' = v)"
by (simp add: spec.return_def spec.post.strongest.action)

setup Sign.mandatory_path "term"

lemma none:
  shows "spec.post.strongest (spec.term.none P) = "
by (clarsimp simp: spec.term.none_def spec.post.strongest.Sup spec.post.strongest.singleton fun_eq_iff)

lemma all:
  assumes "spec.idle  P"
  shows "spec.post.strongest (spec.term.all P) = "
by (rule top_le[OF order.trans[OF _ spec.post.strongest.mono[OF spec.term.all.mono[OF assms]]]])
   (simp add: spec.term.all.idle spec.post.strongest.Sup spec.post.strongest.return Sup_fun_def top_fun_def)

lemma closed:
  assumes "spec.idle  P"
  assumes "P  spec.term.closed _"
  shows "spec.post.strongest P = "
by (metis spec.post.strongest.term.all[OF assms(1)] spec.term.all.closed_conv[OF assms(2)])

setup Sign.parent_path

lemma bind:
  shows "spec.post.strongest (f  g)
       = spec.post.strongest (v. spec.pre (spec.post.strongest f v)  g v)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (auto 0 4 simp: spec.post.strongest_alt_def spec.singleton.le_conv
                elim!: spec.singleton.bind_le)
  show "?rhs  ?lhs"
    by (force simp: spec.post.strongest_alt_def spec.singleton.le_conv trace.split_all
              dest: spec.bind.continueI)
qed

lemma rel:
  shows "spec.post.strongest (spec.rel r) = "
by (simp add: spec.rel_def spec.post.strongest.term.all spec.idle.kleene.star_le)

lemma pre:
  shows "spec.post.strongest (spec.pre P) = (λv s'. s. P s)"
by (auto simp: spec.post.strongest_alt_def spec.singleton.le_conv trace.split_Ex fun_eq_iff
       intro!: exI[where x="[(undefined, s)]" for s])

lemma post:
  shows "spec.post.strongest (spec.post Q) = Q"
by (auto simp: spec.post.strongest_alt_def spec.singleton.le_conv trace.split_Ex fun_eq_iff
       intro!: exI[where x="[]"])

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Initial steps\label{sec:safety_logic-initial_steps} ›

text‹

The initial transition of a process.

›

setup Sign.mandatory_path "spec"

definition initial_steps :: "('a, 's, 'v) spec  ('a, 's) steps" where
  "initial_steps P = {(a, s, s'). s, [(a, s')], None  P}"

setup Sign.mandatory_path "initial_steps"

lemma steps_le:
  shows "spec.initial_steps P  spec.steps P  UNIV × Id"
by (fastforce simp: spec.initial_steps_def spec.steps.member_conv split: if_splits)

lemma galois:
  shows "r  spec.initial_steps P  spec.idle  P  spec.action ({()} × r)    P" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (auto simp: spec.action_def spec.initial_steps_def spec.bind.SupL spec.bind.supL
                   spec.bind.singletonL spec.singleton.not_bot spec.term.none.singleton)
  show "?rhs  ?lhs"
    by (auto simp: spec.initial_steps_def spec.idle.action_le spec.idle.bind_le_conv
            elim!: order.trans[rotated]
            intro: spec.action.stepI spec.bind.incompleteI)
qed

lemma bot:
  shows "spec.initial_steps  = {}"
by (auto simp: spec.initial_steps_def spec.singleton.not_bot)

lemma top:
  shows "spec.initial_steps  = UNIV"
by (auto simp: spec.initial_steps_def spec.singleton.not_bot)

lemma monotone:
  shows "mono spec.initial_steps"
by (force intro: monoI simp: spec.initial_steps_def)

lemmas strengthen[strg] = st_monotone[OF spec.initial_steps.monotone]
lemmas mono = monotoneD[OF spec.initial_steps.monotone]

lemma Sup:
  shows "spec.initial_steps (X) = (spec.initial_steps ` X)"
by (auto simp: spec.initial_steps_def)

lemma Inf:
  shows "spec.initial_steps (X) = (spec.initial_steps ` X)"
by (auto simp: spec.initial_steps_def le_Inf_iff)

lemma idle:
  shows "spec.initial_steps spec.idle = UNIV × Id"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

lemma action:
  shows "spec.initial_steps (spec.action F) = snd ` F  UNIV × Id"
by (auto simp: spec.initial_steps_def spec.singleton.action_le_conv trace.steps'.step_conv)

lemma return:
  shows "spec.initial_steps (spec.return v) = UNIV × Id"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

lemma bind:
  shows "spec.initial_steps (f  g)
       = spec.initial_steps f
          spec.initial_steps (v. spec.pre (spec.post.strongest (f  spec.return v) v)  g v)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (fastforce simp: spec.initial_steps_def spec.post.strongest_alt_def spec.singleton.le_conv
                        trace.split_all Cons_eq_append_conv trace.natural.simps
                 elim!: spec.singleton.bind_le)
  show "?rhs  ?lhs"
    by (auto simp: spec.initial_steps_def spec.post.strongest_alt_def spec.singleton.le_conv
                   trace.split_all spec.bind.incompleteI order.trans[OF _ spec.bind.continueI, rotated])
qed

lemma rel:
  shows "spec.initial_steps (spec.rel r) = r  UNIV × Id"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

lemma pre:
  shows "spec.initial_steps (spec.pre P) = UNIV × Pre P"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

lemma post:
  shows "spec.initial_steps (spec.post Q) = UNIV"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

setup Sign.mandatory_path "term"

lemma none:
  shows "spec.initial_steps (spec.term.none P) = spec.initial_steps P"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv)

lemma all:
  shows "spec.initial_steps (spec.term.all P) = spec.initial_steps P"
by (auto simp: spec.initial_steps_def spec.singleton.le_conv order.trans[rotated]
               spec.singleton.mono trace.less_eq_None)

setup Sign.parent_path

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Heyting implication\label{sec:safety_logic-heyting} ›

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "singleton"

lemma heyting_le_conv:
  shows "σ  P H Q  (σ'σ. σ'  P  σ'  Q)" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (meson order.trans heyting.mp spec.singleton.mono)
  show "?rhs  ?lhs"
    by (rule spec.singleton.exhaust[of P])
       (clarsimp simp: heyting inf_Sup spec.singleton.inf trace.less_eq_take_def spec.singleton_le_conv;
        metis spec.singleton.simps(1) trace.take.naturalE(2))
qed

setup Sign.parent_path

text‹

Connect the generic definition of Heyting implication to a concrete one in the model.

›

lift_definition heyting :: "('a, 's, 'v) spec  ('a, 's, 'v) spec  ('a, 's, 'v) spec" is
  downwards.imp
by (rule raw.spec.closed.downwards_imp)

lemma heyting_alt_def:
  shows "(H) = (spec.heyting :: __('a, 's, 'v) spec)"
proof -
  have "P  spec.heyting Q R  P  Q  R" for P Q R :: "('a, 's, 'v) spec"
    by transfer (simp add: raw.spec.closed.heyting_downwards_imp)
  with heyting show ?thesis by (intro fun_eqI antisym; fast)
qed

declare spec.heyting.transfer[transfer_rule del]

setup Sign.mandatory_path "heyting"

lemma transfer_alt[transfer_rule]:
  shows "rel_fun (pcr_spec (=) (=) (=)) (rel_fun (pcr_spec (=) (=) (=)) (pcr_spec (=) (=) (=))) downwards.imp (H)"
by (simp add: spec.heyting.transfer spec.heyting_alt_def)

text‹

An example due to citet‹p504› in "AbadiMerz:1995"
where the (TLA) model validates a theorem that is not intuitionistically valid.
This is ``some kind of linearity'' and intuitively encodes disjunction elimination.

›

lemma linearity:
  fixes Q :: "(_, _, _) spec"
  shows "((P H Q) H R)  ((Q H P) H R)  R"
by transfer
   (clarsimp simp: downwards.imp_def;
    meson downwards.closed_in[OF _ _ raw.spec.closed.downwards_closed] trace.less_eq_same_cases order.refl)

lemma SupR:
  fixes P :: "(_, _, _) spec"
  assumes "X  {}"
  shows "P H (xX. Q x) = (xX. P H Q x)" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI heyting.SUPR_le])
  show "σ  ?rhs" if "σ  ?lhs" for σ
  proof(cases "σ  P")
    case True with σ  ?lhs show ?thesis by (simp add: heyting inf.absorb1)
  next
    case False show ?thesis
    proof(cases "trace.init σ, [], None  P")
      case True with ¬ σ  P
      obtain j where "ij. trace.take i σ  P"
                 and "¬ trace.take (Suc j) σ  P"
        using ex_least_nat_less[where P="λi. ¬trace.take i σ  P" and n="Suc (length (trace.rest σ))"]
        by (clarsimp simp: less_Suc_eq_le simp flip: trace.take.Ex_all)
      with σ  ?lhs show ?thesis
        by (clarsimp simp: spec.singleton.heyting_le_conv dest!: spec[where x="trace.take j σ"])
           (metis not_less_eq_eq order_refl spec.singleton.mono spec.singleton_le_ext_conv
                  trace.less_eq_takeE trace.take.mono)
    next
      case False with X  {} σ  ?lhs show ?thesis
        by (clarsimp simp: spec.singleton.heyting_le_conv simp flip: ex_in_conv)
           (metis "trace.take.0" spec.singleton.takeI trace.less_eq_take_def trace.take.sel(1))
    qed
  qed
qed

lemma cont:
  fixes P :: "(_, _, _) spec"
  shows "cont Sup (≤) Sup (≤) ((H) P)"
by (rule contI) (simp add: spec.heyting.SupR[where Q=id, simplified])

lemma mcont:
  fixes P :: "(_, _, _) spec"
  shows "mcont Sup (≤) Sup (≤) ((H) P)"
by (simp add: mcontI[OF _ spec.heyting.cont])

lemmas mcont2mcont[cont_intro] = mcont2mcont[OF spec.heyting.mcont, of luba orda Q P] for luba orda Q P

lemma non_triv:
  shows "P H   P  spec.idle  P" (is "?lhs  ?rhs")
proof(rule iffI)
  show "?lhs  ?rhs"
    by (rule spec.singleton.exhaust[of P])
       (fastforce dest: spec[where x="x, [], None" for x]
                  simp: spec.idle_def heyting_def heyting.inf_Sup_distrib trace.split_all
                        spec.singleton.inf spec.singleton_le_conv trace.less_eq_None trace.natural.simps)
  have "spec.idle H   spec.idle"
    by (fastforce intro: spec.singleton_le_extI
                   dest: spec[where x="trace.T (trace.init σ) [] None" for σ]
                   simp: spec.singleton.le_conv spec.singleton.heyting_le_conv spec.singleton.not_bot trace.less_eq_None)
  then show ?lhs if ?rhs
    by - (strengthen ord_to_strengthen(2)[OF ?rhs])
qed

lemma post:
  shows "spec.post Q H spec.post Q' = spec.post (λv s. Q v s  Q' v s)" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (auto intro: spec.singleton_le_extI simp: spec.singleton.heyting_le_conv spec.singleton.le_conv split: option.splits)
  show "?rhs  ?lhs"
    by (auto simp add: heyting simp flip: spec.post.conj intro: spec.post.mono)
qed

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemma heyting:
  shows "spec.invmap af sf vf (P H Q) = spec.invmap af sf vf P H spec.invmap af sf vf Q" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs" by (simp add: heyting heyting.detachment spec.invmap.mono flip: spec.invmap.inf)
  show "?rhs  ?lhs"
    by (simp add: heyting heyting.detachment spec.map.inf_distr flip: spec.map_invmap.galois spec.invmap.inf)
       (simp add: spec.invmap.mono spec.map_invmap.galois)
qed

setup Sign.parent_path

setup Sign.mandatory_path "term"

lemma heyting_noneL_allR_mono:
  fixes P :: "(_, _, 'v) spec"
  fixes Q :: "(_, _, 'v) spec"
  shows "spec.term.none P H Q = P H spec.term.all Q" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: heyting spec.term.none.inf flip: spec.term.galois) (simp add: heyting.uncurry)
  show "?rhs  ?lhs"
    by (simp add: heyting heyting.discharge spec.term.closed.none_inf_monomorphic spec.term.galois)
qed

setup Sign.mandatory_path "all"

lemma heyting: ―‹ polymorphic constspec.term.all
  fixes P :: "(_, _, 'v) spec"
  fixes Q :: "(_, _, 'v) spec"
  shows "(spec.term.all (P H Q) :: (_, _, 'w) spec)
       = spec.term.none P H spec.term.all Q" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (simp add: heyting spec.term.none.inf flip: spec.term.galois)
       (metis heyting.detachment(2) le_inf_iff spec.term.none.contractive spec.term.none.inf(2))
  have "spec.term.none (spec.term.none P H spec.term.all Q :: (_, _, 'w) spec)
           spec.term.none (spec.term.none P :: (_, _, 'w) spec)
      Q"
    by (metis heyting.detachment(2) inf_sup_ord(2) spec.term.galois spec.term.none.inf(2))
  then show "?rhs  ?lhs"
    by (simp add: heyting flip: spec.term.galois)
       (metis spec.term.cl_def spec.term.all.monomorphic spec.term.none_all
              heyting.detachment(2) spec.term.heyting_noneL_allR_mono)
qed

setup Sign.parent_path

setup Sign.mandatory_path "none"

lemma heyting_le:
  shows "spec.term.none (P H Q)  spec.term.all P H spec.term.none Q"
by (simp add: spec.term.galois spec.term.all.heyting heyting.mono spec.term.all.expansive)

setup Sign.parent_path

setup Sign.mandatory_path "closed"

lemma heyting:
  assumes "Q  spec.term.closed _"
  shows "P H Q  spec.term.closed _"
by (rule spec.term.closed_clI)
   (simp add: spec.term.all.heyting spec.term.heyting_noneL_allR_mono spec.term.all.monomorphic
        flip: spec.term.all.closed_conv[OF assms])

setup Sign.parent_path

setup Sign.parent_path

setup Sign.parent_path


subsection‹ Miscellaneous algebra ›

setup Sign.mandatory_path "spec"

setup Sign.mandatory_path "steps"

lemma bind:
  shows "spec.steps (f  g)
       = spec.steps f  (v. spec.steps (spec.pre (spec.post.strongest f v)  g v))" (is "?lhs = ?rhs")
proof(rule antisym)
 show "?lhs  ?rhs"
  unfolding spec.rel.galois
  by (rule spec.singleton_le_extI)
     (fastforce elim!: spec.singleton.bind_le
                 simp: spec.singleton.le_conv spec.steps.member_conv trace.steps'.append
                       spec.post.strongest_alt_def)
  show "?rhs  ?lhs"
    by (fastforce simp: spec.post.strongest_alt_def spec.bind_def spec.continue_def
                        spec.steps.term.none spec.steps.Sup spec.steps.sup spec.steps.singleton
                        spec.steps.member_conv spec.singleton.le_conv trace.split_Ex trace.steps'.append)
qed

setup Sign.parent_path

setup Sign.mandatory_path "map"

lemma idle:
  shows "spec.map af sf vf spec.idle = spec.pre (λs. s  range sf)  spec.idle" (is "?lhs = ?rhs")
proof(rule antisym)
  show "?lhs  ?rhs"
    by (auto simp: spec.idle_def spec.map.Sup spec.map.singleton spec.singleton.pre_le_conv)
  show "?rhs  ?lhs"
    by (auto simp: spec.idle_def spec.pre_def trace.split_all image_image inf_Sup Sup_inf
                   spec.map.Sup spec.map.singleton spec.singleton.inf
            elim!: trace.less_eqE)
qed

lemma return:
  fixes F :: "('v × 'a × 's × 's) set"
  shows "spec.map af sf vf (spec.return v)
       = spec.pre (λs. s  range sf)  spec.return (vf v)" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (force simp: spec.return_def spec.action_def spec.idle_def
                    spec.map.Sup spec.map.sup spec.map.singleton spec.singleton.pre_le_conv)
  fix σ
  assume "σ  ?rhs"
  then obtain s where "trace.init σ = sf s" by (clarsimp simp: spec.singleton.le_conv)
  with σ  ?rhs show "σ  ?lhs"
    by (simp add: spec.singleton.le_conv exI[where x="trace.T s [] (Some v)"]
                  spec.singleton_le_conv trace.natural_def trace.less_eq_None
            flip: trace.natural'.eq_Nil_conv
           split: option.split_asm)
qed

lemma kleene_star_le:
  fixes P :: "('a, 's, unit) spec"
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "unit  unit"
  shows "spec.map af sf vf (spec.kleene.star P)  spec.kleene.star (spec.map af sf vf P)" (is "_  ?rhs")
proof(induct rule: spec.kleene.star.fixp_induct[where P="λR. spec.map af sf vf (R P)  ?rhs", case_names adm bot step])
  case (step R) show ?case
    apply (simp add: spec.map.sup spec.map.return order.trans[OF _ spec.kleene.epsilon_star_le])
    apply (subst spec.kleene.star.simps)
    apply (strengthen ord_to_strengthen(1)[OF spec.map.bind_le])
    apply (strengthen ord_to_strengthen(1)[OF step])
    apply (simp add: fun_unit_id[where f=vf])
    done
qed (simp_all add: spec.map.bot)

lemma rel_le:
  shows "spec.map af sf vf (spec.rel r)  spec.rel (map_prod af (map_prod sf sf) ` r)"
apply (simp add: spec.rel_def spec.term.none.map flip: spec.term.galois)
apply (simp add: spec.rel.act_def flip: spec.term.none.map[where vf=id])
apply (strengthen ord_to_strengthen(1)[OF spec.map.kleene_star_le])
apply (strengthen ord_to_strengthen(1)[OF spec.map.action_le])
apply (strengthen ord_to_strengthen(1)[OF spec.term.none.contractive])
apply (auto intro!: monotoneD[OF spec.kleene.monotone_star] spec.action.mono)
done

text‹

General lemmas for constspec.map are elusive. We relate it to constspec.rel, constspec.pre
and constspec.post under a somewhat weak constraint. Intuitively we ask that, for
distinct representations (s0 and s0'›) of an abstract state (sf s0 where sf s0' = sf s0),
if agent a› can evolve s0 to s1 according to r› ((a, s0, s1) ∈ r›) then there is an agent
a'› where af a' = af a› that can evolve s0'› to an s1'› which represents the same abstract
state (sf s1' = sf s1).

All injective sf› satisfy this condition.

›

context
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
begin

context
  fixes r :: "('a, 's) steps"
  assumes step_cong: "a s0 s1 s0'. (a, s0, s1)  r  sf s1  sf s0  sf s0' = sf s0
                                 (a' s1'. af a' = af a  sf s1' = sf s1  (a', s0', s1')  r)"
begin

private lemma map_relE[consumes 1]:
  fixes xs :: "('b × 't) list"
  assumes "trace.steps' s xs  map_prod af (map_prod sf sf) ` r"
  obtains (Idle) "snd ` set xs  {s}"
        | (Step) s' xs'
          where "sf s' = s"
           and "trace.natural' s xs = map (map_prod af sf) xs'"
           and "trace.steps' s' xs'  r"
using assms
proof(atomize_elim, induct xs rule: prefix_induct)
  case (snoc xs x) show ?case
  proof(cases "snd x = trace.final' s xs")
    case True with snoc(2,3) show ?thesis
      by (fastforce simp: trace.steps'.append trace.natural'.append)
  next
    case False
    with snoc(2,3) consider
        (idle) "snd ` set xs  {s}"
      | (step) s' xs'
         where "sf s' = s"
           and "trace.natural' s xs = map (map_prod af sf) xs'"
           and "trace.steps' s' xs'  r"
      by (auto 0 0 simp: trace.steps'.append)
    then show ?thesis
    proof cases
      case idle with snoc(3) show ?thesis
        by (cases x)
           (clarsimp simp: trace.steps'.append trace.natural'.append Cons_eq_map_conv
                simp flip: trace.natural'.eq_Nil_conv ex_simps
                    split: if_splits;
            metis)
    next
      case (step s xs') with False snoc(3) step_cong show ?thesis
        by (cases x)
           (clarsimp simp: trace.steps'.append trace.natural'.append append_eq_map_conv Cons_eq_map_conv
                simp flip: ex_simps
                   intro!: exI[where x=s] exI[where x=xs'];
            metis trace.final'.map trace.final'.natural')
    qed
  qed
qed simp

lemma rel:
  shows "spec.map af sf vf (spec.rel r)
       = spec.rel (map_prod af (map_prod sf sf) ` r)
         spec.pre (λs. s  range sf)
         spec.post (λv s. v  range vf)" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI spec.singleton_le_extI])
  show "σ  ?rhs" if "σ  ?lhs" for σ
  proof(intro le_infI)
    from that show "σ  spec.rel (map_prod af (map_prod sf sf) ` r)"
      by (force simp: spec.singleton.le_conv spec.steps.singleton trace.steps'.map
                dest: spec.steps.mono)
    from that show "σ  spec.pre (λs. s  range sf)"
      by (fastforce simp: spec.singleton.le_conv spec.singleton_le_conv trace.natural_def
                    elim: trace.less_eqE)
    from that show "σ  spec.post (λv s. v  range vf)"
      by (cases σ) (clarsimp simp: spec.singleton.le_conv split: option.split)
  qed
  show "σ  ?lhs" if "σ  ?rhs" for σ
    using that[simplified, simplified spec.singleton.le_conv, THEN conjunct1]
          that[simplified, simplified spec.singleton.le_conv, THEN conjunct2]
  proof(induct rule: map_relE)
    case Idle then show ?case
      by (cases σ)
         (clarsimp simp: spec.singleton.le_conv;
          force simp: trace.natural.idle trace.natural.simps f_inv_into_f order_le_less
               split: option.split_asm
              intro!: exI[where x="trace.T s [] (map_option (inv vf) (trace.term σ))" for s])
  next
    case (Step s xs)
    from Step(1,3,4) Step(2)[symmetric] show ?case
      by (cases σ)
         (clarsimp simp: spec.singleton.le_conv f_inv_into_f[OF rangeI] trace.natural'.natural'
                         exI[where x="trace.T s xs (map_option (inv vf) (trace.term σ))"]
                  split: option.split_asm)
  qed
qed

lemma pre:
  fixes P :: "'t pred"
  shows "spec.map af sf vf (spec.pre (λs. P (sf s)))
       = spec.pre (λs. P s  s  range sf)  spec.post (λv s. s  range sf  v  range vf)
           spec.rel (range af × range sf × range sf)" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (simp add: spec.map_invmap.galois spec.invmap.pre spec.invmap.post spec.invmap.rel
                  map_prod_vimage_Times vimage_range spec.rel.UNIV)
  fix σ
  assume "σ  ?rhs"
  then obtain s xs
    where "P (sf s)"
      and "trace.init σ = sf s"
      and "case trace.term σ of None  True
               | Some v  trace.final' (trace.init σ) (trace.rest σ)  range sf  v  range vf"
      and "map (map_prod af sf) xs = trace.natural' (sf s) (trace.rest σ)"
    by (clarsimp simp: spec.singleton.le_conv trace.steps'.map_range_conv)
  then show "σ  ?lhs"
    by (cases σ)
       (fastforce intro: exI[where x="trace.T s xs (Some (inv vf (the (trace.term σ))))"]
                         range_eqI[where x="trace.final' s xs"]
                   dest: arg_cong[where f="trace.final' (sf s)"]
                   simp: spec.singleton.le_conv trace.final'.map f_inv_into_f trace.natural'.natural'
                         order.trans[OF spec.singleton.less_eq_None spec.singleton.simps(2)]
                  split: option.split_asm)
qed

lemma post:
  fixes Q :: "'w  't pred"
  shows "spec.map af sf vf (spec.post (λv s. Q (vf v) (sf s)))
       = spec.pre (λs. s  range sf)  spec.post (λv s. s  range sf  Q v s  v  range vf)
           spec.rel (range af × range sf × range sf)" (is "?lhs = ?rhs")
proof(rule antisym[OF _ spec.singleton_le_extI])
  show "?lhs  ?rhs"
    by (simp add: spec.map_invmap.galois spec.invmap.pre spec.invmap.post spec.invmap.rel
                  map_prod_vimage_Times vimage_range spec.rel.UNIV)
  fix σ
  assume "σ  ?rhs"
  then obtain s xs
    where "trace.init σ = sf s"
      and "case trace.term σ of None  True | Some v  trace.final' (trace.init σ) (trace.rest σ)  range sf  Q v (trace.final' (trace.init σ) (trace.rest σ))  v  range vf"
      and "map (map_prod af sf) xs = trace.natural' (sf s) (trace.rest σ)"
    by (clarsimp simp: spec.singleton.le_conv trace.steps'.map_range_conv)
  then show "σ  ?lhs"
    by (cases σ)
       (clarsimp simp: spec.singleton.le_conv trace.natural'.natural'
               intro!: exI[where x="trace.T s xs (map_option (inv vf) (trace.term σ))"]
                split: option.split_asm;
        clarsimp dest!: arg_cong[where f="trace.final' (sf s)"] simp: trace.final'.map;
        metis f_inv_into_f rangeI)
qed

end

end

setup Sign.parent_path

setup Sign.mandatory_path "invmap"

lemma idle:
  shows "spec.invmap af sf vf spec.idle
       = spec.term.none (spec.rel (UNIV × map_prod sf sf -` Id) :: ('a, 's, unit) spec)" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI spec.singleton_le_extI])
  have "sf s = sf s'"
    if "(a, s, s')  trace.steps' s0 xs"
   and "(λx. sf (snd x)) ` set xs  {sf s0}"
   for a :: 'a and s s' s0 :: 's and xs :: "('a × 's) list"
    using that by (induct xs arbitrary: s0) (auto simp: trace.steps'.Cons_eq_if split: if_split_asm)
  then show "σ  ?rhs" if "σ  ?lhs" for σ
    using that by (clarsimp simp: spec.singleton.le_conv image_image)
  have "sf s' = sf s"
    if "trace.steps' s xs  UNIV × map_prod sf sf -` Id"
   and "(a, s')  set xs"
   for a s s' and xs :: "('a × 's) list"
    using that
    by (induct xs arbitrary: s)
       (auto simp: Diff_subset_conv trace.steps'.Cons_eq_if split: if_split_asm)
  then show "σ  ?lhs" if "σ  ?rhs" for σ
    using that by (clarsimp simp: spec.singleton.le_conv)
qed

lemma inf_rel:
  shows "spec.rel (map_prod af (map_prod sf sf) -` (r  UNIV × Id))  spec.invmap af sf vf P = spec.invmap af sf vf (spec.rel r  P)"
    and "spec.invmap af sf vf P  spec.rel (map_prod af (map_prod sf sf) -` (r  UNIV × Id)) = spec.invmap af sf vf (P  spec.rel r)"
by (simp_all add: inf_commute spec.invmap.rel spec.invmap.inf)

lemma action: ―‹ (* could restrict the stuttering expansion to range af› or an arbitrary element of that ›
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
  fixes F :: "('w × 'b × 't × 't) set"
  defines "F'  map_prod id (map_prod af (map_prod sf sf))
                  -` (F  {(v, a', s, s) |v a a' s. (v, a, s, s)  F  ¬surj af})"
  shows "spec.invmap af sf vf (spec.action F)
       =  spec.rel (UNIV × map_prod sf sf -` Id)
           (λ_::unit. spec.action F'
           (λv. spec.rel (UNIV × map_prod sf sf -` Id)
           (λ_::unit. v'vf -` {v}. spec.return v')))" (is "?lhs = ?rhs")
proof(rule antisym[OF spec.singleton_le_extI], unfold spec.singleton.invmap_le_conv)
  have *: "sf x = sf y"
     if "(λx. sf (snd x)) ` set xs  {sf s}"
    and "(a, x, y)  trace.steps' s xs"
    for s a x y and v :: 'v and xs :: "('a × 's) list"
    using that
    by (induct xs arbitrary: s; clarsimp simp: trace.steps'.Cons_eq_if split: if_split_asm; metis)
  show "σ  ?rhs" if "trace.map af sf vf σ  spec.action F" for σ
  proof(cases " (trace.map af sf vf σ) = trace.T (trace.init (trace.map af sf vf σ)) [] None")
    case True then show ?thesis
      by (cases σ)
         (force simp: spec.singleton.le_conv trace.natural_def trace.natural'.eq_Nil_conv image_image
                dest: *
               intro: spec.bind.incompleteI)
  next
    case False with that show ?thesis
    proof(cases rule: spec.singleton.action_not_idle_le_splitE)
      case (return v a) then show ?thesis
        by (cases σ; clarsimp simp: image_image)
           ( rule spec.action.stutterI[where v=v and a="inv af a"]
                  spec.bind.continueI[where ys="[]", simplified]
           | ( fastforce simp: spec.singleton.le_conv F'_def trace.final'.map_idle surj_f_inv_f dest: * )
           )+
    next
      case (step v a ys zs) then show ?thesis
        by (cases σ; clarsimp simp: map_eq_append_conv image_image split: option.split_asm)
           (rule spec.bind.continueI
                 spec.bind.continueI[where xs="[x]" for x, simplified]
                 spec.bind.incompleteI[where g="Sup X" for X]
                 spec.bind.continueI[where ys="[]", simplified]
           | ( rule spec.action.stepI; force simp: F'_def trace.final'.map_idle )
           | ( fastforce simp: spec.singleton.le_conv trace.final'.map_idle dest: * )
           )+
    qed
  qed
  have *: "map_prod af (map_prod sf sf) ` (UNIV × map_prod sf sf -` Id) - UNIV × Id = {}" by blast
  have "(v, a, s, s')  F'  sf s, [(af a, sf s')], None  spec.action F" for v a s s'
    by (auto simp: F'_def spec.action.stepI intro: order.trans[OF spec.idle.minimal_le spec.idle.action_le])
  moreover
  have "(vf v, a, s, s')  F'; sf s' = trace.init σ; σ  spec.pre (λs. s  range sf); σ  spec.return (vf v)
      sf s, (af a, trace.init σ) # trace.rest σ, trace.term σ  spec.action F" for v a s s' σ
    by (auto simp: F'_def spec.action.stepI spec.action.stutterI spec.singleton.le_conv
                   spec.singleton.Cons[where as="trace.rest σ"]
            intro: order.trans[OF spec.idle.minimal_le spec.idle.action_le]
            split: option.split_asm)
  ultimately have "spec.action (map_prod id (map_prod af (map_prod sf sf)) ` F')
                      (λv. xvf -` {v}. spec.pre (λs. s  range sf)  spec.return v)
                   spec.action F"
    by (subst spec.action_def)
       (auto simp: spec.bind.SupL spec.bind.supL spec.bind.singletonL spec.idle_le split_def spec.term.none.singleton)
  then show "?rhs  ?lhs"
    apply (fold spec.map_invmap.galois)
    apply (strengthen ord_to_strengthen(1)[OF spec.map.bind_le])+
    apply (strengthen ord_to_strengthen[OF spec.map.rel_le])
    apply (strengthen ord_to_strengthen[OF spec.map.action_le])
    apply (subst (1 2) spec.rel.minus_Id[where A=UNIV, symmetric])
    apply (simp add: image_image * spec.map.return spec.rel.empty spec.bind.SupL spec.bind.returnL
                     spec.idle.action_le spec.idle.bind_le_conv spec.bind.SUPR_not_empty spec.bind.supR
                     spec.bind.return spec.map.Sup)
    done
qed

lemma return:
  fixes af :: "'a  'b"
  fixes sf :: "'s  't"
  fixes vf :: "'v  'w"
  fixes F :: "('w × 'b × 't × 't) set"
  shows "spec.invmap af sf vf (spec.return v)
       = spec.rel (UNIV × map_prod sf sf -` Id)  (λ_::unit. v'vf -` {v}. spec.return v')"
proof -
  have *: "spec.action ({()} × UNIV × map_prod sf sf -` Id) = spec.rel.act (UNIV × map_prod sf sf -` Id)"
    by (auto simp: spec.rel.act_def intro: spec.action.cong)
  show ?thesis
    apply (subst spec.return_def)
    apply (simp add: spec.invmap.action map_prod_vimage_Times)
    apply (subst sup.absorb1, force)
    apply (simp add: spec.action.return_const[where V="{v}" and W="{()}"] spec.bind.bind spec.bind.return *)
    apply (simp add: spec.rel.wind_bind flip: spec.bind.bind spec.rel.unfoldL)
    done
qed

setup Sign.parent_path

setup Sign.parent_path
(*<*)

end
(*>*)