src/HOL/IOA/meta_theory/IOA.thy
changeset 1151 c820b3cc3df0
parent 1052 e044350bfa52
child 1476 608483c2122a
--- a/src/HOL/IOA/meta_theory/IOA.thy	Wed Jun 21 15:14:58 1995 +0200
+++ b/src/HOL/IOA/meta_theory/IOA.thy	Wed Jun 21 15:47:10 1995 +0200
@@ -63,9 +63,9 @@
 defs
 
 state_trans_def
-  "state_trans asig R == \
-  \  (!triple. triple:R --> fst(snd(triple)):actions(asig)) & \
-  \  (!a. (a:inputs(asig)) --> (!s1. ? s2. (s1,a,s2):R))"
+  "state_trans asig R == 
+     (!triple. triple:R --> fst(snd(triple)):actions(asig)) & 
+     (!a. (a:inputs(asig)) --> (!s1. ? s2. (s1,a,s2):R))"
 
 
 asig_of_def   "asig_of == fst"
@@ -73,9 +73,9 @@
 trans_of_def  "trans_of == (snd o snd)"
 
 ioa_def
-  "IOA(ioa) == (is_asig(asig_of(ioa))      &                            \
-  \             (~ starts_of(ioa) = {})    &                            \
-  \             state_trans (asig_of ioa) (trans_of ioa))"
+  "IOA(ioa) == (is_asig(asig_of(ioa))      &                            
+                (~ starts_of(ioa) = {})    &                            
+                state_trans (asig_of ioa) (trans_of ioa))"
 
 
 (* An execution fragment is modelled with a pair of sequences:
@@ -83,15 +83,15 @@
  * Finite executions have None actions from some point on.
  *******)
 is_execution_fragment_def
-  "is_execution_fragment A ex ==                                        \
-  \  let act = fst(ex); state = snd(ex)                                 \
-  \  in !n a. (act(n)=None --> state(Suc(n)) = state(n)) &              \
-  \           (act(n)=Some(a) --> (state(n),a,state(Suc(n))):trans_of(A))"
+  "is_execution_fragment A ex ==                                        
+     let act = fst(ex); state = snd(ex)                                 
+     in !n a. (act(n)=None --> state(Suc(n)) = state(n)) &              
+              (act(n)=Some(a) --> (state(n),a,state(Suc(n))):trans_of(A))"
 
 
 executions_def
-  "executions(ioa) == {e. snd e 0:starts_of(ioa) &                      \
-\                        is_execution_fragment ioa e}"
+  "executions(ioa) == {e. snd e 0:starts_of(ioa) &                      
+                        is_execution_fragment ioa e}"
 
   
 reachable_def
@@ -103,10 +103,10 @@
 
 (* Restrict the trace to those members of the set s *)
 filter_oseq_def
-  "filter_oseq p s ==                                                   \
-\   (%i.case s(i)                                                       \
-\         of None => None                                               \
-\          | Some(x) => if p x then Some x else None)"
+  "filter_oseq p s ==                                                   
+   (%i.case s(i)                                                       
+         of None => None                                               
+          | Some(x) => if p x then Some x else None)"
 
 
 mk_trace_def
@@ -115,13 +115,13 @@
 
 (* Does an ioa have an execution with the given trace *)
 has_trace_def
-  "has_trace ioa b ==                                               \
-\     (? ex:executions(ioa). b = mk_trace ioa (fst ex))"
+  "has_trace ioa b ==                                               
+     (? ex:executions(ioa). b = mk_trace ioa (fst ex))"
 
 normal_form_def
-  "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) &   \
-\                    (!j. j ~: range(f) --> nf(j)= None) &   \
-\                    (!i. nf(i)=None --> (nf (Suc i)) = None)   "
+  "NF(tr) == @nf. ? f. mono(f) & (!i. nf(i)=tr(f(i))) &   
+                    (!j. j ~: range(f) --> nf(j)= None) &   
+                    (!i. nf(i)=None --> (nf (Suc i)) = None)   "
   
 (* All the traces of an ioa *)
 
