Theory Mndetprefix

(*<*)
―‹ ******************************************************************** 
 * Project         : HOL-CSP - A Shallow Embedding of CSP in  Isabelle/HOL
 * Version         : 2.0
 *
 * Author          : Burkhart Wolff, Safouan Taha, Lina Ye.
 *                   (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff)
 *
 * This file       : A Combined CSP Theory
 *
 * Copyright (c) 2009 Université Paris-Sud, France
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are
 * met:
 *
 *     * Redistributions of source code must retain the above copyright
 *       notice, this list of conditions and the following disclaimer.
 *
 *     * Redistributions in binary form must reproduce the above
 *       copyright notice, this list of conditions and the following
 *       disclaimer in the documentation and/or other materials provided
 *       with the distribution.
 *
 *     * Neither the name of the copyright holders nor the names of its
 *       contributors may be used to endorse or promote products derived
 *       from this software without specific prior written permission.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
 * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
 * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
 * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
 * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
 * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
 * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
 * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
 * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
 * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 ******************************************************************************›
(*>*)

theory Mndetprefix
  imports Process Stop Mprefix Ndet
begin

section‹Multiple non deterministic operator›

definition
        Mndetprefix   :: "[ set,    process]   process" 
        where   "Mndetprefix A P  if A = {} 
                                   then STOP
                                   else Abs_process( xA.  (x  P x),
                                                     xA.  𝒟(x  P x))"
syntax
  "_Mndetprefix"       :: "pttrn  'a set  'a process  'a process" 
                          ("(3__  / _)" [0, 0, 70] 70)

translations
  "xA P"  "CONST Mndetprefix A (λx. P)"

lemma mt_Mndetprefix[simp] : "Mndetprefix {} P = STOP"
  unfolding Mndetprefix_def   by simp

lemma Mndetprefix_is_process : "A  {}  is_process ( xA. (x  P x),  xA.   𝒟(x  P x))"
  unfolding is_process_def FAILURES_def DIVERGENCES_def
  apply auto
  using is_processT1 apply auto[1]
  using is_processT2 apply blast
  using is_processT3_SR apply blast
  using is_processT4 apply blast  
  using is_processT5_S1 apply blast
  using is_processT6 apply blast
  using is_processT7 apply blast
  using NF_ND apply auto[1]
  using is_processT9 by blast  

lemma T_Mndetprefix1 : "𝒯 (Mndetprefix {} P) = {[]}"
  unfolding Mndetprefix_def by(simp add: T_STOP)

lemma rep_abs_Mndetprefix[simp]: "A  {} 
     (Rep_process (Abs_process(xA. (x  P x),xA. 𝒟 (x  P x)))) = 
      (xA. (x  P x), xA. 𝒟 (x  P x))"
  apply(subst Process.process.Abs_process_inverse)
  by(auto intro: Mndetprefix_is_process[simplified])
  
lemma T_Mndetprefix: "A{}  𝒯 (Mndetprefix A P) = ( xA. 𝒯 (x  P x))"
  unfolding Mndetprefix_def
  apply(simp, subst Traces_def, simp add: TRACES_def FAILURES_def)
  apply(auto intro: Mndetprefix_is_process[simplified])
  unfolding TRACES_def FAILURES_def apply(cases "A = {}") 
   apply(auto intro: F_T D_T simp: Nil_elem_T)
  using NF_NT by blast

lemma F_Mndetprefix1 : " (Mndetprefix {} P) = {(s, X). s = []}"
  unfolding Mndetprefix_def by(simp add: F_STOP)

lemma F_Mndetprefix: "A{}   (Mndetprefix A P) = ( xA. (x  P x))"
  unfolding Mndetprefix_def 
  by(simp, subst Failures_def, auto simp : FAILURES_def F_Mndetprefix1)


lemma D_Mndetprefix1 : "𝒟 (Mndetprefix {} P) = {}"
  unfolding Mndetprefix_def by(simp add: D_STOP)

lemma D_Mndetprefix : "A{}  𝒟(Mndetprefix A P) = ( xA. 𝒟 (x  P x))"
  unfolding Mndetprefix_def 
  apply(simp, subst D_def, subst Process.process.Abs_process_inverse)
   by(auto intro: Mndetprefix_is_process[simplified] simp: DIVERGENCES_def)
  

text‹ Thus we know now, that Mndetprefix yields processes. Direct consequences are the following
  distributivities: ›

lemma Mndetprefix_unit : "( x  {a}  P x)  = (a  P a)" 
  by(auto simp : Process.Process_eq_spec F_Mndetprefix D_Mndetprefix)

lemma Mndetprefix_Un_distrib : "A {}  B {}  ( x  AB  P x) = (( x  A  P x)  ( x  B  P x))"
  by(auto simp : Process.Process_eq_spec F_Ndet D_Ndet F_Mndetprefix D_Mndetprefix)

text‹ The two lemmas @{thm Mndetprefix_unit} and @{thm Mndetprefix_Un_distrib} together give us that @{const Mndetprefix} 
      can be represented by a fold in the finite case. ›                                         

lemma Mndetprefix_distrib_unit : "A-{a}  {}   ( x  insert a A  P x) = ((a  P a)  ( x  A-{a}  P x))"
  by (metis Un_Diff_cancel insert_is_Un insert_not_empty Mndetprefix_Un_distrib Mndetprefix_unit)

subsection‹Finite case Continuity›

text‹ This also implies that Mndetprefix is continuous for the
      finite @{term A} and an arbitrary body @{term f}: ›

lemma Mndetprefix_cont_finite[simp]:
assumes "finite A"
 and    "x. cont (f x)"
shows   "cont (λy.  z  A  f z y)"
proof(rule Finite_Set.finite_induct[OF `finite A`])
  show  "cont (λy. z{}  f z y)" by auto
