src/HOL/HOLCF/IOA/ABP/Impl.thy
changeset 40774 0437dbc127b3
parent 35174 e15040ae75d7
child 40945 b8703f63bfb2
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/HOLCF/IOA/ABP/Impl.thy	Sat Nov 27 16:08:10 2010 -0800
@@ -0,0 +1,35 @@
+(*  Title:      HOLCF/IOA/ABP/Impl.thy
+    Author:     Olaf Müller
+*)
+
+header {* The implementation *}
+
+theory Impl
+imports Sender Receiver Abschannel
+begin
+
+types
+  'm impl_state = "'m sender_state * 'm receiver_state * 'm packet list * bool list"
+  (*  sender_state   *  receiver_state   *    srch_state  * rsch_state *)
+
+definition
+ impl_ioa :: "('m action, 'm impl_state)ioa" where
+ "impl_ioa = (sender_ioa || receiver_ioa || srch_ioa || rsch_ioa)"
+
+definition
+ sen :: "'m impl_state => 'm sender_state" where
+ "sen = fst"
+
+definition
+ rec :: "'m impl_state => 'm receiver_state" where
+ "rec = fst o snd"
+
+definition
+ srch :: "'m impl_state => 'm packet list" where
+ "srch = fst o snd o snd"
+
+definition
+ rsch :: "'m impl_state => bool list" where
+ "rsch = snd o snd o snd"
+
+end