src/HOL/UNITY/Channel.ML
author paulson
Thu, 02 Jul 1998 16:53:55 +0200
changeset 5111 8f4b72f0c15d
parent 5069 3ea049f7979d
child 5232 e5a7cdd07ea5
permissions -rw-r--r--
Uncurried functions LeadsTo and reach Deleted leading parameters thanks to new Goal command
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Channel
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Unordered Channel
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming" (1994), section 13.3
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
open Channel;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    13
AddIffs [skip];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    14
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
(*None represents "infinity" while Some represents proper integers*)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    17
Goalw [minSet_def] "minSet A = Some x --> x : A";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
by (Simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
by (fast_tac (claset() addIs [LeastI]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
qed_spec_mp "minSet_eq_SomeD";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    22
Goalw [minSet_def] " minSet{} = None";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
qed_spec_mp "minSet_empty";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
Addsimps [minSet_empty];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    26
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    27
Goalw [minSet_def] "x:A ==> minSet A = Some (LEAST x. x: A)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    29
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    30
qed_spec_mp "minSet_nonempty";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    32
Goal
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
    "leadsTo Acts (minSet -`` {Some x}) (minSet -`` (Some``greaterThan x))";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
by (rtac leadsTo_weaken 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
by (rtac ([UC2, UC1] MRS PSP) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    38
by Safe_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    39
by (auto_tac (claset() addDs [minSet_eq_SomeD], 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    40
	      simpset() addsimps [le_def, nat_neq_iff]));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
qed "minSet_greaterThan";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    44
(*The induction*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    45
Goal "leadsTo Acts (UNIV-{{}}) (minSet -`` (Some``atLeast y))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    46
by (rtac leadsTo_weaken_R 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
by (res_inst_tac  [("l", "y"), ("f", "the o minSet"), ("B", "{}")]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    48
     greaterThan_bounded_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    49
by Safe_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
by (dtac minSet_nonempty 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    52
by (Asm_full_simp_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    53
by (rtac (minSet_greaterThan RS leadsTo_weaken) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
by Safe_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    55
by (ALLGOALS Asm_full_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
by (dtac minSet_nonempty 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
by (Asm_full_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
val lemma = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    61
Goal "!!y::nat. leadsTo Acts (UNIV-{{}}) {s. y ~: s}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    62
by (rtac (lemma RS leadsTo_weaken_R) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    63
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
by (forward_tac [minSet_nonempty] 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
by (asm_full_simp_tac (simpset() addsimps [Suc_le_eq]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    66
by (blast_tac (claset() addDs [Suc_le_lessD, not_less_Least]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
qed "Channel_progress";