author | oheimb |
Thu, 30 Aug 2001 15:47:30 +0200 | |
changeset 11507 | 4b32a46ffd29 |
parent 10834 | a7897aebbffc |
child 11701 | 3d51fbf81c17 |
permissions | -rw-r--r-- |
7186 | 1 |
(* Title: HOL/UNITY/Lift_prog.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1999 University of Cambridge |
|
5 |
||
6 |
lift_prog, etc: replication of components |
|
7 |
*) |
|
8 |
||
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
9 |
Lift_prog = Rename + |
7186 | 10 |
|
11 |
constdefs |
|
12 |
||
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
13 |
insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
14 |
"insert_map i z f k == if k<i then f k |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
15 |
else if k=i then z |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
16 |
else f(k-1)" |
7186 | 17 |
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
18 |
delete_map :: "[nat, nat=>'b] => (nat=>'b)" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
19 |
"delete_map i g k == if k<i then g k else g (Suc k)" |
7186 | 20 |
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
21 |
lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
22 |
"lift_map i == %(s,(f,uu)). (insert_map i s f, uu)" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
23 |
|
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
24 |
drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
25 |
"drop_map i == %(g, uu). (g i, (delete_map i g, uu))" |
7186 | 26 |
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
27 |
lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set" |
10834 | 28 |
"lift_set i A == lift_map i ` A" |
7186 | 29 |
|
8251
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
30 |
lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program" |
9be357df93d4
New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents:
8122
diff
changeset
|
31 |
"lift i == rename (lift_map i)" |
7186 | 32 |
|
33 |
(*simplifies the expression of specifications*) |
|
34 |
constdefs |
|
35 |
sub :: ['a, 'a=>'b] => 'b |
|
7947 | 36 |
"sub == %i f. f i" |
7186 | 37 |
|
38 |
||
39 |
end |