| author | nipkow | 
| Wed, 29 Mar 2000 15:09:51 +0200 | |
| changeset 8604 | c99e0024050c | 
| parent 8256 | 6ba8fa2b0638 | 
| child 9403 | aad13b59b8d9 | 
| 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 |   rename_act :: "['a => 'b, ('a*'a) set] => ('b*'b) set"
 | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 13 | "rename_act h == extend_act (%(x,u::unit). h x)" | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 14 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 15 | (**OR | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 16 |       "rename_act h == %act. UN (s,s'): act.  {(h s, h s')}"
 | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 17 | "rename_act h == %act. (prod_fun h h) `` act" | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 18 | **) | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 19 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 20 | rename :: "['a => 'b, 'a program] => 'b program" | 
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 21 | "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 | 22 | |
| 
6ba8fa2b0638
Rename: theory for applying a bijection over states to a UNITY program
 paulson parents: diff
changeset | 23 | end |