@@ -134,10 +134,10 @@
 *)
   
 compat_asigs_def
-  "compat_asigs a1 a2 ==                                                \
- \ (((outputs(a1) Int outputs(a2)) = {}) &                              \
- \  ((internals(a1) Int actions(a2)) = {}) &                            \
- \  ((internals(a2) Int actions(a1)) = {}))"
+  "compat_asigs a1 a2 ==                                                
+   (((outputs(a1) Int outputs(a2)) = {}) &                              
+    ((internals(a1) Int actions(a2)) = {}) &                            
+    ((internals(a2) Int actions(a1)) = {}))"
 
 
 compat_ioas_def
@@ -145,52 +145,52 @@
 
 
 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))))"
+  "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
-  "(ioa1 || ioa2) ==                                                    \
-  \    (asig_comp (asig_of ioa1) (asig_of ioa2),                        \
-  \     {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)},        \
-  \     {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))        \
-  \          in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) & \
-  \             (if a:actions(asig_of(ioa1)) then                       \
-  \                (fst(s),a,fst(t)):trans_of(ioa1)                     \
-  \              else fst(t) = fst(s))                                  \
-  \             &                                                       \
-  \             (if a:actions(asig_of(ioa2)) then                       \
-  \                (snd(s),a,snd(t)):trans_of(ioa2)                     \
-  \              else snd(t) = snd(s))})"
+  "(ioa1 || ioa2) ==                                                    
+       (asig_comp (asig_of ioa1) (asig_of ioa2),                        
+        {pr. fst(pr):starts_of(ioa1) & snd(pr):starts_of(ioa2)},        
+        {tr. let s = fst(tr); a = fst(snd(tr)); t = snd(snd(tr))        
+             in (a:actions(asig_of(ioa1)) | a:actions(asig_of(ioa2))) & 
+                (if a:actions(asig_of(ioa1)) then                       
+                   (fst(s),a,fst(t)):trans_of(ioa1)                     
+                 else fst(t) = fst(s))                                  
+                &                                                       
+                (if a:actions(asig_of(ioa2)) then                       
+                   (snd(s),a,snd(t)):trans_of(ioa2)                     
+                 else snd(t) = snd(s))})"
 
 
 restrict_asig_def
-  "restrict_asig asig actns ==                                          \
-\    (inputs(asig) Int actns, outputs(asig) Int actns,                  \
-\     internals(asig) Un (externals(asig) - actns))"
+  "restrict_asig asig actns ==                                          
+    (inputs(asig) Int actns, outputs(asig) Int actns,                  
+     internals(asig) Un (externals(asig) - actns))"
 
 
 restrict_def
-  "restrict ioa actns ==                                               \
-\    (restrict_asig (asig_of ioa) actns, starts_of(ioa), trans_of(ioa))"
+  "restrict ioa actns ==                                               
+    (restrict_asig (asig_of ioa) actns, starts_of(ioa), trans_of(ioa))"
 
 
 ioa_implements_def
-  "ioa_implements ioa1 ioa2 ==   \
-\  ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) &   \
-\     (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) &   \
-\      traces(ioa1) <= traces(ioa2))"
+  "ioa_implements ioa1 ioa2 ==   
+  ((inputs(asig_of(ioa1)) = inputs(asig_of(ioa2))) &   
+     (outputs(asig_of(ioa1)) = outputs(asig_of(ioa2))) &   
+      traces(ioa1) <= traces(ioa2))"
  
 rename_def 
-"rename ioa ren ==  \
-\  (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))},         \
-\    {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))},        \
-\    {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}),     \
-\              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 ioa ren ==  
+  (({b. ? x. Some(x)= ren(b) & x : inputs(asig_of(ioa))},         
+    {b. ? x. Some(x)= ren(b) & x : outputs(asig_of(ioa))},        
+    {b. ? x. Some(x)= ren(b) & x : internals(asig_of(ioa))}),     
+              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)})"
 
 end