src/HOLCF/IOA/meta_theory/Automata.thy
changeset 17233 41eee2e7b465
parent 14981 e73f8140af78
child 19741 f65265d71426
--- a/src/HOLCF/IOA/meta_theory/Automata.thy	Fri Sep 02 15:54:47 2005 +0200
+++ b/src/HOLCF/IOA/meta_theory/Automata.thy	Fri Sep 02 17:23:59 2005 +0200
@@ -1,22 +1,22 @@
 (*  Title:      HOLCF/IOA/meta_theory/Automata.thy
     ID:         $Id$
     Author:     Olaf Müller, Konrad Slind, Tobias Nipkow
+*)
 
-The I/O automata of Lynch and Tuttle in HOLCF.
-*)   
-
-		       
-Automata = Asig +
+header {* The I/O automata of Lynch and Tuttle in HOLCF *}
 
-default type
- 
+theory Automata
+imports Asig
+begin
+
+defaultsort type
+
 types
-   ('a,'s)transition       =    "'s * 'a * 's"
-   ('a,'s)ioa              =    "'a signature * 's set * ('a,'s)transition set * 
-                                 (('a set) set) * (('a set) set)"
+  ('a, 's) transition = "'s * 'a * 's"
+  ('a, 's) ioa = "'a signature * 's set * ('a,'s)transition set * ('a set set) * ('a set set)"
 
 consts
- 
+
   (* IO automata *)
 
   asig_of        ::"('a,'s)ioa => 'a signature"
@@ -26,10 +26,10 @@
   sfair_of       ::"('a,'s)ioa => ('a set) set"
 
   is_asig_of     ::"('a,'s)ioa => bool"
-  is_starts_of	 ::"('a,'s)ioa => bool"
-  is_trans_of	 ::"('a,'s)ioa => bool"
-  input_enabled	 ::"('a,'s)ioa => bool"
-  IOA	         ::"('a,'s)ioa => bool"
+  is_starts_of   ::"('a,'s)ioa => bool"
+  is_trans_of    ::"('a,'s)ioa => bool"
+  input_enabled  ::"('a,'s)ioa => bool"
+  IOA            ::"('a,'s)ioa => bool"
 
   (* constraints for fair IOA *)
 
@@ -41,7 +41,7 @@
   enabled        ::"('a,'s)ioa => 'a => 's => bool"
   Enabled    ::"('a,'s)ioa => 'a set => 's => bool"
 
-  (* action set keeps enabled until probably disabled by itself *) 
+  (* action set keeps enabled until probably disabled by itself *)
 
   en_persistent  :: "('a,'s)ioa => 'a set => bool"
 
@@ -61,7 +61,7 @@
 
   (* hiding and restricting *)
   hide_asig     :: "['a signature, 'a set] => 'a signature"
-  hide          :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
+  "hide"        :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
   restrict_asig :: "['a signature, 'a set] => 'a signature"
   restrict      :: "[('a,'s)ioa, 'a set] => ('a,'s)ioa"
 
@@ -70,7 +70,7 @@
   rename        :: "('a, 'b)ioa => ('c => 'a option) => ('c,'b)ioa"
 
 
-syntax 
+syntax
 
   "_trans_of"  :: "'s => 'a => ('a,'s)ioa => 's => bool"  ("_ -_--_-> _" [81,81,81,81] 100)
   "reachable"  :: "[('a,'s)ioa, 's] => bool"
@@ -84,15 +84,15 @@
 
 syntax (xsymbols)
 
-  "_trans_of"  :: "'s => 'a => ('a,'s)ioa => 's => bool"  
-                  ("_ \\<midarrow>_\\<midarrow>_\\<longrightarrow> _" [81,81,81,81] 100)
-  "op ||"         ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr "\\<parallel>" 10)
+  "_trans_of"  :: "'s => 'a => ('a,'s)ioa => 's => bool"
+                  ("_ \<midarrow>_\<midarrow>_\<longrightarrow> _" [81,81,81,81] 100)
+  "op ||"         ::"[('a,'s)ioa, ('a,'t)ioa] => ('a,'s*'t)ioa"  (infixr "\<parallel>" 10)
 
 
-inductive "reachable C" 
-   intrs  
-    reachable_0  "s:(starts_of C) ==> s : reachable C"
-    reachable_n  "[|s:reachable C; (s,a,t):trans_of C|] ==> t:reachable C"
+inductive "reachable C"
+   intros
+    reachable_0:  "s:(starts_of C) ==> s : reachable C"
+    reachable_n:  "[|s:reachable C; (s,a,t):trans_of C|] ==> t:reachable C"
 
 
 translations
