src/HOL/UNITY/Traces.ML
author paulson
Fri, 31 Jul 1998 18:46:55 +0200
changeset 5232 e5a7cdd07ea5
parent 5132 24f992a25adc
child 5240 bbcd79ef7cf2
permissions -rw-r--r--
Tidied; uses records
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     1
(*  Title:      HOL/UNITY/Traces
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     2
    ID:         $Id$
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     4
    Copyright   1998  University of Cambridge
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     5
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     6
Definitions of
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     7
  * traces: the possible execution traces
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     8
  * reachable: the set of reachable states
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
     9
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    10
*)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    11
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    12
(****
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    13
Now simulate the inductive definition (illegal due to paired arguments)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    14
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    15
inductive "reachable(Init,Acts)"
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    16
  intrs 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    17
    Init  "s: Init ==> s : reachable(Init,Acts)"
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    18
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    19
    Acts  "[| act: Acts;  s : reachable(Init,Acts);  (s,s'): act |]
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    20
	   ==> s' : reachable(Init,Acts)"
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    22
This amounts to an equivalence proof for the definition actually used, 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    23
****)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    24
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    26
(** reachable: Deriving the Introduction rules **)
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    28
Goal "s: Init ==> s : reachable(Init,Acts)";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    29
by (simp_tac (simpset() addsimps [reachable_def]) 1);
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    30
by (blast_tac (claset() addIs traces.intrs) 1);
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    31
qed "reachable_Init";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    34
Goal "[| act: Acts;  s : reachable(Init,Acts) |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    35
\     ==> (s,s'): act --> s' : reachable(Init,Acts)";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    36
by (asm_full_simp_tac (simpset() addsimps [reachable_def]) 1);
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    37
by (etac exE 1);
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5132
diff changeset
    38
by (etac traces.induct 1);
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    39
by (ALLGOALS Asm_simp_tac);
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    40
by (ALLGOALS (blast_tac (claset() addIs traces.intrs)));
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    41
qed_spec_mp "reachable_Acts";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    42
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    43
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    44
val major::prems = 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    45
Goalw [reachable_def] 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    46
  "[| za : reachable(Init,Acts);  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    47
\     !!s. s : Init ==> P s;      \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    48
\     !!act s s'.  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    49
\        [| act : Acts; s : reachable(Init,Acts); P s; (s, s') : act |]  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    50
\        ==> P s' |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    51
\  ==> P za";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    52
by (cut_facts_tac [major] 1);
5132
24f992a25adc isatool expandshort;
wenzelm
parents: 5111
diff changeset
    53
by Auto_tac;
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5132
diff changeset
    54
by (etac traces.induct 1);
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    55
by (ALLGOALS (blast_tac (claset() addIs prems)));
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    56
qed "reachable_induct";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    57
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    58
structure reachable = 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    59
  struct 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    60
  val Init = reachable_Init
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    61
  val Acts = reachable_Acts
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    62
  val intrs = [reachable_Init, reachable_Acts]
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    63
  val induct = reachable_induct
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    64
  end;
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
(*The strongest invariant is the set of all reachable states!*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    68
Goalw [stable_def, constrains_def]
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    69
    "[| Init<=A;  stable Acts A |] ==> reachable(Init,Acts) <= A";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    70
by (rtac subsetI 1);
5232
e5a7cdd07ea5 Tidied; uses records
paulson
parents: 5132
diff changeset
    71
by (etac reachable.induct 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    72
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    73
qed "strongest_invariant";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    74
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    75
Goal "stable Acts (reachable(Init,Acts))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    76
by (REPEAT (blast_tac (claset() addIs [stableI, constrainsI]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    77
				addIs reachable.intrs) 1));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    78
qed "stable_reachable";