Theory POTS

(*<*)
―‹ ********************************************************************
 * Project         : HOL-CSPM - Architectural operators for HOL-CSP
 *
 * Author          : Benoît Ballenghien, Safouan Taha, Burkhart Wolff
 *
 * This file       : POTS example
 *
 * Copyright (c) 2023 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.
 ******************************************************************************›
(*>*)


chapter‹ Example: Plain Old Telephone System›

text‹The "Plain Old Telephone Service is a standard medium-size example for
     architectural modeling of a concurrent system. 

     Plain old telephone service (POTS), or plain ordinary telephone system,[1] 
     is a retronym for voice-grade telephone service employing analog signal  transmission over 
     copper loops. POTS was the standard service offering from telephone companies from 1876 until 
     1988[2] in the United States when the Integrated Services Digital Network (ISDN) Basic Rate 
     Interface (BRI) was introduced, followed by cellular telephone systems, and voice over 
     IP (VoIP). POTS remains the basic form of residential and small business service connection 
     to the telephone network in many parts of the world. The term reflects the technology that has 
     been available since the introduction of the public telephone system in the late 19th century, 
     in a form mostly unchanged despite the introduction of Touch-Tone dialing, electronic telephone 
     exchanges and fiber-optic communication into the public switched telephone network (PSTN).


     C.f. wikipedia 🌐‹https://en.wikipedia.org/wiki/Plain_old_telephone_service›.
› (* rework this small text *)


theory POTS                                               
  imports CSPM 
begin 

text ‹We need to see typint as a classcpo.›
―‹We may replace this instantiation by an import of "HOLCF-Library.Int_Discrete"›
instantiation int :: discrete_cpo
begin

definition below_int_def:
  "(x::int)  y  x = y"

instance proof
qed (rule below_int_def)

end



section‹ The Alphabet and Basic Types of POTS ›

text‹Underlying terminology apparent in the acronyms: 
 T-side (target side, callee side)
 O-side (originator (?) side, caller side)›

datatype MtcO = Osetup | Odiscon_o
datatype MctO = Obusy  | Oalert    | Oconnect | Odiscon_t
datatype MtcT = Tbusy  | Talert    | Tconnect | Tdiscon_t
datatype MctT = Tsetup | Tdiscon_o

type_synonym Phones = int


datatype channels = tcO Phones × MtcO            ―‹  ›       
                  | ctO Phones × MctO            
                  | tcT Phones × MtcT × Phones 
                  | ctT Phones × MctT × Phones 
                  | tcOdial    Phones × Phones   
                  | StartReject Phones            ―‹ phone x rejects from now on to be called ›
                  | EndReject   Phones            ―‹ phone x accepts from now on to be called ›
                  | terminal    Phones 
                  | off_hook    Phones     
                  | on_hook     Phones      
                  | digits     Phones × Phones ―‹ communication relation: x calls y ›     
                  | tone_ring   Phones    
                  | tone_quiet  Phones   
                  | tone_busy   Phones
                  | tone_dial   Phones    
                  | connected   Phones


locale POTS = 
  fixes min_phones :: int
    and max_phones :: int
    and VisibleEvents :: channels set
  assumes min_phones_g_1[simp]          :          1  min_phones
      and max_phones_g_min_phones[simp] : min_phones < max_phones
begin

definition phones :: Phones set where phones  {min_phones ..  max_phones}

lemma nonempty_phones[simp]: phones  {}
  and finite_phones[simp]: finite phones
  and at_least_two_phones[simp]: 2  card phones
  and not_singl_phone[simp]: phones - {p}  {}
  apply (simp_all add: phones_def)
  using max_phones_g_min_phones apply linarith+
  by (metis atLeastAtMost_iff less_le_not_le max_phones_g_min_phones order_refl singletonD subsetD)



definition  EventsIPhone :: Phones  channels set
  where    EventsIPhone u  {tone_ring u, tone_quiet u, tone_busy u, tone_dial u, connected u}
definition  EventsUser :: Phones  channels set
  where    EventsUser u  {off_hook u, on_hook u}  {x .  n. x = digits (u, n)}



section‹Auxilliaries to Substructure the Specification›

abbreviation Sliding ::  process   process   process (infixl  78)
  where P  Q  (P  Q)  Q
