| 4776 |      1 | (*  Title:      HOL/UNITY/Channel
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
 | 
|  |      4 |     Copyright   1998  University of Cambridge
 | 
|  |      5 | 
 | 
|  |      6 | Unordered Channel
 | 
|  |      7 | 
 | 
|  |      8 | From Misra, "A Logic for Concurrent Programming" (1994), section 13.3
 | 
|  |      9 | *)
 | 
|  |     10 | 
 | 
|  |     11 | (*None represents "infinity" while Some represents proper integers*)
 | 
| 5111 |     12 | Goalw [minSet_def] "minSet A = Some x --> x : A";
 | 
| 4776 |     13 | by (Simp_tac 1);
 | 
|  |     14 | by (fast_tac (claset() addIs [LeastI]) 1);
 | 
|  |     15 | qed_spec_mp "minSet_eq_SomeD";
 | 
|  |     16 | 
 | 
| 5069 |     17 | Goalw [minSet_def] " minSet{} = None";
 | 
| 4776 |     18 | by (Asm_simp_tac 1);
 | 
|  |     19 | qed_spec_mp "minSet_empty";
 | 
|  |     20 | Addsimps [minSet_empty];
 | 
|  |     21 | 
 | 
| 5111 |     22 | Goalw [minSet_def] "x:A ==> minSet A = Some (LEAST x. x: A)";
 | 
| 4776 |     23 | by (ALLGOALS Asm_simp_tac);
 | 
|  |     24 | by (Blast_tac 1);
 | 
|  |     25 | qed_spec_mp "minSet_nonempty";
 | 
|  |     26 | 
 | 
| 6536 |     27 | Goal "F : (minSet -`` {Some x}) leadsTo (minSet -`` (Some``greaterThan x))";
 | 
| 4776 |     28 | by (rtac leadsTo_weaken 1);
 | 
| 6700 |     29 | by (res_inst_tac [("x1","x")] ([UC2, UC1] MRS psp) 1);
 | 
| 4776 |     30 | by Safe_tac;
 | 
|  |     31 | by (auto_tac (claset() addDs [minSet_eq_SomeD], 
 | 
| 6700 |     32 | 	      simpset() addsimps [linorder_neq_iff]));
 | 
| 4776 |     33 | qed "minSet_greaterThan";
 | 
|  |     34 | 
 | 
|  |     35 | (*The induction*)
 | 
| 6536 |     36 | Goal "F : (UNIV-{{}}) leadsTo (minSet -`` (Some``atLeast y))";
 | 
| 4776 |     37 | by (rtac leadsTo_weaken_R 1);
 | 
|  |     38 | by (res_inst_tac  [("l", "y"), ("f", "the o minSet"), ("B", "{}")]
 | 
|  |     39 |      greaterThan_bounded_induct 1);
 | 
|  |     40 | by Safe_tac;
 | 
|  |     41 | by (ALLGOALS Asm_simp_tac);
 | 
|  |     42 | by (dtac minSet_nonempty 2);
 | 
|  |     43 | by (Asm_full_simp_tac 2);
 | 
|  |     44 | by (rtac (minSet_greaterThan RS leadsTo_weaken) 1);
 | 
|  |     45 | by Safe_tac;
 | 
|  |     46 | by (ALLGOALS Asm_full_simp_tac);
 | 
|  |     47 | by (dtac minSet_nonempty 1);
 | 
|  |     48 | by (Asm_full_simp_tac 1);
 | 
|  |     49 | val lemma = result();
 | 
|  |     50 | 
 | 
|  |     51 | 
 | 
| 6536 |     52 | Goal "!!y::nat. F : (UNIV-{{}}) leadsTo {s. y ~: s}";
 | 
| 4776 |     53 | by (rtac (lemma RS leadsTo_weaken_R) 1);
 | 
|  |     54 | by (Clarify_tac 1);
 | 
|  |     55 | by (forward_tac [minSet_nonempty] 1);
 | 
|  |     56 | by (asm_full_simp_tac (simpset() addsimps [Suc_le_eq]) 1);
 | 
|  |     57 | by (blast_tac (claset() addDs [Suc_le_lessD, not_less_Least]) 1);
 | 
|  |     58 | qed "Channel_progress";
 |