@@ -105,72 +105,68 @@
   "out A"         == "outputs (asig_of A)"
   "local A"       == "locals (asig_of A)"
 
-
-
 defs
 
 (* --------------------------------- IOA ---------------------------------*)
 
+asig_of_def:   "asig_of == fst"
+starts_of_def: "starts_of == (fst o snd)"
+trans_of_def:  "trans_of == (fst o snd o snd)"
+wfair_of_def:  "wfair_of == (fst o snd o snd o snd)"
+sfair_of_def:  "sfair_of == (snd o snd o snd o snd)"
+
+is_asig_of_def:
+  "is_asig_of A == is_asig (asig_of A)"
+
+is_starts_of_def:
+  "is_starts_of A ==  (~ starts_of A = {})"
+
+is_trans_of_def:
+  "is_trans_of A ==
+    (!triple. triple:(trans_of A) --> fst(snd(triple)):actions(asig_of A))"
+
+input_enabled_def:
+  "input_enabled A ==
+    (!a. (a:inputs(asig_of A)) --> (!s1. ? s2. (s1,a,s2):(trans_of A)))"
 
 
-asig_of_def   "asig_of == fst"
-starts_of_def "starts_of == (fst o snd)"
-trans_of_def  "trans_of == (fst o snd o snd)"
-wfair_of_def  "wfair_of == (fst o snd o snd o snd)"
-sfair_of_def  "sfair_of == (snd o snd o snd o snd)"
-
-is_asig_of_def
-  "is_asig_of A == is_asig (asig_of A)" 
-
-is_starts_of_def 
-  "is_starts_of A ==  (~ starts_of A = {})"
-
-is_trans_of_def
-  "is_trans_of A == 
-    (!triple. triple:(trans_of A) --> fst(snd(triple)):actions(asig_of A))"
-
-input_enabled_def
-  "input_enabled A ==
-    (!a. (a:inputs(asig_of A)) --> (!s1. ? s2. (s1,a,s2):(trans_of A)))" 
-
-
-ioa_def
-  "IOA A == (is_asig_of A    &                            
-             is_starts_of A  &                            
+ioa_def:
+  "IOA A == (is_asig_of A    &
+             is_starts_of A  &
              is_trans_of A   &
              input_enabled A)"
 
 
-invariant_def "invariant A P == (!s. reachable A s --> P(s))"
+invariant_def: "invariant A P == (!s. reachable A s --> P(s))"
 
 
 (* ------------------------- parallel composition --------------------------*)
 
 
-compatible_def
-  "compatible A B ==  
-  (((out A Int out B) = {}) &                              
-   ((int A Int act B) = {}) &                            
+compatible_def:
+  "compatible A B ==
+  (((out A Int out B) = {}) &
+   ((int A Int act B) = {}) &
    ((int B Int act A) = {}))"
 
-asig_comp_def
-  "asig_comp a1 a2 ==                                                   
-     (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)),      
-       (outputs(a1) Un outputs(a2)),                                   
+asig_comp_def:
+  "asig_comp a1 a2 ==
+     (((inputs(a1) Un inputs(a2)) - (outputs(a1) Un outputs(a2)),
+       (outputs(a1) Un outputs(a2)),
        (internals(a1) Un internals(a2))))"
 
-par_def
-  "(A || B) ==                                                    
-      (asig_comp (asig_of A) (asig_of B),                        
-       {pr. fst(pr):starts_of(A) & snd(pr):starts_of(B)},        
-       {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))        
-            in (a:act A | a:act B) & 
-               (if a:act A then                       
-                  (fst(s),a,fst(t)):trans_of(A)                     
-                else fst(t) = fst(s))                                  
-               &                                                       
-               (if a:act B then                       
-                  (snd(s),a,snd(t)):trans_of(B)                     
+par_def:
+  "(A || B) ==
+      (asig_comp (asig_of A) (asig_of B),
+       {pr. fst(pr):starts_of(A) & snd(pr):starts_of(B)},
+       {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))
+            in (a:act A | a:act B) &
+               (if a:act A then
+                  (fst(s),a,fst(t)):trans_of(A)
+                else fst(t) = fst(s))
+               &
+               (if a:act B then
+                  (snd(s),a,snd(t)):trans_of(B)
                 else snd(t) = snd(s))},
         wfair_of A Un wfair_of B,
         sfair_of A Un sfair_of B)"
