| author | paulson | 
| Wed, 19 Jun 2002 12:48:55 +0200 | |
| changeset 13225 | b6fc6e4a0a24 | 
| parent 9403 | aad13b59b8d9 | 
| child 13790 | 8d7e9fce8c50 | 
| permissions | -rw-r--r-- | 
| 8256 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 1 | (* Title: HOL/UNITY/Rename.thy | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 4 | Copyright 2000 University of Cambridge | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 5 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 6 | Renaming of state sets | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 7 | *) | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 8 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 9 | Rename = Extend + | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 10 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 11 | constdefs | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 12 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 13 | rename :: "['a => 'b, 'a program] => 'b program" | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 14 | "rename h == extend (%(x,u::unit). h x)" | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 15 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 16 | end |