src/HOL/ex/InductiveInvariant.thy
author blanchet
Tue, 31 Aug 2010 20:24:28 +0200
changeset 38942 e10c11971fa7
parent 32960 69916a850301
child 44013 5cfc1c36ae97
permissions -rw-r--r--
"try" -- a new diagnosis tool that tries to apply several methods in parallel
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
     1
(*
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
     2
    Author:     Sava Krsti\'{c} and John Matthews
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     3
*)
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     4
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     5
header {* Some of the results in Inductive Invariants for Nested Recursion *}
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15636
diff changeset
     7
theory InductiveInvariant imports Main begin
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
     8
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
     9
text {* A formalization of some of the results in \emph{Inductive
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
    10
  Invariants for Nested Recursion}, by Sava Krsti\'{c} and John
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
    11
  Matthews.  Appears in the proceedings of TPHOLs 2003, LNCS
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
    12
  vol. 2758, pp. 253-269. *}
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    13
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    14
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    15
text "S is an inductive invariant of the functional F with respect to the wellfounded relation r."
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    16
19736
wenzelm
parents: 17388
diff changeset
    17
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    18
  indinv :: "('a * 'a) set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool" where
19736
wenzelm
parents: 17388
diff changeset
    19
  "indinv r S F = (\<forall>f x. (\<forall>y. (y,x) : r --> S y (f y)) --> S x (F f x))"
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    20
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    21
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    22
text "S is an inductive invariant of the functional F on set D with respect to the wellfounded relation r."
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    23
19736
wenzelm
parents: 17388
diff changeset
    24
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 19736
diff changeset
    25
  indinv_on :: "('a * 'a) set => 'a set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool" where
19736
wenzelm
parents: 17388
diff changeset
    26
  "indinv_on r D S F = (\<forall>f. \<forall>x\<in>D. (\<forall>y\<in>D. (y,x) \<in> r --> S y (f y)) --> S x (F f x))"
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    27
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    28
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    29
text "The key theorem, corresponding to theorem 1 of the paper. All other results
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    30
      in this theory are proved using instances of this theorem, and theorems
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    31
      derived from this theorem."
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    32
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    33
theorem indinv_wfrec:
19736
wenzelm
parents: 17388
diff changeset
    34
  assumes wf:  "wf r" and
wenzelm
parents: 17388
diff changeset
    35
          inv: "indinv r S F"
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    36
  shows        "S x (wfrec r F x)"
19736
wenzelm
parents: 17388
diff changeset
    37
  using wf
wenzelm
parents: 17388
diff changeset
    38
proof (induct x)
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    39
  fix x
19736
wenzelm
parents: 17388
diff changeset
    40
  assume  IHYP: "!!y. (y,x) \<in> r \<Longrightarrow> S y (wfrec r F y)"
wenzelm
parents: 17388
diff changeset
    41
  then have     "!!y. (y,x) \<in> r \<Longrightarrow> S y (cut (wfrec r F) r x y)" by (simp add: tfl_cut_apply)
wenzelm
parents: 17388
diff changeset
    42
  with inv have "S x (F (cut (wfrec r F) r x) x)" by (unfold indinv_def, blast)
wenzelm
parents: 17388
diff changeset
    43
  thus "S x (wfrec r F x)" using wf by (simp add: wfrec)
15636
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    44
qed
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    45
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    46
theorem indinv_on_wfrec:
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    47
  assumes WF:  "wf r" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    48
          INV: "indinv_on r D S F" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    49
          D:   "x\<in>D"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    50
  shows        "S x (wfrec r F x)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    51
apply (insert INV D indinv_wfrec [OF WF, of "% x y. x\<in>D --> S x y"])
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    52
by (simp add: indinv_on_def indinv_def)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    53
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    54
theorem ind_fixpoint_on_lemma:
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    55
  assumes WF:  "wf r" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    56
         INV: "\<forall>f. \<forall>x\<in>D. (\<forall>y\<in>D. (y,x) \<in> r --> S y (wfrec r F y) & f y = wfrec r F y)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    57
                               --> S x (wfrec r F x) & F f x = wfrec r F x" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    58
           D: "x\<in>D"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    59
  shows "F (wfrec r F) x = wfrec r F x & S x (wfrec r F x)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    60
proof (rule indinv_on_wfrec [OF WF _ D, of "% a b. F (wfrec r F) a = b & wfrec r F a = b & S a b" F, simplified])
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    61
  show "indinv_on r D (%a b. F (wfrec r F) a = b & wfrec r F a = b & S a b) F"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    62
  proof (unfold indinv_on_def, clarify)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    63
    fix f x
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    64
    assume A1: "\<forall>y\<in>D. (y, x) \<in> r --> F (wfrec r F) y = f y & wfrec r F y = f y & S y (f y)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    65
    assume D': "x\<in>D"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    66
    from A1 INV [THEN spec, of f, THEN bspec, OF D']
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    67
      have "S x (wfrec r F x)" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    68
           "F f x = wfrec r F x" by auto
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    69
    moreover
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    70
    from A1 have "\<forall>y\<in>D. (y, x) \<in> r --> S y (wfrec r F y)" by auto
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    71
    with D' INV [THEN spec, of "wfrec r F", simplified]
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    72
      have "F (wfrec r F) x = wfrec r F x" by blast
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    73
    ultimately show "F (wfrec r F) x = F f x & wfrec r F x = F f x & S x (F f x)" by auto
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    74
  qed
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    75
qed
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    76
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    77
theorem ind_fixpoint_lemma:
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    78
  assumes WF:  "wf r" and
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    79
         INV: "\<forall>f x. (\<forall>y. (y,x) \<in> r --> S y (wfrec r F y) & f y = wfrec r F y)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    80
                         --> S x (wfrec r F x) & F f x = wfrec r F x"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    81
  shows "F (wfrec r F) x = wfrec r F x & S x (wfrec r F x)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    82
apply (rule ind_fixpoint_on_lemma [OF WF _ UNIV_I, simplified])
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    83
by (rule INV)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    84
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    85
theorem tfl_indinv_wfrec:
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    86
"[| f == wfrec r F; wf r; indinv r S F |]
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    87
 ==> S x (f x)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    88
by (simp add: indinv_wfrec)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    89
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    90
theorem tfl_indinv_on_wfrec:
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    91
"[| f == wfrec r F; wf r; indinv_on r D S F; x\<in>D |]
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    92
 ==> S x (f x)"
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    93
by (simp add: indinv_on_wfrec)
57c437b70521 converted from DOS to UNIX format
paulson
parents: 14244
diff changeset
    94
17388
495c799df31d tuned headers etc.;
wenzelm
parents: 16417
diff changeset
    95
end