src/HOL/UNITY/Comp/Progress.thy
author wenzelm
Sat, 18 Jul 2015 20:54:56 +0200
changeset 60754 02924903a6fd
parent 58889 5b7a9633cfa8
child 63146 f1ecba0272f9
permissions -rw-r--r--
prefer tactics with explicit context;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
37936
1e4c5015a72e updated some headers;
wenzelm
parents: 35416
diff changeset
     1
(*  Title:      HOL/UNITY/Comp/Progress.thy
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     3
    Copyright   2003  University of Cambridge
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     4
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     5
David Meier's thesis
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     6
*)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     7
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 46577
diff changeset
     8
section{*Progress Set Examples*}
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
     9
18556
dc39832e9280 added explicit paths to required theories
paulson
parents: 16417
diff changeset
    10
theory Progress imports "../UNITY_Main" begin
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    11
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    12
subsection {*The Composition of Two Single-Assignment Programs*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    13
text{*Thesis Section 4.4.2*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    14
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 28869
diff changeset
    15
definition FF :: "int program" where
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 37936
diff changeset
    16
    "FF = mk_total_program (UNIV, {range (\<lambda>x. (x, x+1))}, UNIV)"
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    17
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 28869
diff changeset
    18
definition GG :: "int program" where
46577
e5438c5797ae tuned proofs;
wenzelm
parents: 37936
diff changeset
    19
    "GG = mk_total_program (UNIV, {range (\<lambda>x. (x, 2*x))}, UNIV)"
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    20
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    21
subsubsection {*Calculating @{term "wens_set FF (atLeast k)"}*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    22
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    23
lemma Domain_actFF: "Domain (range (\<lambda>x::int. (x, x + 1))) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    24
by force
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    25
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    26
lemma FF_eq:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    27
      "FF = mk_program (UNIV, {range (\<lambda>x. (x, x+1))}, UNIV)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    28
by (simp add: FF_def mk_total_program_def totalize_def totalize_act_def
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    29
              program_equalityI Domain_actFF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    30
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    31
lemma wp_actFF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    32
     "wp (range (\<lambda>x::int. (x, x + 1))) (atLeast k) = atLeast (k - 1)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    33
by (force simp add: wp_def)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    34
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    35
lemma wens_FF: "wens FF (range (\<lambda>x. (x, x+1))) (atLeast k) = atLeast (k - 1)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    36
by (force simp add: FF_eq wens_single_eq wp_actFF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    37
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    38
lemma single_valued_actFF: "single_valued (range (\<lambda>x::int. (x, x + 1)))"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    39
by (force simp add: single_valued_def)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    40
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    41
lemma wens_single_finite_FF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    42
     "wens_single_finite (range (\<lambda>x. (x, x+1))) (atLeast k) n =
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    43
      atLeast (k - int n)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    44
apply (induct n, simp)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    45
apply (force simp add: wens_FF
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    46
            def_wens_single_finite_Suc_eq_wens [OF FF_eq single_valued_actFF])
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    47
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    48
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    49
lemma wens_single_FF_eq_UNIV:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    50
     "wens_single (range (\<lambda>x::int. (x, x + 1))) (atLeast k) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    51
apply (auto simp add: wens_single_eq_Union)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    52
apply (rule_tac x="nat (k-x)" in exI)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    53
apply (simp add: wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    54
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    55
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    56
lemma wens_set_FF:
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    57
     "wens_set FF (atLeast k) = insert UNIV (atLeast ` atMost k)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    58
apply (auto simp add: wens_set_single_eq [OF FF_eq single_valued_actFF]
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    59
                      wens_single_FF_eq_UNIV wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    60
apply (erule notE)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    61
apply (rule_tac x="nat (k-xa)" in range_eqI)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    62
apply (simp add: wens_single_finite_FF)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    63
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    64
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    65
subsubsection {*Proving @{term "FF \<in> UNIV leadsTo atLeast (k::int)"}*}
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    66
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    67
lemma atLeast_ensures: "FF \<in> atLeast (k - 1) ensures atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    68
apply (simp add: Progress.wens_FF [symmetric] wens_ensures)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    69
apply (simp add: wens_ensures FF_eq)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    70
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    71
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    72
lemma atLeast_leadsTo: "FF \<in> atLeast (k - int n) leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    73
apply (induct n)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    74
 apply (simp_all add: leadsTo_refl)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    75
apply (rule_tac A = "atLeast (k - int n - 1)" in leadsTo_weaken_L)
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    76
 apply (blast intro: leadsTo_Trans atLeast_ensures, force) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    77
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    78
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    79
lemma UN_atLeast_UNIV: "(\<Union>n. atLeast (k - int n)) = UNIV"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    80
apply auto 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    81
apply (rule_tac x = "nat (k - x)" in exI, simp) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    82
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    83
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    84
lemma FF_leadsTo: "FF \<in> UNIV leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    85
apply (subst UN_atLeast_UNIV [symmetric]) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    86
apply (rule leadsTo_UN [OF atLeast_leadsTo]) 
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    87
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    88
13866
b42d7983a822 More "progress set" material
paulson
parents: 13851
diff changeset
    89
text{*Result (4.39): Applying the leadsTo-Join Theorem*}
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    90
theorem "FF\<squnion>GG \<in> atLeast 0 leadsTo atLeast (k::int)"
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    91
apply (subgoal_tac "FF\<squnion>GG \<in> (atLeast 0 \<inter> atLeast 0) leadsTo atLeast k")
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    92
 apply simp
13866
b42d7983a822 More "progress set" material
paulson
parents: 13851
diff changeset
    93
apply (rule_tac T = "atLeast 0" in leadsTo_Join)
28869
191cbfac6c9a Strange. The proof worked in the 2008 release. In order to make it work now, the last line of the proof must be moved up two places. In other words, the first proof step is now returning its subgoals in a different order from before.
paulson
parents: 28866
diff changeset
    94
  apply (rule leadsTo_weaken_L [OF FF_leadsTo], simp) 
191cbfac6c9a Strange. The proof worked in the 2008 release. In order to make it work now, the last line of the proof must be moved up two places. In other words, the first proof step is now returning its subgoals in a different order from before.
paulson
parents: 28866
diff changeset
    95
 apply (simp add: awp_iff_constrains FF_def, safety)
191cbfac6c9a Strange. The proof worked in the 2008 release. In order to make it work now, the last line of the proof must be moved up two places. In other words, the first proof step is now returning its subgoals in a different order from before.
paulson
parents: 28866
diff changeset
    96
apply (simp add: awp_iff_constrains GG_def wens_set_FF, safety)
13846
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    97
done
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    98
b2c494d76012 new examples theory
paulson
parents:
diff changeset
    99
end