equal
deleted
inserted
replaced
1 (* Title: HOL/IOA/meta_theory/Solve.thy |
|
2 ID: $Id$ |
|
3 Author: Tobias Nipkow & Konrad Slind |
|
4 Copyright 1994 TU Muenchen |
|
5 |
|
6 Weak possibilities mapping (abstraction) |
|
7 *) |
|
8 |
|
9 Solve = IOA + |
|
10 |
|
11 consts |
|
12 |
|
13 is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool" |
|
14 |
|
15 defs |
|
16 |
|
17 is_weak_pmap_def |
|
18 "is_weak_pmap(f,C,A) == |
|
19 (!s:starts_of(C). f(s):starts_of(A)) & |
|
20 (!s t a. reachable(C,s) & |
|
21 <s,a,t>:trans_of(C) |
|
22 --> if(a:externals(asig_of(C)), |
|
23 <f(s),a,f(t)>:trans_of(A), |
|
24 f(s)=f(t)))" |
|
25 |
|
26 end |
|