author | kleing |
Sat, 08 Apr 2006 15:24:21 +0200 | |
changeset 19360 | f47412f922ab |
parent 17233 | 41eee2e7b465 |
permissions | -rw-r--r-- |
4559 | 1 |
(* Title: HOLCF/IOA/meta_theory/TL.ML |
2 |
ID: $Id$ |
|
19360
f47412f922ab
converted Müller to Mueller to make smlnj 110.58 work
kleing
parents:
17233
diff
changeset
|
3 |
Author: Olaf Mueller |
4559 | 4 |
*) |
5 |
||
6 |
||
5068 | 7 |
Goal "[] <> (.~ P) = (.~ <> [] P)"; |
5132 | 8 |
by (rtac ext 1); |
4559 | 9 |
by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1); |
5132 | 10 |
by Auto_tac; |
5976 | 11 |
qed"simple"; |
4559 | 12 |
|
5068 | 13 |
Goal "nil |= [] P"; |
4559 | 14 |
by (asm_full_simp_tac (simpset() addsimps [satisfies_def, |
15 |
Box_def,tsuffix_def,suffix_def,nil_is_Conc])1); |
|
16 |
qed"Boxnil"; |
|
17 |
||
5068 | 18 |
Goal "~(nil |= <> P)"; |
4559 | 19 |
by (simp_tac (simpset() addsimps [Diamond_def,satisfies_def,NOT_def])1); |
20 |
by (cut_inst_tac [] Boxnil 1); |
|
21 |
by (asm_full_simp_tac (simpset() addsimps [satisfies_def])1); |
|
22 |
qed"Diamondnil"; |
|
23 |
||
5068 | 24 |
Goal "(<> F) s = (? s2. tsuffix s2 s & F s2)"; |
4559 | 25 |
by (simp_tac (simpset() addsimps [Diamond_def,NOT_def,Box_def])1); |
26 |
qed"Diamond_def2"; |
|
27 |
||
28 |
||
29 |
||
30 |
section "TLA Axiomatization by Merz"; |
|
31 |
||
32 |
(* ---------------------------------------------------------------- *) |
|
33 |
(* TLA Axiomatization by Merz *) |
|
34 |
(* ---------------------------------------------------------------- *) |
|
35 |
||
5068 | 36 |
Goal "suffix s s"; |
4559 | 37 |
by (simp_tac (simpset() addsimps [suffix_def])1); |
38 |
by (res_inst_tac [("x","nil")] exI 1); |
|
5132 | 39 |
by Auto_tac; |
4559 | 40 |
qed"suffix_refl"; |
41 |
||
5068 | 42 |
Goal "s~=UU & s~=nil --> (s |= [] F .--> F)"; |
4559 | 43 |
by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1); |
44 |
by (REPEAT (rtac impI 1)); |
|
45 |
by (eres_inst_tac [("x","s")] allE 1); |
|
46 |
by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_refl])1); |
|
47 |
qed"reflT"; |
|
48 |
||
49 |
||
6161 | 50 |
Goal "[| suffix y x ; suffix z y |] ==> suffix z x"; |
4559 | 51 |
by (asm_full_simp_tac (simpset() addsimps [suffix_def])1); |
5132 | 52 |
by Auto_tac; |
4559 | 53 |
by (res_inst_tac [("x","s1 @@ s1a")] exI 1); |
5132 | 54 |
by Auto_tac; |
4559 | 55 |
by (simp_tac (simpset() addsimps [Conc_assoc]) 1); |
56 |
qed"suffix_trans"; |
|
57 |
||
5068 | 58 |
Goal "s |= [] F .--> [] [] F"; |
4559 | 59 |
by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def,tsuffix_def])1); |
5132 | 60 |
by Auto_tac; |
61 |
by (dtac suffix_trans 1); |
|
62 |
by (assume_tac 1); |
|
4559 | 63 |
by (eres_inst_tac [("x","s2a")] allE 1); |
5132 | 64 |
by Auto_tac; |
4559 | 65 |
qed"transT"; |
66 |
||
67 |
||
5068 | 68 |
Goal "s |= [] (F .--> G) .--> [] F .--> [] G"; |
4559 | 69 |
by (simp_tac (simpset() addsimps [satisfies_def,IMPLIES_def,Box_def])1); |
70 |
qed"normalT"; |
|
71 |
||
72 |
||
73 |
section "TLA Rules by Lamport"; |
|
74 |
||
75 |
(* ---------------------------------------------------------------- *) |
|
76 |
(* TLA Rules by Lamport *) |
|
77 |
(* ---------------------------------------------------------------- *) |
|
78 |
||
6161 | 79 |
Goal "validT P ==> validT ([] P)"; |
4559 | 80 |
by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Box_def,tsuffix_def])1); |
81 |
qed"STL1a"; |
|
82 |
||
6161 | 83 |
Goal "valid P ==> validT (Init P)"; |
4559 | 84 |
by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,Init_def])1); |
85 |
qed"STL1b"; |
|
86 |
||
6161 | 87 |
Goal "valid P ==> validT ([] (Init P))"; |
5132 | 88 |
by (rtac STL1a 1); |
89 |
by (etac STL1b 1); |
|
4559 | 90 |
qed"STL1"; |
91 |
||
92 |
(* Note that unlift and HD is not at all used !!! *) |
|
6161 | 93 |
Goal "valid (P .--> Q) ==> validT ([] (Init P) .--> [] (Init Q))"; |
4559 | 94 |
by (asm_full_simp_tac (simpset() addsimps [valid_def,validT_def,satisfies_def,IMPLIES_def,Box_def,Init_def])1); |
95 |
qed"STL4"; |
|
96 |
||
97 |
||
98 |
section "LTL Axioms by Manna/Pnueli"; |
|
99 |
||
100 |
(* ---------------------------------------------------------------- *) |
|
101 |
(* LTL Axioms by Manna/Pnueli *) |
|
102 |
(* ---------------------------------------------------------------- *) |
|
103 |
||
104 |
||
5068 | 105 |
Goalw [tsuffix_def,suffix_def] |
10835 | 106 |
"s~=UU & s~=nil --> tsuffix s2 (TL$s) --> tsuffix s2 s"; |
5132 | 107 |
by Auto_tac; |
4559 | 108 |
by (Seq_case_simp_tac "s" 1); |
109 |
by (res_inst_tac [("x","a>>s1")] exI 1); |
|
5132 | 110 |
by Auto_tac; |
4559 | 111 |
qed_spec_mp"tsuffix_TL"; |
112 |
||
113 |
val tsuffix_TL2 = conjI RS tsuffix_TL; |
|
114 |
||
4833 | 115 |
Delsplits[split_if]; |
5068 | 116 |
Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def,AND_def,Box_def] |
4559 | 117 |
"s~=UU & s~=nil --> (s |= [] F .--> (F .& (Next ([] F))))"; |
5132 | 118 |
by Auto_tac; |
4559 | 119 |
(* []F .--> F *) |
120 |
by (eres_inst_tac [("x","s")] allE 1); |
|
121 |
by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_refl])1); |
|
122 |
(* []F .--> Next [] F *) |
|
4833 | 123 |
by (asm_full_simp_tac (simpset() addsplits [split_if]) 1); |
5132 | 124 |
by Auto_tac; |
125 |
by (dtac tsuffix_TL2 1); |
|
4559 | 126 |
by (REPEAT (atac 1)); |
5132 | 127 |
by Auto_tac; |
4559 | 128 |
qed"LTL1"; |
4833 | 129 |
Addsplits[split_if]; |
4559 | 130 |
|
131 |
||
5068 | 132 |
Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] |
4559 | 133 |
"s |= .~ (Next F) .--> (Next (.~ F))"; |
4833 | 134 |
by (Asm_full_simp_tac 1); |
4559 | 135 |
qed"LTL2a"; |
136 |
||
5068 | 137 |
Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] |
4559 | 138 |
"s |= (Next (.~ F)) .--> (.~ (Next F))"; |
4833 | 139 |
by (Asm_full_simp_tac 1); |
4559 | 140 |
qed"LTL2b"; |
141 |
||
5068 | 142 |
Goalw [Next_def,satisfies_def,NOT_def,IMPLIES_def] |
4559 | 143 |
"ex |= (Next (F .--> G)) .--> (Next F) .--> (Next G)"; |
4833 | 144 |
by (Asm_full_simp_tac 1); |
4559 | 145 |
qed"LTL3"; |
146 |
||
5068 | 147 |
Goalw [Next_def,satisfies_def,Box_def,NOT_def,IMPLIES_def] |
4559 | 148 |
"s |= [] (F .--> Next F) .--> F .--> []F"; |
4833 | 149 |
by (Asm_full_simp_tac 1); |
5132 | 150 |
by Auto_tac; |
4559 | 151 |
|
152 |
by (asm_full_simp_tac (simpset() addsimps [tsuffix_def,suffix_def])1); |
|
5132 | 153 |
by Auto_tac; |
4559 | 154 |
|
155 |
||
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
156 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
157 |
|
6161 | 158 |
Goal "[| validT (P .--> Q); validT P |] ==> validT Q"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
159 |
by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,IMPLIES_def])1); |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
160 |
qed"ModusPonens"; |
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
161 |
|
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
162 |
(* works only if validT is defined without restriction to s~=UU, s~=nil *) |
6161 | 163 |
Goal "validT P ==> validT (Next P)"; |
4577
674b0b354feb
added thms wrt weakening and strengthening in Abstraction;
mueller
parents:
4559
diff
changeset
|
164 |
by (asm_full_simp_tac (simpset() addsimps [validT_def,satisfies_def,Next_def])1); |
4681 | 165 |
(* qed"NextTauto"; *) |