next
  fix A fix a 
  assume "cont (λy. zA  f z y)" and "a  A"
  show   "cont (λy. zinsert a A  f z y)"
  proof(cases "A={}")
    case True
    then show ?thesis by(simp add: Mndetprefix_unit True `x. cont (f x)`)
  next
    case False
    have *  : "A-{a}  {}" by (simp add: False a  A)
    have ** : "A-{a} = A"  by (simp add: a  A)
    show ?thesis
      apply(simp only: Mndetprefix_distrib_unit[OF *], simp only: **)  
      by (simp add: cont (λy. zA  f z y) assms(2))
  qed
qed

subsection‹General case Continuity›

lemma mono_Mndetprefix[simp] : "monofun (Mndetprefix (A::'a set))" 
proof(cases "A={}")
  case True
  then show ?thesis by(auto simp: monofun_def)
next
  case False
  then show ?thesis apply(simp add: monofun_def, intro allI impI) 
    unfolding le_approx_def
    proof(simp add:T_Mndetprefix F_Mndetprefix D_Mndetprefix, intro conjI)
      fix x::"'a  'a process" and y::"'a  'a process"  
      assume "A  {}" and "x  y"
      show "(xA.  𝒟 (x y x))  (xaA. 𝒟 (xa  x xa))" 
        by (metis (mono_tags, lifting) SUP_mono x  y fun_below_iff le_approx1 mono_Mprefix0 write0_def)
    next
      fix x::"'a  'a process" and y::"'a   'a process"  
      assume *:"A  {}" and **:"x  y"
      have *** : "z. z  A  x z  y z " by (simp add: x  y fun_belowD)
      with * show "s. (xaA. s  𝒟 (xa  x xa))  Ra (Mndetprefix A x) s = Ra (Mndetprefix A y) s"
        unfolding Ra_def
        by (auto simp:proc_ord2a mono_Mprefix0 write0_def F_Mndetprefix) 
           (meson le_approx2 mono_Mprefix0 write0_def)+
    next
      fix x::"'a  'a process" and y::"'a   'a process"  
      assume *:"A  {}" and "x  y"
      have *** : "z. z  A  (z  x z)  (z  y z) "
        by (metis x  y fun_below_iff mono_Mprefix0 write0_def)
      with * show "min_elems (xaA. 𝒟 (xa  x xa))  (xA. 𝒯 (x  y x))"
        unfolding min_elems_def apply auto 
        by (metis Set.set_insert elem_min_elems insert_subset le_approx3 le_less min_elems5)
    qed
qed  

lemma Mndetprefix_chainpreserving: "chain Y  chain (λi. ( z  A  Y i z))"
  apply(rule chainI, rename_tac i)
  apply(frule_tac i="i" in chainE)
  by (simp add: below_fun_def mono_Mprefix0 write0_def monofunE)

lemma contlub_Mndetprefix : "contlub (Mndetprefix A)"
proof(cases "A={}")
  case True
  then show ?thesis by(auto simp: contlub_def)
next
  case False
  show ?thesis
  proof(rule contlubI, rule proc_ord_proc_eq_spec)
    fix Y :: "nat  'a  'a process"
    assume a:"chain Y"
    show "(xA  (i. Y i) x)  (i. xA  Y i x)"
    proof(simp add:le_approx_def, intro conjI allI impI)
      show "𝒟 (i. xA  Y i x)  𝒟 (xA  Lub Y x)"
        using a False D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] 
              limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] 
        

        apply (auto simp add:write0_def D_Mprefix D_LUB[OF ch2ch_fun[OF a]] 
                             limproc_is_thelub_fun[OF a] D_Mndetprefix) 

        by (metis (mono_tags, opaque_lifting) event.inject)
    next
      fix s :: "'a event list"
      assume "s  𝒟 (xA  Lub Y x)"
      show "Ra (xA  Lub Y x) s = Ra (i. xA  Y i x) s" 
        unfolding Ra_def 
        using a False F_LUB[OF Mndetprefix_chainpreserving[OF a], of A] 
              limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] 
        apply (auto simp add:write0_def F_Mprefix F_LUB[OF ch2ch_fun[OF a]] 
                             limproc_is_thelub_fun[OF a] F_Mndetprefix)
        by (metis (mono_tags, opaque_lifting) event.inject)
    next
      show "min_elems (𝒟 (xA  Lub Y x))  𝒯 (i. xA  Y i x)"
        unfolding min_elems_def
        using a False limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A]
              D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] 
              F_LUB[OF Mndetprefix_chainpreserving[OF a], of A] 
        by (auto simp add:D_T write0_def D_Mprefix F_Mprefix D_Mndetprefix F_Mndetprefix 
                          D_LUB[OF ch2ch_fun[OF a]] F_LUB[OF ch2ch_fun[OF a]] 
                          limproc_is_thelub_fun[OF a])
    qed
  next
    fix Y :: "nat  'a  'a process"
    assume a:"chain Y"      
    show "𝒟 (xA  (i. Y i) x)  𝒟 (i. xA  Y i x)"
        using a False D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] 
              limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] 
        by (auto simp add:write0_def D_Mprefix D_Mndetprefix D_LUB[OF ch2ch_fun[OF a]] 
                           limproc_is_thelub_fun[OF a])
  qed
qed

lemma Mndetprefix_cont[simp]: "(x. cont (f x))  cont (λy. ( z  A  (f z y)))"
  apply(rule_tac f = "λz y. (f y z)" in Cont.cont_compose, rule monocontlub2cont) 
  by (auto intro: mono_Mndetprefix contlub_Mndetprefix Fun_Cpo.cont2cont_lambda)

end