@@ -178,84 +174,84 @@
 
 (* ------------------------ hiding -------------------------------------------- *)
 
-restrict_asig_def
-  "restrict_asig asig actns ==                                          
-    (inputs(asig) Int actns, 
-     outputs(asig) Int actns,                  
+restrict_asig_def:
+  "restrict_asig asig actns ==
+    (inputs(asig) Int actns,
+     outputs(asig) Int actns,
      internals(asig) Un (externals(asig) - actns))"
 
-(* Notice that for wfair_of and sfair_of nothing has to be changed, as 
-   changes from the outputs to the internals does not touch the locals as 
+(* Notice that for wfair_of and sfair_of nothing has to be changed, as
+   changes from the outputs to the internals does not touch the locals as
    a whole, which is of importance for fairness only *)
 
-restrict_def
-  "restrict A actns ==                                               
-    (restrict_asig (asig_of A) actns, 
-     starts_of A, 
+restrict_def:
+  "restrict A actns ==
+    (restrict_asig (asig_of A) actns,
+     starts_of A,
      trans_of A,
      wfair_of A,
      sfair_of A)"
 
-hide_asig_def
-  "hide_asig asig actns ==                                          
-    (inputs(asig) - actns, 
-     outputs(asig) - actns,                  
+hide_asig_def:
+  "hide_asig asig actns ==
+    (inputs(asig) - actns,
+     outputs(asig) - actns,
      internals(asig) Un actns)"
 
-hide_def
-  "hide A actns ==                                               
-    (hide_asig (asig_of A) actns, 
-     starts_of A, 
+hide_def:
+  "hide A actns ==
+    (hide_asig (asig_of A) actns,
+     starts_of A,
      trans_of A,
      wfair_of A,
      sfair_of A)"
 
 (* ------------------------- renaming ------------------------------------------- *)
-  
-rename_set_def
-  "rename_set A ren == {b. ? x. Some x = ren b & x : A}" 
+
+rename_set_def:
+  "rename_set A ren == {b. ? x. Some x = ren b & x : A}"
 
-rename_def 
-"rename ioa ren ==  
-  ((rename_set (inp ioa) ren,         
-    rename_set (out ioa) ren,        
-    rename_set (int ioa) ren),     
-   starts_of ioa,                                            
-   {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))    
-        in                                                      
+rename_def:
+"rename ioa ren ==
+  ((rename_set (inp ioa) ren,
+    rename_set (out ioa) ren,
+    rename_set (int ioa) ren),
+   starts_of ioa,
+   {tr. let s = fst(tr); a = fst(snd(tr));  t = snd(snd(tr))
+        in
         ? x. Some(x) = ren(a) & (s,x,t):trans_of ioa},
    {rename_set s ren | s. s: wfair_of ioa},
    {rename_set s ren | s. s: sfair_of ioa})"
 
 (* ------------------------- fairness ----------------------------- *)
 
-fairIOA_def
-  "fairIOA A == (! S : wfair_of A. S<= local A) & 
+fairIOA_def:
+  "fairIOA A == (! S : wfair_of A. S<= local A) &
                 (! S : sfair_of A. S<= local A)"
 
-input_resistant_def
-  "input_resistant A == ! W : sfair_of A. ! s a t. 
+input_resistant_def:
+  "input_resistant A == ! W : sfair_of A. ! s a t.
                         reachable A s & reachable A t & a:inp A &
                         Enabled A W s & s -a--A-> t
                         --> Enabled A W t"
 
-enabled_def
+enabled_def:
   "enabled A a s == ? t. s-a--A-> t"
 
-Enabled_def
+Enabled_def:
   "Enabled A W s == ? w:W. enabled A w s"
 
-en_persistent_def
-  "en_persistent A W == ! s a t. Enabled A W s & 
-                                 a ~:W & 
-                                 s -a--A-> t 
+en_persistent_def:
+  "en_persistent A W == ! s a t. Enabled A W s &
+                                 a ~:W &
+                                 s -a--A-> t
                                  --> Enabled A W t"
-was_enabled_def
+was_enabled_def:
   "was_enabled A a t == ? s. s-a--A-> t"
 
-set_was_enabled_def
+set_was_enabled_def:
   "set_was_enabled A W t == ? w:W. was_enabled A w t"
 
+ML {* use_legacy_bindings (the_context ()) *}
 
 end
-