| author | wenzelm | 
| Wed, 29 Jun 2005 15:13:25 +0200 | |
| changeset 16595 | e64fb2cf50cb | 
| parent 14981 | e73f8140af78 | 
| child 17233 | 41eee2e7b465 | 
| permissions | -rw-r--r-- | 
| 4559 | 1 | (* Title: HOLCF/IOA/meta_theory/TLS.ML | 
| 2 | ID: $Id$ | |
| 12218 | 3 | Author: Olaf Müller | 
| 4559 | 4 | |
| 12218 | 5 | Temporal Logic of Steps -- tailored for I/O automata. | 
| 5976 | 6 | *) | 
| 4559 | 7 | |
| 4815 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 8 | (* global changes to simpset() and claset(), repeated from Traces.ML *) | 
| 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 9 | Delsimps (ex_simps @ all_simps); | 
| 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 10 | Delsimps [split_paired_Ex]; | 
| 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 11 | Addsimps [Let_def]; | 
| 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 12 | claset_ref() := claset() delSWrapper "split_all_tac"; | 
| 
b8a32ef742d9
removed split_all_tac from claset() globally within IOA
 oheimb parents: 
4577diff
changeset | 13 | |
| 4559 | 14 | |
| 15 | (* ---------------------------------------------------------------- *) | |
| 16 | (* ex2seqC *) | |
| 17 | (* ---------------------------------------------------------------- *) | |
| 18 | ||
| 5068 | 19 | Goal "ex2seqC = (LAM ex. (%s. case ex of \ | 
| 4559 | 20 | \ nil => (s,None,s)>>nil \ | 
| 21 | \ | x##xs => (flift1 (%pr. \ | |
| 10835 | 22 | \ (s,Some (fst pr), snd pr)>> (ex2seqC$xs) (snd pr)) \ | 
| 23 | \ $x) \ | |
| 4559 | 24 | \ ))"; | 
| 25 | by (rtac trans 1); | |
| 26 | by (rtac fix_eq2 1); | |
| 27 | by (rtac ex2seqC_def 1); | |
| 28 | by (rtac beta_cfun 1); | |
| 29 | by (simp_tac (simpset() addsimps [flift1_def]) 1); | |
| 30 | qed"ex2seqC_unfold"; | |
| 31 | ||
| 10835 | 32 | Goal "(ex2seqC $UU) s=UU"; | 
| 4559 | 33 | by (stac ex2seqC_unfold 1); | 
| 34 | by (Simp_tac 1); | |
| 35 | qed"ex2seqC_UU"; | |
| 36 | ||
| 10835 | 37 | Goal "(ex2seqC $nil) s = (s,None,s)>>nil"; | 
| 4559 | 38 | by (stac ex2seqC_unfold 1); | 
| 39 | by (Simp_tac 1); | |
| 40 | qed"ex2seqC_nil"; | |
| 41 | ||
| 10835 | 42 | Goal "(ex2seqC $((a,t)>>xs)) s = \ | 
| 43 | \ (s,Some a,t)>> ((ex2seqC$xs) t)"; | |
| 4559 | 44 | by (rtac trans 1); | 
| 45 | by (stac ex2seqC_unfold 1); | |
| 7229 
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
 wenzelm parents: 
6161diff
changeset | 46 | by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1); | 
| 
6773ba0c36d5
renamed Cons to Consq in order to avoid clash with List.Cons;
 wenzelm parents: 
