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