| author | wenzelm | 
| Wed, 13 Sep 2000 22:29:17 +0200 | |
| changeset 9950 | 879e88b1e552 | 
| parent 8128 | 3a5864b465e2 | 
| child 10293 | 354e885b3e0a | 
| permissions | -rw-r--r-- | 
| 8044 | 1  | 
(* Title: HOL/UNITY/ELT  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1999 University of Cambridge  | 
|
5  | 
||
6  | 
leadsTo strengthened with a specification of the allowable sets transient parts  | 
|
| 
8122
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
7  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
8  | 
TRY INSTEAD (to get rid of the {} and to gain strong induction)
 | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
9  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
10  | 
  elt :: "['a set set, 'a program, 'a set] => ('a set) set"
 | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
11  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
12  | 
inductive "elt CC F B"  | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
13  | 
intrs  | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
14  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
15  | 
Weaken "A <= B ==> A : elt CC F B"  | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
16  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
17  | 
ETrans "[| F : A ensures A'; A-A' : CC; A' : elt CC F B |]  | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
18  | 
==> A : elt CC F B"  | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
19  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
20  | 
    Union  "{A. A: S} : Pow (elt CC F B) ==> (Union S) : elt CC F B"
 | 
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
21  | 
|
| 
 
b43ad07660b9
working version, with Alloc now working on the same state space as the whole
 
paulson 
parents: 
8072 
diff
changeset
 | 
22  | 
monos Pow_mono  | 
| 8044 | 23  | 
*)  | 
24  | 
||
25  | 
ELT = Project +  | 
|
26  | 
||
27  | 
consts  | 
|
28  | 
||
29  | 
(*LEADS-TO constant for the inductive definition*)  | 
|
30  | 
  elt :: "['a set set, 'a program] => ('a set * 'a set) set"
 | 
|
31  | 
||
32  | 
||
33  | 
inductive "elt CC F"  | 
|
34  | 
intrs  | 
|
35  | 
||
| 
8072
 
5b95377d7538
removing the "{} : CC" requirement for leadsTo[CC]
 
paulson 
parents: 
8055 
diff
changeset
 | 
36  | 
    Basis  "[| F : A ensures B;  A-B : (insert {} CC) |] ==> (A,B) : elt CC F"
 | 
| 8044 | 37  | 
|
38  | 
Trans "[| (A,B) : elt CC F; (B,C) : elt CC F |] ==> (A,C) : elt CC F"  | 
|
39  | 
||
40  | 
    Union  "{(A,B) | A. A: S} : Pow (elt CC F) ==> (Union S, B) : elt CC F"
 | 
|
41  | 
||
42  | 
monos Pow_mono  | 
|
43  | 
||
44  | 
||
45  | 
constdefs  | 
|
| 8128 | 46  | 
|
47  | 
(*the set of all sets determined by f alone*)  | 
|
48  | 
givenBy :: "['a => 'b] => 'a set set"  | 
|
49  | 
"givenBy f == range (%B. f-`` B)"  | 
|
| 8044 | 50  | 
|
51  | 
(*visible version of the LEADS-TO relation*)  | 
|
52  | 
leadsETo :: "['a set, 'a set set, 'a set] => 'a program set"  | 
|
53  | 
                                        ("(3_/ leadsTo[_]/ _)" [80,0,80] 80)
 | 
|
54  | 
    "leadsETo A CC B == {F. (A,B) : elt CC F}"
 | 
|
55  | 
||
56  | 
LeadsETo :: "['a set, 'a set set, 'a set] => 'a program set"  | 
|
57  | 
                                        ("(3_/ LeadsTo[_]/ _)" [80,0,80] 80)
 | 
|
58  | 
"LeadsETo A CC B ==  | 
|
59  | 
      {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) `` CC] B}"
 | 
|
60  | 
||
61  | 
end  |