| 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);
 | 
|  |     13 | by (rtac ball_constrains_UN 1);
 | 
|  |     14 | by (Simp_tac 1);
 | 
|  |     15 | qed "stable_FP_Orig_Int";
 | 
|  |     16 | 
 | 
|  |     17 | 
 | 
|  |     18 | val prems = goalw thy [FP_Orig_def, stable_def]
 | 
| 5648 |     19 |     "(!!B. F : stable (A Int B)) ==> A <= FP_Orig F";
 | 
| 4776 |     20 | by (blast_tac (claset() addIs prems) 1);
 | 
|  |     21 | qed "FP_Orig_weakest";
 | 
|  |     22 | 
 | 
|  |     23 | 
 | 
| 5648 |     24 | Goal "F : stable (FP F Int B)";
 | 
|  |     25 | by (subgoal_tac "FP F Int B = (UN x:B. FP F Int {x})" 1);
 | 
| 4776 |     26 | by (Blast_tac 2);
 | 
|  |     27 | by (asm_simp_tac (simpset() addsimps [Int_insert_right]) 1);
 | 
|  |     28 | by (rewrite_goals_tac [FP_def, stable_def]);
 | 
|  |     29 | by (rtac ball_constrains_UN 1);
 | 
|  |     30 | by (Simp_tac 1);
 | 
|  |     31 | qed "stable_FP_Int";
 | 
|  |     32 | 
 | 
| 5648 |     33 | Goal "FP F <= FP_Orig F";
 | 
| 4776 |     34 | by (rtac (stable_FP_Int RS FP_Orig_weakest) 1);
 | 
|  |     35 | val lemma1 = result();
 | 
|  |     36 | 
 | 
| 5648 |     37 | Goalw [FP_Orig_def, FP_def] "FP_Orig F <= FP F";
 | 
| 4776 |     38 | by (Clarify_tac 1);
 | 
|  |     39 | by (dres_inst_tac [("x", "{x}")] spec 1);
 | 
|  |     40 | by (asm_full_simp_tac (simpset() addsimps [Int_insert_right]) 1);
 | 
|  |     41 | val lemma2 = result();
 | 
|  |     42 | 
 | 
| 5648 |     43 | Goal "FP F = FP_Orig F";
 | 
| 4776 |     44 | by (rtac ([lemma1,lemma2] MRS equalityI) 1);
 | 
|  |     45 | qed "FP_equivalence";
 | 
|  |     46 | 
 | 
|  |     47 | val [prem] = goal thy
 | 
| 5648 |     48 |     "(!!B. F : stable (A Int B)) ==> A <= FP F";
 | 
| 4776 |     49 | by (simp_tac (simpset() addsimps [FP_equivalence, prem RS FP_Orig_weakest]) 1);
 | 
|  |     50 | qed "FP_weakest";
 | 
|  |     51 | 
 | 
| 5069 |     52 | Goalw [FP_def, stable_def, constrains_def]
 | 
| 5648 |     53 |     "-(FP F) = (UN act: Acts F. -{s. act^^{s} <= {s}})";
 | 
| 4776 |     54 | by (Blast_tac 1);
 | 
|  |     55 | qed "Compl_FP";
 | 
|  |     56 | 
 | 
| 5648 |     57 | Goal "A - (FP F) = (UN act: Acts F. A - {s. act^^{s} <= {s}})";
 | 
| 4776 |     58 | by (simp_tac (simpset() addsimps [Diff_eq, Compl_FP]) 1);
 | 
|  |     59 | qed "Diff_FP";
 | 
|  |     60 | 
 |