| 6470 |      1 | (*  Title:      HOLCF/IOA/TrivEx.thy
 | 
|  |      2 |     ID:         $Id$
 | 
| 12218 |      3 |     Author:     Olaf Müller
 | 
| 6470 |      4 | *)
 | 
|  |      5 | 
 | 
| 17244 |      6 | header {* Trivial Abstraction Example with fairness *}
 | 
|  |      7 | 
 | 
|  |      8 | theory TrivEx2
 | 
|  |      9 | imports IOA Abstraction
 | 
|  |     10 | begin
 | 
| 6470 |     11 | 
 | 
|  |     12 | datatype action = INC
 | 
|  |     13 | 
 | 
|  |     14 | consts
 | 
|  |     15 | 
 | 
|  |     16 | C_asig   ::  "action signature"
 | 
| 17244 |     17 | C_trans  :: "(action, nat)transition set"
 | 
|  |     18 | C_ioa    :: "(action, nat)ioa"
 | 
|  |     19 | C_live_ioa :: "(action, nat)live_ioa"
 | 
| 6470 |     20 | 
 | 
|  |     21 | A_asig   :: "action signature"
 | 
| 17244 |     22 | A_trans  :: "(action, bool)transition set"
 | 
|  |     23 | A_ioa    :: "(action, bool)ioa"
 | 
|  |     24 | A_live_ioa :: "(action, bool)live_ioa"
 | 
| 6470 |     25 | 
 | 
|  |     26 | h_abs    :: "nat => bool"
 | 
|  |     27 | 
 | 
|  |     28 | defs
 | 
|  |     29 | 
 | 
| 17244 |     30 | C_asig_def:
 | 
| 6470 |     31 |   "C_asig == ({},{INC},{})"
 | 
|  |     32 | 
 | 
| 17244 |     33 | C_trans_def: "C_trans ==
 | 
|  |     34 |  {tr. let s = fst(tr);
 | 
|  |     35 |           t = snd(snd(tr))
 | 
|  |     36 |       in case fst(snd(tr))
 | 
|  |     37 |       of
 | 
| 6470 |     38 |       INC       => t = Suc(s)}"
 | 
|  |     39 | 
 | 
| 17244 |     40 | C_ioa_def: "C_ioa ==
 | 
| 6470 |     41 |  (C_asig, {0}, C_trans,{},{})"
 | 
|  |     42 | 
 | 
| 17244 |     43 | C_live_ioa_def:
 | 
| 6470 |     44 |   "C_live_ioa == (C_ioa, WF C_ioa {INC})"
 | 
|  |     45 | 
 | 
| 17244 |     46 | A_asig_def:
 | 
| 6470 |     47 |   "A_asig == ({},{INC},{})"
 | 
|  |     48 | 
 | 
| 17244 |     49 | A_trans_def: "A_trans ==
 | 
|  |     50 |  {tr. let s = fst(tr);
 | 
|  |     51 |           t = snd(snd(tr))
 | 
|  |     52 |       in case fst(snd(tr))
 | 
|  |     53 |       of
 | 
| 6470 |     54 |       INC       => t = True}"
 | 
|  |     55 | 
 | 
| 17244 |     56 | A_ioa_def: "A_ioa ==
 | 
| 6470 |     57 |  (A_asig, {False}, A_trans,{},{})"
 | 
|  |     58 | 
 | 
| 17244 |     59 | A_live_ioa_def:
 | 
| 6470 |     60 |   "A_live_ioa == (A_ioa, WF A_ioa {INC})"
 | 
|  |     61 | 
 | 
|  |     62 | 
 | 
| 17244 |     63 | h_abs_def:
 | 
| 6470 |     64 |   "h_abs n == n~=0"
 | 
|  |     65 | 
 | 
| 17244 |     66 | axioms
 | 
| 6470 |     67 | 
 | 
| 17244 |     68 | MC_result:
 | 
| 6470 |     69 |   "validLIOA (A_ioa,WF A_ioa {INC}) (<>[] <%(b,a,c). b>)"
 | 
|  |     70 | 
 | 
| 19740 |     71 | 
 | 
|  |     72 | lemma h_abs_is_abstraction:
 | 
|  |     73 | "is_abstraction h_abs C_ioa A_ioa"
 | 
|  |     74 | apply (unfold is_abstraction_def)
 | 
|  |     75 | apply (rule conjI)
 | 
|  |     76 | txt {* start states *}
 | 
|  |     77 | apply (simp (no_asm) add: h_abs_def starts_of_def C_ioa_def A_ioa_def)
 | 
|  |     78 | txt {* step case *}
 | 
|  |     79 | apply (rule allI)+
 | 
|  |     80 | apply (rule imp_conj_lemma)
 | 
|  |     81 | apply (simp (no_asm) add: trans_of_def C_ioa_def A_ioa_def C_trans_def A_trans_def)
 | 
|  |     82 | apply (induct_tac "a")
 | 
|  |     83 | apply (simp (no_asm) add: h_abs_def)
 | 
|  |     84 | done
 | 
|  |     85 | 
 | 
|  |     86 | 
 | 
|  |     87 | lemma Enabled_implication:
 | 
|  |     88 |     "!!s. Enabled A_ioa {INC} (h_abs s) ==> Enabled C_ioa {INC} s"
 | 
|  |     89 |   apply (unfold Enabled_def enabled_def h_abs_def A_ioa_def C_ioa_def A_trans_def
 | 
|  |     90 |     C_trans_def trans_of_def)
 | 
|  |     91 |   apply auto
 | 
|  |     92 |   done
 | 
|  |     93 | 
 | 
|  |     94 | 
 | 
|  |     95 | lemma h_abs_is_liveabstraction:
 | 
|  |     96 | "is_live_abstraction h_abs (C_ioa, WF C_ioa {INC}) (A_ioa, WF A_ioa {INC})"
 | 
|  |     97 | apply (unfold is_live_abstraction_def)
 | 
|  |     98 | apply auto
 | 
|  |     99 | txt {* is_abstraction *}
 | 
|  |    100 | apply (rule h_abs_is_abstraction)
 | 
|  |    101 | txt {* temp_weakening *}
 | 
|  |    102 | apply (tactic "abstraction_tac 1")
 | 
|  |    103 | apply (erule Enabled_implication)
 | 
|  |    104 | done
 | 
|  |    105 | 
 | 
|  |    106 | 
 | 
|  |    107 | lemma TrivEx2_abstraction:
 | 
|  |    108 |   "validLIOA C_live_ioa (<>[] <%(n,a,m). n~=0>)"
 | 
|  |    109 | apply (unfold C_live_ioa_def)
 | 
|  |    110 | apply (rule AbsRuleT2)
 | 
|  |    111 | apply (rule h_abs_is_liveabstraction)
 | 
|  |    112 | apply (rule MC_result)
 | 
|  |    113 | apply (tactic "abstraction_tac 1")
 | 
|  |    114 | apply (simp add: h_abs_def)
 | 
|  |    115 | done
 | 
| 17244 |    116 | 
 | 
|  |    117 | end
 |