| author | huffman | 
| Mon, 09 Nov 2009 15:51:32 -0800 | |
| changeset 33587 | 54f98d225163 | 
| parent 25135 | 4f8176c940cf | 
| child 35174 | e15040ae75d7 | 
| permissions | -rw-r--r-- | 
| 4565 | 1 | (* Title: HOLCF/IOA/meta_theory/Simulations.thy | 
| 2 | ID: $Id$ | |
| 12218 | 3 | Author: Olaf Müller | 
| 4565 | 4 | *) | 
| 5 | ||
| 17233 | 6 | header {* Simulations in HOLCF/IOA *}
 | 
| 4565 | 7 | |
| 17233 | 8 | theory Simulations | 
| 9 | imports RefCorrectness | |
| 10 | begin | |
| 11 | ||
| 12 | defaultsort type | |
| 4565 | 13 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 14 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 15 |   is_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 16 | "is_simulation R C A = | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 17 |    ((!s:starts_of C. R``{s} Int starts_of A ~= {}) &
 | 
| 17233 | 18 | (!s s' t a. reachable C s & | 
| 4565 | 19 | s -a--C-> t & | 
| 17233 | 20 | (s,s') : R | 
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 21 | --> (? t' ex. (t,t'):R & move A ex s' a t')))" | 
| 4565 | 22 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 23 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 24 |   is_backward_simulation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 25 | "is_backward_simulation R C A = | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 26 |    ((!s:starts_of C. R``{s} <= starts_of A) &
 | 
| 17233 | 27 | (!s t t' a. reachable C s & | 
| 4565 | 28 | s -a--C-> t & | 
| 17233 | 29 | (t,t') : R | 
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 30 | --> (? ex s'. (s,s'):R & move A ex s' a t')))" | 
| 4565 | 31 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 32 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 33 |   is_forw_back_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 34 | "is_forw_back_simulation R C A = | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 35 | ((!s:starts_of C. ? S'. (s,S'):R & S'<= starts_of A) & | 
| 17233 | 36 | (!s S' t a. reachable C s & | 
| 4565 | 37 | s -a--C-> t & | 
| 17233 | 38 | (s,S') : R | 
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 39 | --> (? T'. (t,T'):R & (! t':T'. ? s':S'. ? ex. move A ex s' a t'))))" | 
| 4565 | 40 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 41 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 42 |   is_back_forw_simulation :: "[('s1 * 's2 set)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 43 | "is_back_forw_simulation R C A = | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 44 |    ((!s:starts_of C. ! S'. (s,S'):R --> S' Int starts_of A ~={}) &
 | 
| 17233 | 45 | (!s t T' a. reachable C s & | 
| 4565 | 46 | s -a--C-> t & | 
| 17233 | 47 | (t,T') : R | 
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 48 | --> (? S'. (s,S'):R & (! s':S'. ? t':T'. ? ex. move A ex s' a t'))))" | 
| 4565 | 49 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 50 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 51 |   is_history_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 52 | "is_history_relation R C A = (is_simulation R C A & | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 53 | is_ref_map (%x.(@y. (x,y):(R^-1))) A C)" | 
| 4565 | 54 | |
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 55 | definition | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 56 |   is_prophecy_relation :: "[('s1 * 's2)set,('a,'s1)ioa,('a,'s2)ioa] => bool" where
 | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 57 | "is_prophecy_relation R C A = (is_backward_simulation R C A & | 
| 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 58 | is_ref_map (%x.(@y. (x,y):(R^-1))) A C)" | 
| 17233 | 59 | |
| 19741 | 60 | |
| 61 | lemma set_non_empty: "(A~={}) = (? x. x:A)"
 | |
| 62 | apply auto | |
| 63 | done | |
| 64 | ||
| 65 | lemma Int_non_empty: "(A Int B ~= {}) = (? x. x: A & x:B)"
 | |
| 66 | apply (simp add: set_non_empty) | |
| 67 | done | |
| 68 | ||
| 69 | ||
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 70 | lemma Sim_start_convert: | 
| 19741 | 71 | "(R``{x} Int S ~= {}) = (? y. (x,y):R & y:S)"
 | 
| 72 | apply (unfold Image_def) | |
| 73 | apply (simp add: Int_non_empty) | |
| 74 | done | |
| 75 | ||
| 76 | declare Sim_start_convert [simp] | |
| 77 | ||
| 78 | ||
| 25135 
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
 wenzelm parents: 
19741diff
changeset | 79 | lemma ref_map_is_simulation: | 
| 19741 | 80 | "!! f. is_ref_map f C A ==> is_simulation {p. (snd p) = f (fst p)} C A"
 | 
| 81 | ||
| 82 | apply (unfold is_ref_map_def is_simulation_def) | |
| 83 | apply simp | |
| 84 | done | |
| 17233 | 85 | |
| 4565 | 86 | end |