6161diff
changeset | 47 | by (asm_full_simp_tac (simpset() addsimps [Consq_def,flift1_def]) 1); | 
| 4559 | 48 | qed"ex2seqC_cons"; | 
| 49 | ||
| 50 | Addsimps [ex2seqC_UU,ex2seqC_nil,ex2seqC_cons]; | |
| 51 | ||
| 52 | ||
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 53 | |
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 54 | Addsimps [mkfin_UU,mkfin_nil,mkfin_cons]; | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 55 | |
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 56 | Goal "ex2seq (s, UU) = (s,None,s)>>nil"; | 
| 4559 | 57 | by (simp_tac (simpset() addsimps [ex2seq_def]) 1); | 
| 58 | qed"ex2seq_UU"; | |
| 59 | ||
| 5068 | 60 | Goal "ex2seq (s, nil) = (s,None,s)>>nil"; | 
| 4559 | 61 | by (simp_tac (simpset() addsimps [ex2seq_def]) 1); | 
| 62 | qed"ex2seq_nil"; | |
| 63 | ||
| 5068 | 64 | Goal "ex2seq (s, (a,t)>>ex) = (s,Some a,t) >> ex2seq (t, ex)"; | 
| 4559 | 65 | by (simp_tac (simpset() addsimps [ex2seq_def]) 1); | 
| 66 | qed"ex2seq_cons"; | |
| 67 | ||
| 68 | Delsimps [ex2seqC_UU,ex2seqC_nil,ex2seqC_cons]; | |
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 69 | Addsimps [ex2seq_UU,ex2seq_nil, ex2seq_cons]; | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 70 | |
| 4559 | 71 | |
| 5068 | 72 | Goal "ex2seq exec ~= UU & ex2seq exec ~= nil"; | 
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 73 | by (pair_tac "exec" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 74 | by (Seq_case_simp_tac "y" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 75 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 76 | qed"ex2seq_nUUnnil"; | 
| 4559 | 77 | |
| 78 | ||
| 79 | (* ----------------------------------------------------------- *) | |
| 80 | (* Interface TL -- TLS *) | |
| 81 | (* ---------------------------------------------------------- *) | |
| 82 | ||
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 83 | |
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 84 | (* uses the fact that in executions states overlap, which is lost in | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 85 | after the translation via ex2seq !! *) | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 86 | |
| 5068 | 87 | Goalw [Init_def,Next_def,temp_sat_def,satisfies_def,IMPLIES_def,AND_def] | 
| 6161 | 88 | "[| ! s a t. (P s) & s-a--A-> t --> (Q t) |]\ | 
| 4559 | 89 | \ ==> ex |== (Init (%(s,a,t). P s) .& Init (%(s,a,t). s -a--A-> t) \ | 
| 90 | \ .--> (Next (Init (%(s,a,t).Q s))))"; | |
| 91 | ||
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 92 | by (clarify_tac set_cs 1); | 
| 9385 | 93 | by (asm_full_simp_tac (simpset() addsplits [split_if]) 1); | 
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 94 | (* TL = UU *) | 
| 6161 | 95 | by (rtac conjI 1); | 
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 96 | by (pair_tac "ex" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 97 | by (Seq_case_simp_tac "y" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 98 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 99 | by (Seq_case_simp_tac "s" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 100 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 101 | (* TL = nil *) | 
| 6161 | 102 | by (rtac conjI 1); | 
| 4559 | 103 | by (pair_tac "ex" 1); | 
| 104 | by (Seq_case_simp_tac "y" 1); | |
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 105 | by (asm_full_simp_tac (simpset() addsimps [unlift_def])1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 106 | by (Fast_tac 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 107 | by (asm_full_simp_tac (simpset() addsimps [unlift_def])1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 108 | by (Fast_tac 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 109 | by (asm_full_simp_tac (simpset() addsimps [unlift_def])1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 110 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 111 | by (Seq_case_simp_tac "s" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 112 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 113 | (* TL =cons *) | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 114 | by (asm_full_simp_tac (simpset() addsimps [unlift_def])1); | 
| 4559 | 115 | |
| 5677 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 116 | by (pair_tac "ex" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 117 | by (Seq_case_simp_tac "y" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 118 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 119 | by (Seq_case_simp_tac "s" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 120 | by (Fast_tac 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 121 | by (Fast_tac 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 122 | by (pair_tac "a" 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 123 | by (Fast_tac 1); | 
| 
4feffde494cf
another little bug ;-) and minor changes in TLS.*;
 mueller parents: 
5068diff
changeset | 124 | qed"TL_TLS"; |