src/HOL/UNITY/Simple/Channel.thy
author obua
Sun, 09 May 2004 23:04:36 +0200
changeset 14722 8e739a6eaf11
parent 13806 fd40c9d9076b
child 16417 9bc16273c2d4
permissions -rw-r--r--
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Channel
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     2
    ID:         $Id$
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     5
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     6
Unordered Channel
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     7
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming" (1994), section 13.3
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
     9
*)
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    10
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    11
theory Channel = UNITY_Main:
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    12
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    13
types state = "nat set"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    14
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    15
consts
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    16
  F :: "state program"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    17
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    18
constdefs
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    19
  minSet :: "nat set => nat option"
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    20
    "minSet A == if A={} then None else Some (LEAST x. x \<in> A)"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    21
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    22
axioms
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    23
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    24
  UC1:  "F \<in> (minSet -` {Some x}) co (minSet -` (Some`atLeast x))"
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    25
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    26
  (*  UC1  "F \<in> {s. minSet s = x} co {s. x \<subseteq> minSet s}"  *)
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    27
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    28
  UC2:  "F \<in> (minSet -` {Some x}) leadsTo {s. x \<notin> s}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    29
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    30
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    31
(*None represents "infinity" while Some represents proper integers*)
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    32
lemma minSet_eq_SomeD: "minSet A = Some x ==> x \<in> A"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    33
apply (unfold minSet_def)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    34
apply (simp split: split_if_asm)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    35
apply (fast intro: LeastI)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    36
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    37
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    38
lemma minSet_empty [simp]: " minSet{} = None"
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    39
by (unfold minSet_def, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    40
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    41
lemma minSet_nonempty: "x \<in> A ==> minSet A = Some (LEAST x. x \<in> A)"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    42
by (unfold minSet_def, auto)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    43
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    44
lemma minSet_greaterThan:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    45
     "F \<in> (minSet -` {Some x}) leadsTo (minSet -` (Some`greaterThan x))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    46
apply (rule leadsTo_weaken)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    47
apply (rule_tac x1=x in psp [OF UC2 UC1], safe)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    48
apply (auto dest: minSet_eq_SomeD simp add: linorder_neq_iff)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    49
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    50
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    51
(*The induction*)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    52
lemma Channel_progress_lemma:
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    53
     "F \<in> (UNIV-{{}}) leadsTo (minSet -` (Some`atLeast y))"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    54
apply (rule leadsTo_weaken_R)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    55
apply (rule_tac l = y and f = "the o minSet" and B = "{}" 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    56
       in greaterThan_bounded_induct, safe)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    57
apply (simp_all (no_asm_simp))
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    58
apply (drule_tac [2] minSet_nonempty)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    59
 prefer 2 apply simp 
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    60
apply (rule minSet_greaterThan [THEN leadsTo_weaken], safe)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    61
apply simp_all
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    62
apply (drule minSet_nonempty, simp)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    63
done
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    64
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    65
13806
fd40c9d9076b more tidying
paulson
parents: 13785
diff changeset
    66
lemma Channel_progress: "!!y::nat. F \<in> (UNIV-{{}}) leadsTo {s. y \<notin> s}"
13785
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    67
apply (rule Channel_progress_lemma [THEN leadsTo_weaken_R], clarify)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    68
apply (frule minSet_nonempty)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    69
apply (auto dest: Suc_le_lessD not_less_Least)
e2fcd88be55d Partial conversion of UNITY to Isar new-style theories
paulson
parents: 12919
diff changeset
    70
done
11195
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    71
65ede8dfe304 reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff changeset
    72
end