src/HOL/UNITY/Comp.ML
author paulson
Fri, 22 Oct 1999 18:35:20 +0200
changeset 7915 c7fd7eb3b0ef
parent 7878 43b03d412b82
child 7947 b999c1ab9327
permissions -rw-r--r--
ALMOST working version: LocalTo results commented out
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Comp.thy
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     5
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     6
Composition
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     7
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     8
From Chandy and Sanders, "Reasoning About Program Composition"
a12b25c53df1 composition theory
paulson
parents:
diff changeset
     9
*)
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    10
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    11
(*** component ***)
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    12
6738
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    13
Goalw [component_def]
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    14
     "H <= F | H <= G ==> H <= (F Join G)";
7361
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    15
by Auto_tac;
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    16
by (res_inst_tac [("x", "G Join Ga")] exI 1);
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    17
by (res_inst_tac [("x", "G Join F")] exI 2);
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    18
by (auto_tac (claset(), simpset() addsimps Join_ac));
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    19
qed "componentI";
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    20
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7340
diff changeset
    21
Goalw [component_def]
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    22
     "(F <= G) = (Init G <= Init F & Acts F <= Acts G)";
6738
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    23
by (force_tac (claset() addSIs [exI, program_equalityI], 
7537
875754b599df working snapshot
paulson
parents: 7399
diff changeset
    24
	       simpset()) 1);
6738
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    25
qed "component_eq_subset";
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    26
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    27
Goalw [component_def] "SKIP <= F";
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5968
diff changeset
    28
by (force_tac (claset() addIs [Join_SKIP_left], simpset()) 1);
5612
e981ca6f7332 Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents: 5597
diff changeset
    29
qed "component_SKIP";
e981ca6f7332 Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents: 5597
diff changeset
    30
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    31
Goalw [component_def] "F <= (F :: 'a program)";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6018
diff changeset
    32
by (blast_tac (claset() addIs [Join_SKIP_right]) 1);
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    33
qed "component_refl";
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    34
5612
e981ca6f7332 Finished proofs to end of section 5.1 of Chandy and Sanders
paulson
parents: 5597
diff changeset
    35
AddIffs [component_SKIP, component_refl];
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    36
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    37
Goal "F <= SKIP ==> F = SKIP";
6833
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    38
by (auto_tac (claset() addSIs [program_equalityI],
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    39
	      simpset() addsimps [component_eq_subset]));
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    40
qed "SKIP_minimal";
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    41
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    42
Goalw [component_def] "F <= (F Join G)";
5968
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    43
by (Blast_tac 1);
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    44
qed "component_Join1";
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    45
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    46
Goalw [component_def] "G <= (F Join G)";
5968
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    47
by (simp_tac (simpset() addsimps [Join_commute]) 1);
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    48
by (Blast_tac 1);
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    49
qed "component_Join2";
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    50
7915
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    51
Goal "F<=G ==> F Join G = G";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    52
by (auto_tac (claset(), simpset() addsimps [component_def, Join_left_absorb]));
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    53
qed "Join_absorb1";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    54
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    55
Goal "G<=F ==> F Join G = F";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    56
by (auto_tac (claset(), simpset() addsimps Join_ac@[component_def]));
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    57
qed "Join_absorb2";
c7fd7eb3b0ef ALMOST working version: LocalTo results commented out
paulson
parents: 7878
diff changeset
    58
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    59
Goalw [component_def] "i : I ==> (F i) <= (JN i:I. (F i))";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6018
diff changeset
    60
by (blast_tac (claset() addIs [JN_absorb]) 1);
5968
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    61
qed "component_JN";
06f9dbfff032 renamed vars
paulson
parents: 5804
diff changeset
    62
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    63
Goalw [component_def] "[| F <= G; G <= H |] ==> F <= (H :: 'a program)";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6018
diff changeset
    64
by (blast_tac (claset() addIs [Join_assoc RS sym]) 1);
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    65
qed "component_trans";
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    66
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    67
Goal "[| F <= G; G <= F |] ==> F = (G :: 'a program)";
6738
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    68
by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1);
06189132c67b component_eq_subset: a neat characterization of "component"
paulson
parents: 6703
diff changeset
    69
by (blast_tac (claset() addSIs [program_equalityI]) 1);
6703
paulson
parents: 6646
diff changeset
    70
qed "component_antisym";
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    71
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5706
diff changeset
    72
Goalw [component_def]
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7594
diff changeset
    73
      "F <= H = (EX G. F Join G = H & Disjoint UNIV F G)";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6018
diff changeset
    74
by (blast_tac (claset() addSIs [Diff_Disjoint, Join_Diff2]) 1);
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5706
diff changeset
    75
qed "component_eq";
5597
a12b25c53df1 composition theory
paulson
parents:
diff changeset
    76
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    77
Goal "((F Join G) <= H) = (F <= H & G <= H)";
7537
875754b599df working snapshot
paulson
parents: 7399
diff changeset
    78
by (simp_tac (simpset() addsimps [component_eq_subset]) 1);
6833
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    79
by (Blast_tac 1);
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    80
qed "Join_component_iff";
15d6c121d75f many new guarantees laws
paulson
parents: 6822
diff changeset
    81
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    82
Goal "[| F <= G; G : A co B |] ==> F : A co B";
7386
1048bc161c05 a new theorem
paulson
parents: 7361
diff changeset
    83
by (auto_tac (claset(), 
1048bc161c05 a new theorem
paulson
parents: 7361
diff changeset
    84
	      simpset() addsimps [constrains_def, component_eq_subset]));
1048bc161c05 a new theorem
paulson
parents: 7361
diff changeset
    85
qed "component_constrains";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6018
diff changeset
    86
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7386
diff changeset
    87
bind_thm ("program_less_le", strict_component_def RS meta_eq_to_obj_eq);