| author | paulson <lp15@cam.ac.uk> | 
| Thu, 26 Apr 2018 22:47:04 +0100 | |
| changeset 68048 | 0b4fb9fd91b1 | 
| parent 63146 | f1ecba0272f9 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 37936 | 1 | (* Title: HOL/UNITY/Detects.thy | 
| 8334 | 2 | Author: Tanja Vos, Cambridge University Computer Laboratory | 
| 3 | Copyright 2000 University of Cambridge | |
| 4 | ||
| 5 | Detects definition (Section 3.8 of Chandy & Misra) using LeadsTo | |
| 6 | *) | |
| 7 | ||
| 63146 | 8 | section\<open>The Detects Relation\<close> | 
| 13798 | 9 | |
| 16417 | 10 | theory Detects imports FP SubstAx begin | 
| 8334 | 11 | |
| 57488 | 12 | definition Detects :: "['a set, 'a set] => 'a program set" (infixl "Detects" 60) | 
| 13 | where "A Detects B = (Always (-A \<union> B)) \<inter> (B LeadsTo A)" | |
| 14 | ||
| 15 | definition Equality :: "['a set, 'a set] => 'a set" (infixl "<==>" 60) | |
| 16 | where "A <==> B = (-A \<union> B) \<inter> (A \<union> -B)" | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 17 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 18 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 19 | (* Corollary from Sectiom 3.6.4 *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 20 | |
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 21 | lemma Always_at_FP: | 
| 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 22 | "[|F \<in> A LeadsTo B; all_total F|] ==> F \<in> Always (-((FP F) \<inter> A \<inter> -B))" | 
| 61635 | 23 | supply [[simproc del: boolean_algebra_cancel_inf]] inf_compl_bot_right[simp del] | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 24 | apply (rule LeadsTo_empty) | 
| 13805 | 25 | apply (subgoal_tac "F \<in> (FP F \<inter> A \<inter> - B) LeadsTo (B \<inter> (FP F \<inter> -B))") | 
| 26 | apply (subgoal_tac [2] " (FP F \<inter> A \<inter> - B) = (A \<inter> (FP F \<inter> -B))") | |
| 27 | apply (subgoal_tac "(B \<inter> (FP F \<inter> -B)) = {}")
 | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 28 | apply auto | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 29 | apply (blast intro: PSP_Stable stable_imp_Stable stable_FP_Int) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 30 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 31 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 32 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 33 | lemma Detects_Trans: | 
| 13805 | 34 | "[| F \<in> A Detects B; F \<in> B Detects C |] ==> F \<in> A Detects C" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 35 | apply (unfold Detects_def Int_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 36 | apply (simp (no_asm)) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 37 | apply safe | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 38 | apply (rule_tac [2] LeadsTo_Trans, auto) | 
| 13805 | 39 | apply (subgoal_tac "F \<in> Always ((-A \<union> B) \<inter> (-B \<union> C))") | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 40 | apply (blast intro: Always_weaken) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 41 | apply (simp add: Always_Int_distrib) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 42 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 43 | |
| 13805 | 44 | lemma Detects_refl: "F \<in> A Detects A" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 45 | apply (unfold Detects_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 46 | apply (simp (no_asm) add: Un_commute Compl_partition subset_imp_LeadsTo) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 47 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 48 | |
| 13805 | 49 | lemma Detects_eq_Un: "(A<==>B) = (A \<inter> B) \<union> (-A \<inter> -B)" | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 50 | by (unfold Equality_def, blast) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 51 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 52 | (*Not quite antisymmetry: sets A and B agree in all reachable states *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 53 | lemma Detects_antisym: | 
| 13805 | 54 | "[| F \<in> A Detects B; F \<in> B Detects A|] ==> F \<in> Always (A <==> B)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 55 | apply (unfold Detects_def Equality_def) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 56 | apply (simp add: Always_Int_I Un_commute) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 57 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 58 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 59 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 60 | (* Theorem from Section 3.8 *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 61 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 62 | lemma Detects_Always: | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 63 | "[|F \<in> A Detects B; all_total F|] ==> F \<in> Always (-(FP F) \<union> (A <==> B))" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 64 | apply (unfold Detects_def Equality_def) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 65 | apply (simp add: Un_Int_distrib Always_Int_distrib) | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 66 | apply (blast dest: Always_at_FP intro: Always_weaken) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 67 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 68 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 69 | (* Theorem from exercise 11.1 Section 11.3.1 *) | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 70 | |
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 71 | lemma Detects_Imp_LeadstoEQ: | 
| 13805 | 72 | "F \<in> A Detects B ==> F \<in> UNIV LeadsTo (A <==> B)" | 
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 73 | apply (unfold Detects_def Equality_def) | 
| 13812 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 paulson parents: 
13805diff
changeset | 74 | apply (rule_tac B = B in LeadsTo_Diff) | 
| 13805 | 75 | apply (blast intro: Always_LeadsToI subset_imp_LeadsTo) | 
| 76 | apply (blast intro: Always_LeadsTo_weaken) | |
| 13785 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 77 | done | 
| 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 paulson parents: 
8334diff
changeset | 78 | |
| 8334 | 79 | |
| 80 | end | |
| 81 |