src/HOL/UNITY/Detects.ML
author wenzelm
Fri, 09 Nov 2001 00:09:47 +0100
changeset 12114 a8e860c86252
parent 8334 7896bcbd8641
permissions -rw-r--r--
eliminated old "symbols" syntax, use "xsymbols" instead;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8334
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Detects
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     2
    ID:         $Id$
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     3
    Author:     Tanja Vos, Cambridge University Computer Laboratory
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     4
    Copyright   2000  University of Cambridge
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     5
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     6
Detects definition (Section 3.8 of Chandy & Misra) using LeadsTo
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     7
*)
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     8
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
     9
(* Corollary from Sectiom 3.6.4 *)
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    10
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    11
Goal "F: A LeadsTo B ==> F : Always (-((FP F) Int A Int -B))";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    12
by (rtac LeadsTo_empty 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    13
by (subgoal_tac "F : (FP F Int A Int - B) LeadsTo (B Int (FP F Int -B))" 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    14
by (subgoal_tac "(FP F Int A Int - B) = (A Int (FP F Int -B))" 2);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    15
by (subgoal_tac "(B Int (FP F Int -B)) = {}" 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    16
by Auto_tac;
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    17
by (blast_tac (claset() addIs [PSP_Stable, stable_imp_Stable, 
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    18
			       stable_FP_Int]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    19
qed "Always_at_FP";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    20
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    21
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    22
Goalw [Detects_def, Int_def]
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    23
     "[| F : A Detects B; F : B Detects C |] ==> F : A Detects C";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    24
by (Simp_tac 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    25
by Safe_tac;
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    26
by (rtac LeadsTo_Trans 2);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    27
by Auto_tac;
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    28
by (subgoal_tac "F : Always ((-A Un B) Int (-B Un C))" 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    29
by (simp_tac (simpset() addsimps [Always_Int_distrib]) 2);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    30
by Auto_tac;
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    31
by (blast_tac (claset() addIs [Always_weaken]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    32
qed "Detects_Trans";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    33
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    34
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    35
Goalw [Detects_def] "F : A Detects A";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    36
by (simp_tac (simpset() addsimps [Un_commute, Compl_partition, 
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    37
				  subset_imp_LeadsTo]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    38
qed "Detects_refl";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    39
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    40
Goalw [Equality_def] "(A<==>B) = (A Int B) Un (-A Int -B)";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    41
by (Blast_tac 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    42
qed "Detects_eq_Un";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    43
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    44
(*Not quite antisymmetry: sets A and B agree in all reachable states *)
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    45
Goalw [Detects_def, Equality_def]
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    46
     "[| F : A Detects B;  F : B Detects A|] ==> F : Always (A <==> B)";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    47
by (asm_full_simp_tac (simpset() addsimps [Always_Int_I, Un_commute]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    48
qed "Detects_antisym";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    49
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    50
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    51
(* Theorem from Section 3.8 *)
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    52
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    53
Goalw [Detects_def, Equality_def]
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    54
     "F : A Detects B ==> F : Always ((-(FP F)) Un (A <==> B))";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    55
by (simp_tac (simpset() addsimps [Un_Int_distrib, Always_Int_distrib]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    56
by (blast_tac (claset() addDs [Always_at_FP]
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    57
			addIs [Always_weaken]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    58
qed "Detects_Always";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    59
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    60
(* Theorem from exercise 11.1 Section 11.3.1 *)
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    61
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    62
Goalw [Detects_def, Equality_def]
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    63
     "F : A Detects B ==> F : UNIV LeadsTo (A <==> B)";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    64
by (res_inst_tac [("B", "B")]  LeadsTo_Diff 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    65
by (blast_tac (claset() addIs [Always_LeadsTo_weaken]) 2);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    66
by (blast_tac (claset() addIs [Always_LeadsToI, subset_imp_LeadsTo]) 1);
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    67
qed "Detects_Imp_LeadstoEQ";
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    68
7896bcbd8641 Added Tanja's Detects and Reachability theories. Also
paulson
parents:
diff changeset
    69