src/HOL/Library/While_Combinator.thy
author wenzelm
Sat, 03 Feb 2001 17:43:05 +0100
changeset 11047 10c51288b00d
parent 10997 e14029f92770
child 11284 981ea92a86dd
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10251
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Library/While.thy
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     3
    Author:     Tobias Nipkow
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     4
    Copyright   2000 TU Muenchen
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     5
*)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     6
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     7
header {*
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     8
 \title{A general ``while'' combinator}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
     9
 \author{Tobias Nipkow}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    10
*}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    11
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    12
theory While_Combinator = Main:
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    13
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    14
text {*
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    15
 We define a while-combinator @{term while} and prove: (a) an
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    16
 unrestricted unfolding law (even if while diverges!)  (I got this
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    17
 idea from Wolfgang Goerigk), and (b) the invariant rule for reasoning
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    18
 about @{term while}.
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    19
*}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    20
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    21
consts while_aux :: "('a => bool) \<times> ('a => 'a) \<times> 'a => 'a"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    22
recdef while_aux
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    23
  "same_fst (\<lambda>b. True) (\<lambda>b. same_fst (\<lambda>c. True) (\<lambda>c.
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    24
      {(t, s).  b s \<and> c s = t \<and>
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    25
        \<not> (\<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))}))"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    26
  "while_aux (b, c, s) =
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    27
    (if (\<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1)))
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    28
      then arbitrary
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    29
      else if b s then while_aux (b, c, c s)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    30
      else s)"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    31
10774
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    32
recdef_tc while_aux_tc: while_aux
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    33
  apply (rule wf_same_fst)
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    34
  apply (rule wf_same_fst)
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    35
  apply (simp add: wf_iff_no_infinite_down_chain)
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    36
  apply blast
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    37
  done
4de3a0d3ae28 recdef_tc;
wenzelm
parents: 10673
diff changeset
    38
10251
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    39
constdefs
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    40
  while :: "('a => bool) => ('a => 'a) => 'a => 'a"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    41
  "while b c s == while_aux (b, c, s)"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    42
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    43
lemma while_aux_unfold:
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    44
  "while_aux (b, c, s) =
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    45
    (if \<exists>f. f 0 = s \<and> (\<forall>i. b (f i) \<and> c (f i) = f (i + 1))
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    46
      then arbitrary
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    47
      else if b s then while_aux (b, c, c s)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    48
      else s)"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    49
  apply (rule while_aux_tc [THEN while_aux.simps [THEN trans]])
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    50
   apply (simp add: same_fst_def)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    51
  apply (rule refl)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    52
  done
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    53
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    54
text {*
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    55
 The recursion equation for @{term while}: directly executable!
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    56
*}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    57
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    58
theorem while_unfold:
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    59
    "while b c s = (if b s then while b c (c s) else s)"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    60
  apply (unfold while_def)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    61
  apply (rule while_aux_unfold [THEN trans])
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    62
  apply auto
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    63
  apply (subst while_aux_unfold)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    64
  apply simp
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    65
  apply clarify
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    66
  apply (erule_tac x = "\<lambda>i. f (Suc i)" in allE)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    67
  apply blast
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    68
  done
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    69
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    70
hide const while_aux
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    71
10251
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    72
text {*
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    73
 The proof rule for @{term while}, where @{term P} is the invariant.
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    74
*}
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    75
10653
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
    76
theorem while_rule_lemma[rule_format]:
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    77
  "[| !!s. P s ==> b s ==> P (c s);
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    78
      !!s. P s ==> \<not> b s ==> Q s;
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    79
      wf {(t, s). P s \<and> b s \<and> t = c s} |] ==>
10251
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    80
    P s --> Q (while b c s)"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    81
