| author | haftmann | 
| Tue, 29 Aug 2006 14:31:15 +0200 | |
| changeset 20428 | 67fa1c6ba89e | 
| parent 19764 | 372065f34795 | 
| child 25131 | 2c8caac48ade | 
| permissions | -rw-r--r-- | 
| 7299 | 1 | |
| 17244 | 2 | (* $Id$ *) | 
| 3 | ||
| 4 | theory Ring3 | |
| 5 | imports MuIOAOracle | |
| 6 | begin | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 7 | |
| 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 8 | datatype token = Leader | id0 | id1 | id2 | id3 | id4 | 
| 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 9 | |
| 7299 | 10 | datatype Comm = | 
| 11 | Zero_One token | One_Two token | Two_Zero token | | |
| 12 | Leader_Zero | Leader_One | Leader_Two | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 13 | |
| 7299 | 14 | constdefs | 
| 17244 | 15 | Greater :: "[token, token] => bool" | 
| 7299 | 16 | "Greater x y == | 
| 17 | (x~=Leader & x~=id0 & y=id0) | (x=id4 & y~=id4 & y~=Leader) | | |
| 18 | (x=id2 & y=id1) | (x=id3 & y=id1) | (x=id3 & y=id2)" | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 19 | |
| 7299 | 20 | |
| 21 | (*the ring is the composition of aut0, aut1 and aut2*) | |
| 22 | ||
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 23 | automaton aut0 = | 
| 7299 | 24 | signature | 
| 25 | actions Comm | |
| 26 | inputs "Two_Zero t" | |
| 27 | outputs "Zero_One t","Leader_Zero" | |
| 28 | states | |
| 29 | idf :: token | |
| 30 | initially "idf=id0 | idf = id3" | |
| 31 | transitions | |
| 32 | "Two_Zero t" | |
| 33 | transrel | |
| 34 | "if (t=id0 | t=id3) then (idf'=Leader) else | |
| 35 | (if (Greater t idf) then (idf'=t) else (idf'=idf))" | |
| 36 | "Zero_One t" | |
| 37 | pre "t=idf" | |
| 38 | "Leader_Zero" | |
| 39 | pre "idf=Leader" | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 40 | |
| 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 41 | automaton aut1 = | 
| 7299 | 42 | signature | 
| 43 | actions Comm | |
| 44 | inputs "Zero_One t" | |
| 45 | outputs "One_Two t","Leader_One" | |
| 46 | states | |
| 47 | idf :: token | |
| 48 | initially "idf=id1 | idf=id4" | |
| 49 | transitions | |
| 50 | "Zero_One t" | |
| 51 | transrel | |
| 52 | "if (t=id1 | t=id4) then (idf'=Leader) else | |
| 53 | (if (Greater t idf) then (idf'=t) else (idf'=idf))" | |
| 54 | "One_Two t" | |
| 55 | pre "t=idf" | |
| 56 | "Leader_One" | |
| 57 | pre "idf=Leader" | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 58 | |
| 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 59 | automaton aut2 = | 
| 7299 | 60 | signature | 
| 61 | actions Comm | |
| 62 | inputs "One_Two t" | |
| 63 | outputs "Two_Zero t","Leader_Two" | |
| 64 | states | |
| 65 | idf :: token | |
| 66 | initially "idf=id2" | |
| 67 | transitions | |
| 68 | "One_Two t" | |
| 69 | transrel | |
| 70 | "if (t=id2) then (idf'=Leader) else | |
| 71 | (if (Greater t idf) then (idf'=t) else (idf'=idf))" | |
| 72 | "Two_Zero t" | |
| 73 | pre "t=idf" | |
| 74 | "Leader_Two" | |
| 75 | pre "idf=Leader" | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 76 | |
| 7299 | 77 | automaton ring = compose aut0, aut1, aut2 | 
| 78 | ||
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 79 | |
| 7299 | 80 | (*one_leader expresses property that at most one component declares | 
| 81 | itself to leader*) | |
| 82 | ||
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 83 | automaton one_leader = | 
| 7299 | 84 | signature | 
| 85 | actions Comm | |
| 17244 | 86 | outputs "Zero_One t","One_Two t","Two_Zero t","Leader_Zero","Leader_One","Leader_Two" | 
| 7299 | 87 | states | 
| 88 | leader :: token | |
| 89 | initially "leader=Leader" | |
| 90 | transitions | |
| 91 | "Zero_One t" | |
| 92 | pre "True" | |
| 93 | "One_Two t" | |
| 94 | pre "True" | |
| 95 | "Two_Zero t" | |
| 17244 | 96 | pre "True" | 
| 7299 | 97 | "Leader_Zero" | 
| 98 | pre "leader=id0 | leader=Leader" | |
| 99 | post leader:="id0" | |
| 100 | "Leader_One" | |
| 101 | pre "leader=id1 | leader=Leader" | |
| 17244 | 102 | post leader:="id1" | 
| 7299 | 103 | "Leader_Two" | 
| 104 | pre "leader=id2 | leader=Leader" | |
| 105 | post leader:="id2" | |
| 106 | ||
| 19764 | 107 | lemmas aut_simps = | 
| 108 | Greater_def one_leader_def | |
| 109 | one_leader_asig_def one_leader_trans_def one_leader_initial_def | |
| 110 | ring_def aut0_asig_def aut0_trans_def aut0_initial_def aut0_def | |
| 111 | aut1_asig_def aut1_trans_def aut1_initial_def aut1_def | |
| 112 | aut2_asig_def aut2_trans_def aut2_initial_def aut2_def | |
| 113 | ||
| 114 | (* property to prove: at most one (of 3) members of the ring will | |
| 115 | declare itself to leader *) | |
| 116 | lemma at_most_one_leader: "ring =<| one_leader" | |
| 117 | apply (tactic {* is_sim_tac (thms "aut_simps") 1 *})
 | |
| 118 | apply auto | |
| 119 | done | |
| 6471 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 120 | |
| 
08d12ef5fc19
added translation from IOA to mucalculus and corresponding modelchecker examples;
 mueller parents: diff
changeset | 121 | end |