Theory Sliding_Choice

(*<*)
―‹ ********************************************************************
 * Project         : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL
 * Version         : 2.0
 *
 * Author          : Benoît Ballenghien, Safouan Taha, Burkhart Wolff, Lina Ye.
 *                   (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff)
 *
 * This file       : Sliding choice
 *
 * Copyright (c) 2009 Université Paris-Sud, France
 * Copyright (c) 2025 Université Paris-Saclay, 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.
 ******************************************************************************›
(*>*)


section ‹ Sliding Choice ›

(*<*)
theory  Sliding_Choice
  imports Deterministic_Choice Non_Deterministic_Choice
begin
  (*>*)

subsection ‹Definition›

definition Sliding :: ('a, 'r)processptick  ('a, 'r)processptick  ('a, 'r)processptick (infixl  80)
  where P  Q  (P  Q)  Q



subsection ‹Projections›

lemma F_Sliding:
   (P  Q) =  Q  {(s, X). s  []  (s, X)   P  
                               s = []  (r. s  𝒟 P  tick r  X  [tick r]  𝒯 P)}
  by (auto simp add: Sliding_def F_Ndet F_Det intro: is_processT8 is_processT6_TR_notin)

lemma D_Sliding: 𝒟 (P  Q) = 𝒟 P  𝒟 Q
  by (simp add: Sliding_def D_Ndet D_Det)

lemma T_Sliding: 𝒯 (P  Q) = 𝒯 P  𝒯 Q
  by (simp add: Sliding_def T_Ndet T_Det)

lemmas Sliding_projs = F_Sliding D_Sliding T_Sliding



subsection ‹Properties›

lemma Sliding_id: P  P = P
  by (simp add: Sliding_def)

(* 
text ‹Of course, term‹P ⊳ STOP ≠ STOP› and term‹P ⊳ STOP ≠ P› in general.›
lemma ‹∃P. P ⊳ STOP ≠ STOP ∧ P ⊳ STOP ≠ P›
proof (intro exI)
  show ‹SKIP undefined ⊳ STOP ≠ STOP ∧ SKIP undefined ⊳ STOP ≠ SKIP undefined›
    by (metis Det_STOP Ndet_commute SKIP_F_iff SKIP_Neq_STOP
            STOP_F_iff Sliding_def mono_Ndet_F_left)
qed
   *)




subsection ‹Continuity›

text ‹From the definition, monotony and continuity is obvious.›

lemma mono_Sliding : P  P'  Q  Q'  P  Q  P'  Q'
  unfolding Sliding_def by (simp add: mono_Det mono_Ndet)

lemma Sliding_cont[simp] : cont f  cont g  cont (λx. f x  g x)
  by (simp add: Sliding_def)

(*<*)
end
  (*>*)