| author | wenzelm | 
| Thu, 17 Jan 2002 21:05:58 +0100 | |
| changeset 12807 | 4f2983e39a59 | 
| parent 10834 | a7897aebbffc | 
| permissions | -rw-r--r-- | 
| 4776 | 1 | (* Title: HOL/UNITY/FP | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1998 University of Cambridge | |
| 5 | ||
| 6 | Fixed Point of a Program | |
| 7 | ||
| 8 | From Misra, "A Logic for Concurrent Programming", 1994 | |
| 9 | *) | |
| 10 | ||
| 5648 | 11 | Goalw [FP_Orig_def, stable_def] "F : stable (FP_Orig F Int B)"; | 
| 4776 | 12 | by (stac Int_Union2 1); | 
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
5648diff
changeset | 13 | by (blast_tac (claset() addIs [constrains_UN]) 1); | 
| 4776 | 14 | qed "stable_FP_Orig_Int"; | 
| 15 | ||
| 16 | ||
| 17 | val prems = goalw thy [FP_Orig_def, stable_def] | |
| 5648 | 18 | "(!!B. F : stable (A Int B)) ==> A <= FP_Orig F"; | 
| 4776 | 19 | by (blast_tac (claset() addIs prems) 1); | 
| 20 | qed "FP_Orig_weakest"; | |
| 21 | ||
| 22 | ||
| 5648 | 23 | Goal "F : stable (FP F Int B)"; | 
| 24 | by (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x})" 1);
 | |
| 4776 | 25 | by (Blast_tac 2); | 
| 26 | by (asm_simp_tac (simpset() addsimps [Int_insert_right]) 1); | |
| 27 | by (rewrite_goals_tac [FP_def, stable_def]); | |
| 8334 
7896bcbd8641
Added Tanja's Detects and Reachability theories.  Also
 paulson parents: 
5648diff
changeset | 28 | by (rtac constrains_UN 1); | 
| 4776 | 29 | by (Simp_tac 1); | 
| 30 | qed "stable_FP_Int"; | |
| 31 | ||
| 5648 | 32 | Goal "FP F <= FP_Orig F"; | 
| 4776 | 33 | by (rtac (stable_FP_Int RS FP_Orig_weakest) 1); | 
| 34 | val lemma1 = result(); | |
| 35 | ||
| 5648 | 36 | Goalw [FP_Orig_def, FP_def] "FP_Orig F <= FP F"; | 
| 4776 | 37 | by (Clarify_tac 1); | 
| 38 | by (dres_inst_tac [("x", "{x}")] spec 1);
 | |
| 39 | by (asm_full_simp_tac (simpset() addsimps [Int_insert_right]) 1); | |
| 40 | val lemma2 = result(); | |
| 41 | ||
| 5648 | 42 | Goal "FP F = FP_Orig F"; | 
| 4776 | 43 | by (rtac ([lemma1,lemma2] MRS equalityI) 1); | 
| 44 | qed "FP_equivalence"; | |
| 45 | ||
| 46 | val [prem] = goal thy | |
| 5648 | 47 | "(!!B. F : stable (A Int B)) ==> A <= FP F"; | 
| 4776 | 48 | by (simp_tac (simpset() addsimps [FP_equivalence, prem RS FP_Orig_weakest]) 1); | 
| 49 | qed "FP_weakest"; | |
| 50 | ||
| 5069 | 51 | Goalw [FP_def, stable_def, constrains_def] | 
| 10834 | 52 |     "-(FP F) = (UN act: Acts F. -{s. act``{s} <= {s}})";
 | 
| 4776 | 53 | by (Blast_tac 1); | 
| 54 | qed "Compl_FP"; | |
| 55 | ||
| 10834 | 56 | Goal "A - (FP F) = (UN act: Acts F. A - {s. act``{s} <= {s}})";
 | 
| 4776 | 57 | by (simp_tac (simpset() addsimps [Diff_eq, Compl_FP]) 1); | 
| 58 | qed "Diff_FP"; | |
| 59 |