― ‹This operator is also called Timeout, more studied in future theories.›

abbreviation
  Tside_connected     :: Phones  Phones  channels process
where Tside_connected ts os  
           (ctT!(ts,Tdiscon_o,os)  tcT!(ts,Tdiscon_t,os)  EndReject!tsSKIP)
        (tcT!(ts,Tdiscon_t,os)  ctT!(ts,Tdiscon_o,os)  EndReject!tsSKIP)



abbreviation
  Oside_connected     :: Phones  channels process
where   Oside_connected ts 
            (ctO!(ts,Odiscon_t)  tcO!(ts,Odiscon_o)  EndReject!tsSKIP)
         (tcO!(ts,Odiscon_o)  ctO!(ts,Odiscon_t)  EndReject!tsSKIP)



abbreviation
   Oside1 :: [Phones, Phones]  channels process
where 
  Oside1 ts p   tcOdial!(ts,p)
	                    (ctO!(ts,Oalert)
                          ctO!(ts,Oconnect)
                          (Oside_connected ts))
                      (ctO!(ts,Oconnect) (Oside_connected ts))
                      (ctO!(ts,Obusy)  tcO!(ts,Odiscon_o)  EndReject!ts  SKIP)


definition
  ITside_connected    :: [Phones,Phones,channels process]  channels process
where
 ITside_connected ts os IT  (ctT(ts,Tdiscon_o,os)
                                (  (tone_busy!ts
                                        on_hook!ts
                                        tcT!(ts,Tdiscon_t,os)
                                        EndReject!ts 
                                        IT)
                                    (on_hook!ts
                                        tcT!(ts,Tdiscon_t,os)
                                        EndReject!ts
                                        IT)
                                   ))
                                 (on_hook!ts
                                      tcT!(ts,Tdiscon_t,os)
                                      ctT!(ts,Tdiscon_o,os)
                                      EndReject!ts
                                     IT)


section‹A Telephone ›

fixrec     T        :: Phones  channels process
       and Oside    :: Phones  channels process
       and Tside    :: Phones  channels process
       and NoReject :: Phones  channels process
       and Reject   :: Phones  channels process
where
   T_rec        [simp del]: Tts        = (Tsidets ; Tts)  (Osidets ; Tts)
 | Oside_rec    [simp del]: Osidets    = StartReject!ts 
                                               tcO!(ts,Osetup) 
                                               ( p  phones. Oside1 ts p)
 | Tside_rec    [simp del]: Tsidets    = ctT?(y,z,os)|((y,z)=(ts,Tsetup)) 
                                               StartReject!ts 
                                               (   tcT!(ts,Talert,os)
                                                      tcT!(ts,Tconnect,os)
                                                     (Tside_connected ts os)
                                                   (tcT!(ts,Tconnect,os)
                                                      (Tside_connected ts os)))  
 | NoReject_rec [simp del]: NoRejectts = StartReject!ts  Rejectts
 | Reject_rec   [simp del]: Rejectts   = ctT?(y,z,os)|(y=ts  z=Tsetup  osphones  osts)
                                                   (tcT!(ts,Tbusy,os)  Rejectts)
                                                   (EndReject!ts  NoRejectts)





definition Tel:: Phones  channels process
  where   Tel p  (Tp {StartReject p, EndReject p} NoRejectp) \ {StartReject p, EndReject p}




section‹ A Connector with the Network ›

fixrec     Call      :: Phones  channels process
       and BUSY      :: Phones  Phones  channels process
       and Connected :: Phones  Phones  channels process
