src/HOL/UNITY/Comp/Progress.thy
author paulson
Thu, 02 Jun 2005 13:47:08 +0200
changeset 16184 80617b8d33c5
parent 13866 b42d7983a822
child 16417 9bc16273c2d4
permissions -rw-r--r--
renamed "constrains" to "safety" to avoid keyword clash
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Transformers
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     4
    Copyright   2003  University of Cambridge
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     5
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     6
David Meier's thesis
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     7
*)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     8
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     9
header{*Progress Set Examples*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    10
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    11
theory Progress = UNITY_Main:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    12
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    13
subsection {*The Composition of Two Single-Assignment Programs*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    14
text{*Thesis Section 4.4.2*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    15
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    16
constdefs
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    17
  FF :: "int program"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    18
    "FF == mk_total_program (UNIV, {range (\<lambda>x. (x, x+1))}, UNIV)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    19
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    20
  GG :: "int program"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    21
    "GG == mk_total_program (UNIV, {range (\<lambda>x. (x, 2*x))}, UNIV)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    22
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    23
subsubsection {*Calculating @{term "wens_set FF (atLeast k)"}*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    24
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    25
lemma Domain_actFF: "Domain (range (\<lambda>x::int. (x, x + 1))) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    26
by force
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    27
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    28
lemma FF_eq:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    29
      "FF = mk_program (UNIV, {range (\<lambda>x. (x, x+1))}, UNIV)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    30
by (simp add: FF_def mk_total_program_def totalize_def totalize_act_def
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    31
              program_equalityI Domain_actFF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    32
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    33
lemma wp_actFF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    34
     "wp (range (\<lambda>x::int. (x, x + 1))) (atLeast k) = atLeast (k - 1)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    35
by (force simp add: wp_def)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    36
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    37
lemma wens_FF: "wens FF (range (\<lambda>x. (x, x+1))) (atLeast k) = atLeast (k - 1)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    38
by (force simp add: FF_eq wens_single_eq wp_actFF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    39
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    40
lemma single_valued_actFF: "single_valued (range (\<lambda>x::int. (x, x + 1)))"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    41
by (force simp add: single_valued_def)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    42
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    43
lemma wens_single_finite_FF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    44
     "wens_single_finite (range (\<lambda>x. (x, x+1))) (atLeast k) n =
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    45
      atLeast (k - int n)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    46
apply (induct n, simp)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    47
apply (force simp add: wens_FF
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    48
            def_wens_single_finite_Suc_eq_wens [OF FF_eq single_valued_actFF])
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    49
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    50
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    51
lemma wens_single_FF_eq_UNIV:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    52
     "wens_single (range (\<lambda>x::int. (x, x + 1))) (atLeast k) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    53
apply (auto simp add: wens_single_eq_Union)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    54
apply (rule_tac x="nat (k-x)" in exI)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    55
apply (simp add: wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    56
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    57
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    58
lemma wens_set_FF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    59
     "wens_set FF (atLeast k) = insert UNIV (atLeast ` atMost k)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    60
apply (auto simp add: wens_set_single_eq [OF FF_eq single_valued_actFF]
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    61
                      wens_single_FF_eq_UNIV wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    62
apply (erule notE)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    63
apply (rule_tac x="nat (k-xa)" in range_eqI)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    64
apply (simp add: wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    65
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    66
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    67
subsubsection {*Proving @{term "FF \<in> UNIV leadsTo atLeast (k::int)"}*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    68
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    69
lemma atLeast_ensures: "FF \<in> atLeast (k - 1) ensures atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    70
apply (simp add: Progress.wens_FF [symmetric] wens_ensures)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    71
apply (simp add: wens_ensures FF_eq)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    72
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    73
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    74
lemma atLeast_leadsTo: "FF \<in> atLeast (k - int n) leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    75
apply (induct n)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    76
 apply (simp_all add: leadsTo_refl)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    77
apply (rule_tac A = "atLeast (k - int n - 1)" in leadsTo_weaken_L)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    78
 apply (blast intro: leadsTo_Trans atLeast_ensures, force) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    79
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    80
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    81
lemma UN_atLeast_UNIV: "(\<Union>n. atLeast (k - int n)) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    82
apply auto 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    83
apply (rule_tac x = "nat (k - x)" in exI, simp) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    84
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    85
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    86
lemma FF_leadsTo: "FF \<in> UNIV leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    87
apply (subst UN_atLeast_UNIV [symmetric]) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    88
apply (rule leadsTo_UN [OF atLeast_leadsTo]) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    89
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    90
13866
b42d7983a822 More "progress set" material
paulson
parents: 13851
diff changeset
    91
text{*Result (4.39): Applying the leadsTo-Join Theorem*}
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    92
theorem "FF\<squnion>GG \<in> atLeast 0 leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    93
apply (subgoal_tac "FF\<squnion>GG \<in> (atLeast 0 \<inter> atLeast 0) leadsTo atLeast k")
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    94
 apply simp
13866
b42d7983a822 More "progress set" material
paulson
parents: 13851
diff changeset
    95
apply (rule_tac T = "atLeast 0" in leadsTo_Join)
16184
80617b8d33c5 renamed "constrains" to "safety" to avoid keyword clash
paulson
parents: 13866
diff changeset
    96
  apply (simp add: awp_iff FF_def, safety)
80617b8d33c5 renamed "constrains" to "safety" to avoid keyword clash
paulson
parents: 13866
diff changeset
    97
 apply (simp add: awp_iff GG_def wens_set_FF, safety)
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    98
apply (rule leadsTo_weaken_L [OF FF_leadsTo], simp) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    99
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
   100
b2c494d76012 new examples theory
paulson
parents:
diff changeset
   101
end