proof -
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    82
  case antecedent
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    83
  assume wf: "wf {(t, s). P s \<and> b s \<and> t = c s}"
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    84
  show ?thesis
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    85
    apply (induct s rule: wf [THEN wf_induct])
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    86
    apply simp
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    87
    apply clarify
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    88
    apply (subst while_unfold)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    89
    apply (simp add: antecedent)
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    90
    done
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    91
qed
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
    92
10653
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
    93
theorem while_rule:
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    94
  "[| P s;
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    95
      !!s. [| P s; b s  |] ==> P (c s);
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    96
      !!s. [| P s; \<not> b s  |] ==> Q s;
10997
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
    97
      wf r;
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    98
      !!s. [| P s; b s  |] ==> (c s, s) \<in> r |] ==>
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
    99
   Q (while b c s)"
10653
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   100
apply (rule while_rule_lemma)
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   101
prefer 4 apply assumption
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   102
apply blast
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   103
apply blast
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   104
apply(erule wf_subset)
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   105
apply blast
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   106
done
55f33da63366 small mods.
nipkow
parents: 10269
diff changeset
   107
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   108
text {*
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   109
 \medskip An application: computation of the @{term lfp} on finite
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   110
 sets via iteration.
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   111
*}
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   112
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   113
theorem lfp_conv_while:
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   114
  "[| mono f; finite U; f U = U |] ==>
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   115
    lfp f = fst (while (\<lambda>(A, fA). A \<noteq> fA) (\<lambda>(A, fA). (fA, f fA)) ({}, f {}))"
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   116
apply (rule_tac P = "\<lambda>(A, B). (A \<subseteq> U \<and> B = f A \<and> A \<subseteq> B \<and> B \<subseteq> lfp f)" and
11047
wenzelm
parents: 10997
diff changeset
   117
                r = "((Pow U \<times> UNIV) \<times> (Pow U \<times> UNIV)) \<inter>
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   118
                     inv_image finite_psubset (op - U o fst)" in while_rule)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   119
   apply (subst lfp_unfold)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   120
    apply assumption
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   121
   apply (simp add: monoD)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   122
  apply (subst lfp_unfold)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   123
   apply assumption
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   124
  apply clarsimp
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   125
  apply (blast dest: monoD)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   126
 apply (fastsimp intro!: lfp_lowerbound)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   127
 apply (blast intro: wf_finite_psubset Int_lower2 [THEN [2] wf_subset])
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   128
apply (clarsimp simp add: inv_image_def finite_psubset_def order_less_le)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   129
apply (blast intro!: finite_Diff dest: monoD)
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   130
done
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   131
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   132
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   133
text {*
10997
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   134
 An example of using the @{term while} combinator.\footnote{It is safe
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   135
 to keep the example here, since there is no effect on the current
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   136
 theory.}
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   137
*}
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   138
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   139
theorem "P (lfp (\<lambda>N::int set. {#0} \<union> {(n + #2) mod #6 | n. n \<in> N})) =
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   140
    P {#0, #4, #2}"
10997
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   141
proof -
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   142
  have aux: "!!f A B. {f n | n. A n \<or> B n} = {f n | n. A n} \<union> {f n | n. B n}"
10984
8f49dcbec859 Merged Example into While_Combi
nipkow
parents: 10774
diff changeset
   143
    apply blast
10997
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   144
    done
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   145
  show ?thesis
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   146
    apply (subst lfp_conv_while [where ?U = "{#0, #1, #2, #3, #4, #5}"])
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   147
       apply (rule monoI)
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   148
      apply blast
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   149
     apply simp
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   150
    apply (simp add: aux set_eq_subset)
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   151
    txt {* The fixpoint computation is performed purely by rewriting: *}
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   152
    apply (simp add: while_unfold aux set_eq_subset del: subset_empty)
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   153
    done
e14029f92770 avoid dead code;
wenzelm
parents: 10984
diff changeset
   154
qed
10251
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
   155
5cc44cae9590 A general ``while'' combinator (from main HOL);
wenzelm
parents:
diff changeset
   156
end