| author | wenzelm |
| Fri, 27 Jan 2006 19:03:19 +0100 | |
| changeset 18812 | a4554848b59e |
| parent 17288 | aa3833fb7bee |
| child 19801 | b2af2549efd1 |
| permissions | -rw-r--r-- |
| 4530 | 1 |
(* Title: HOL/IOA/Solve.thy |
|
3078
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
3 |
Author: Tobias Nipkow & Konrad Slind |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
4 |
Copyright 1994 TU Muenchen |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
5 |
*) |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
6 |
|
| 17288 | 7 |
header {* Weak possibilities mapping (abstraction) *}
|
8 |
||
9 |
theory Solve |
|
10 |
imports IOA |
|
11 |
begin |
|
|
3078
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
12 |
|
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
13 |
constdefs |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
14 |
|
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
15 |
is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool"
|
| 17288 | 16 |
"is_weak_pmap f C A == |
17 |
(!s:starts_of(C). f(s):starts_of(A)) & |
|
18 |
(!s t a. reachable C s & |
|
19 |
(s,a,t):trans_of(C) |
|
20 |
--> (if a:externals(asig_of(C)) then |
|
21 |
(f(s),a,f(t)):trans_of(A) |
|
|
3078
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
22 |
else f(s)=f(t)))" |
|
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
23 |
|
| 17288 | 24 |
ML {* use_legacy_bindings (the_context ()) *}
|
25 |
||
|
3078
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
mueller
parents:
diff
changeset
|
26 |
end |