| 6706 |      1 | (*  Title:      HOL/UNITY/Follows
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1998  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | The Follows relation of Charpentier and Sivilotte
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | (*Does this hold for "invariant"?*)
 | 
|  |     10 | Goal "mono h ==> Always {s. f s <= g s} <= Always {s. h (f s) <= h (g s)}";
 | 
|  |     11 | by (asm_simp_tac (simpset() addsimps [Always_eq_includes_reachable]) 1);
 | 
|  |     12 | by (blast_tac (claset() addIs [monoD]) 1);
 | 
|  |     13 | qed "mono_Always_o";
 | 
|  |     14 | 
 | 
|  |     15 | Goalw [Follows_def]
 | 
|  |     16 |     "mono (h::'a::order => 'b::order) \
 | 
|  |     17 | \    ==> (INT j. {s. j <= g s} LeadsTo {s. j <= f s}) <= \
 | 
|  |     18 | \        (INT k. {s. k <= h (g s)} LeadsTo {s. k <= h (f s)})";
 | 
|  |     19 | by Auto_tac;
 | 
|  |     20 | by (rtac single_LeadsTo_I 1);
 | 
|  |     21 | by (dres_inst_tac [("x", "g s")] spec 1);
 | 
|  |     22 | by (etac LeadsTo_weaken 1);
 | 
|  |     23 | by (ALLGOALS (blast_tac (claset() addIs [monoD, order_trans])));
 | 
|  |     24 | qed "mono_LeadsTo_o";
 | 
|  |     25 | 
 | 
| 6809 |     26 | Goalw [Follows_def] "mono h ==> f Fols g <= (h o f) Fols (h o g)";
 | 
| 6706 |     27 | by (Clarify_tac 1);
 | 
|  |     28 | by (asm_full_simp_tac
 | 
|  |     29 |     (simpset() addsimps [impOfSubs mono_Increasing_o,
 | 
|  |     30 | 			 impOfSubs mono_Always_o,
 | 
|  |     31 | 			 impOfSubs mono_LeadsTo_o RS INT_D]) 1);
 | 
|  |     32 | qed "mono_Follows_o";
 | 
|  |     33 | 
 | 
|  |     34 | Goalw [Follows_def]
 | 
| 6809 |     35 |      "[| F : f Fols g;  F: g Fols h |] ==> F : f Fols h";
 | 
| 6706 |     36 | by (asm_full_simp_tac (simpset() addsimps [Always_eq_includes_reachable]) 1);
 | 
|  |     37 | by (blast_tac (claset() addIs [order_trans, LeadsTo_Trans]) 1);
 | 
|  |     38 | qed "Follows_trans";
 | 
|  |     39 | 
 | 
|  |     40 | 
 | 
|  |     41 | (*Can replace "Un" by any sup.  But existing max only works for linorders.*)
 | 
|  |     42 | 
 | 
|  |     43 | Goalw [increasing_def, stable_def, constrains_def]
 | 
|  |     44 |     "[| F : increasing f;  F: increasing g |] \
 | 
|  |     45 | \    ==> F : increasing (%s. (f s) Un (g s))";
 | 
|  |     46 | by Auto_tac;
 | 
|  |     47 | by (dres_inst_tac [("x","f xa")] spec 1);
 | 
|  |     48 | by (dres_inst_tac [("x","g xa")] spec 1);
 | 
|  |     49 | by (blast_tac (claset() addSDs [bspec]) 1);
 | 
|  |     50 | qed "increasing_Un";
 | 
|  |     51 | 
 | 
|  |     52 | Goalw [Increasing_def, Stable_def, Constrains_def, stable_def, constrains_def]
 | 
|  |     53 |     "[| F : Increasing f;  F: Increasing g |] \
 | 
|  |     54 | \    ==> F : Increasing (%s. (f s) Un (g s))";
 | 
|  |     55 | by Auto_tac;
 | 
|  |     56 | by (dres_inst_tac [("x","f xa")] spec 1);
 | 
|  |     57 | by (dres_inst_tac [("x","g xa")] spec 1);
 | 
|  |     58 | by (blast_tac (claset() addSDs [bspec]) 1);
 | 
|  |     59 | qed "Increasing_Un";
 | 
|  |     60 | 
 | 
|  |     61 | 
 | 
|  |     62 | Goal "[| F : Always {s. f' s <= f s}; F : Always {s. g' s <= g s} |] \
 | 
|  |     63 | \     ==> F : Always {s. f' s Un g' s <= f s Un g s}";
 | 
|  |     64 | by (asm_full_simp_tac (simpset() addsimps [Always_eq_includes_reachable]) 1);
 | 
|  |     65 | by (Blast_tac 1);
 | 
|  |     66 | qed "Always_Un";
 | 
|  |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | 
 | 
|  |     70 | Goalw [Increasing_def]
 | 
|  |     71 |      "F : Increasing f ==> F : Stable {s. x <= f s}";
 | 
|  |     72 | by (Blast_tac 1);
 | 
|  |     73 | qed "IncreasingD";
 | 
|  |     74 | 
 | 
|  |     75 | 
 | 
|  |     76 | (*Lemma to re-use the argument that one variable increases (progress)
 | 
|  |     77 |   while the other variable doesn't decrease (safety)*)
 | 
|  |     78 | Goal "[| F : Increasing f; F : Increasing g; \
 | 
|  |     79 | \        F : Increasing g'; F : Always {s. f' s <= f s};\
 | 
|  |     80 | \        ALL k. F : {s. k <= f s} LeadsTo {s. k <= f' s} |]\
 | 
|  |     81 | \     ==> F : {s. k <= f s Un g s} LeadsTo {s. k <= f' s Un g s}";
 | 
|  |     82 | by (rtac single_LeadsTo_I 1);
 | 
|  |     83 | by (dres_inst_tac [("x", "f s")] IncreasingD 1);
 | 
|  |     84 | by (dres_inst_tac [("x", "g s")] IncreasingD 1);
 | 
|  |     85 | by (rtac LeadsTo_weaken 1);
 | 
|  |     86 | by (rtac PSP_Stable 1);
 | 
|  |     87 | by (eres_inst_tac [("x", "f s")] spec 1);
 | 
|  |     88 | by (etac Stable_Int 1);
 | 
|  |     89 | by (assume_tac 1);
 | 
|  |     90 | by (Blast_tac 1);
 | 
|  |     91 | by (Blast_tac 1);
 | 
|  |     92 | qed "Follows_Un_lemma";
 | 
|  |     93 | 
 | 
|  |     94 | Goalw [Follows_def]
 | 
| 6809 |     95 |     "[| F : f' Fols f;  F: g' Fols g |] \
 | 
|  |     96 | \    ==> F : (%s. (f' s) Un (g' s)) Fols (%s. (f s) Un (g s))";
 | 
| 6706 |     97 | by (asm_full_simp_tac (simpset() addsimps [Increasing_Un, Always_Un]) 1);
 | 
|  |     98 | by Auto_tac;
 | 
|  |     99 | by (rtac LeadsTo_Trans 1);
 | 
|  |    100 | by (blast_tac (claset() addIs [Follows_Un_lemma]) 1);
 | 
|  |    101 | (*Weakening is used to exchange Un's arguments*)
 | 
|  |    102 | by (blast_tac (claset() addIs [Follows_Un_lemma RS LeadsTo_weaken]) 1);
 | 
|  |    103 | qed "Follows_Un";
 | 
|  |    104 | 
 |