where
   Call_rec  [simp del]: Callos     = (tcO!  (os,Osetup)  tcOdial?(x,ts)|(x=os)  (BUSYosts)) ; Callos
 | BUSY_rec  [simp del]: BUSYosts  = (if ts = os 
                              then ctO!(os,Obusy)  tcO!(os,Odiscon_o)  SKIP
                              else ctT!(ts,Tsetup,os)
                                   ( (tcT!(ts,Tbusy,os)
                                           ctO!(os,Obusy)
                                           tcO!(os,Odiscon_o)  SKIP)
                                       
                                         (tcT ! (ts,Talert,os)
                                           ctO!(os,Oalert)
                                           tcT!(ts,Tconnect,os)
                                           ctO!(os,Oconnect)
                                           Connectedosts)
                                       
                                         (tcT!(ts,Tconnect,os)
                                           ctO!(os,Oconnect)
                                           Connectedosts)))
 | Connected_rec [simp del]: Connectedosts =  (tcO!(os,Odiscon_o) 
                             (( (ctT!(ts,Tdiscon_o,os)  tcT!(ts,Tdiscon_t,os)  SKIP)
                                 
                                (tcT!(ts,Tdiscon_t,os) ctT!(ts,Tdiscon_o,os)  SKIP)
                              )
                              ; (ctO!(os,Odiscon_t)  SKIP)))
                              
                             (tcT!(ts,Tdiscon_t,os) 
                                     (  (ctO!(os,Odiscon_t) 
                                          ctT!(ts,Tdiscon_o,os) 
                                          tcO!(os,Odiscon_o) 
                                          SKIP )
                                        
                                        (tcO!(os,Odiscon_o) 
                                          ctT!(ts,Tdiscon_o,os) 
                                          ctO!(os,Odiscon_t)
                                          SKIP) 
                                     )
                             )



section‹ Combining NETWORK and TELEPHONES to a SYSTEM ›

definition  NETWORK     :: channels process
  where      NETWORK       (||| os ∈# (mset_set phones). Callos)

definition  TELEPHONES  :: channels process               
  where      TELEPHONES    (||| ts ∈# (mset_set phones). Tel ts)

definition  SYSTEM      :: channels process
  where      SYSTEM        NETWORK VisibleEvents TELEPHONES

text ‹We underline here the usefulness of the architectural operators, especially constMultiSync
      but also constMultiNdet which appears in constOside recursive definition.›




section‹Simple Model of a User ›

fixrec     User      :: Phones  channels process
       and UserSCon  :: Phones  channels process
where
  User_rec[simp del]  : Useru = (off_hook!u 
                         (tone_dial!u 
                          ( p  phones. digits!(u,p)tone_quiet!u
                                          (  (tone_ring!uconnected!uUserSConu)
                                            (connected!uUserSConu)
                                            (tone_busy!uon_hook!uUseru)
                                          )
                          )
                         )
                        (connected!u  UserSConu)
                       )
                         (tone_ring!uoff_hook!uconnected!u UserSConu)
  | UserSCon_rec[simp del]: UserSConu = (tone_busy!u  on_hook!u  Useru)  (on_hook!u  Useru)



fixrec     User_Ndet      :: Phones  channels process
       and UserSCon_Ndet  :: Phones  channels process
where
  User_Ndet_rec[simp del]  : User_Ndetu = (off_hook!u 
                         (tone_dial!u 
                          ( p  phones. digits!(u,p)tone_quiet!u
                                          (  (tone_ring!uconnected!uUserSCon_Ndetu)
                                            (connected!uUserSCon_Ndetu)
                                            (tone_busy!uon_hook!uUser_Ndetu)
                                          )
                          )
                         )
                        (connected!u  UserSCon_Ndetu)
                       )
                         (tone_ring!uoff_hook!uconnected!u UserSCon_Ndetu)
  | UserSCon_Ndet_rec[simp del]: UserSCon_Ndetu = (tone_busy!u  on_hook!u  User_Ndetu)  (on_hook!u  User_Ndetu)



definition  ImplementT          :: Phones  channels process
  where    ImplementT ts  ((Tel ts) EventsIPhone ts  EventsUser ts (Userts))
                            \ (EventsIPhone ts  EventsUser ts)




section ‹ Toplevel Proof-Goals›

text‹ This has been proven in an ancient FDR model for @{term max_phones = 5}...  ›


lemma p  phones. deadlock_free (Tel p) oops
lemma p  phones. deadlock_free_v2 (Callp) oops
lemma deadlock_free_v2 NETWORK oops
lemma deadlock_free_v2 SYSTEM oops
lemma lifelock_free SYSTEM oops 
lemma p  phones. lifelock_free (ImplementT p) oops
lemma p  phones. Tel p FD ImplementT p oops
                                
lemma p  phones. Tel'p  F RUN UNIV oops
text‹this should represent "deterministic" in process-algebraic terms. . .›


end

end