src/HOL/Extraction/Warshall.thy
author berghofe
Sun, 21 Jul 2002 15:44:42 +0200
changeset 13405 d20a4e67afc8
child 13471 aed3aef2a0ca
permissions -rw-r--r--
Examples for program extraction in HOL.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Extraction/Warshall.thy
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer, TU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     4
*)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     5
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     6
header {* Warshall's algorithm *}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     7
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     8
theory Warshall = Main:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     9
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    10
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    11
  Derivation of Warshall's algorithm using program extraction,
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    12
  based on Berger, Schwichtenberg and Seisenberger \cite{Berger-JAR-2001}.
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    13
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    14
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    15
datatype b = T | F
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    16
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    17
theorem b_ind_realizer:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    18
  "R x T \<Longrightarrow> R y F \<Longrightarrow> R (case b of T \<Rightarrow> x | F \<Rightarrow> y) b"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    19
  by (induct b, simp_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    20
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    21
realizers
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    22
  b.induct (P): "\<lambda>P b x y. (case b of T \<Rightarrow> x | F \<Rightarrow> y)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    23
    "\<Lambda>P b x (h: _) y. b_ind_realizer \<cdot> (\<lambda>x b. realizes x (P b)) \<cdot> _ \<cdot> _ \<cdot> _ \<bullet> h"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    24
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    25
consts
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    26
  is_path' :: "('a \<Rightarrow> 'a \<Rightarrow> b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    27
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    28
primrec
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    29
  "is_path' r x [] z = (r x z = T)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    30
  "is_path' r x (y # ys) z = (r x y = T \<and> is_path' r y ys z)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    31
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    32
constdefs
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    33
  is_path :: "(nat \<Rightarrow> nat \<Rightarrow> b) \<Rightarrow> (nat * nat list * nat) \<Rightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    34
    nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    35
  "is_path r p i j k == fst p = j \<and> snd (snd p) = k \<and>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    36
     list_all (\<lambda>x. x < i) (fst (snd p)) \<and>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    37
     is_path' r (fst p) (fst (snd p)) (snd (snd p))"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    38
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
  conc :: "('a * 'a list * 'a) \<Rightarrow> ('a * 'a list * 'a) \<Rightarrow> ('a * 'a list * 'a)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
  "conc p q == (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    41
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    42
theorem is_path'_snoc [simp]:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
  "\<And>x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \<and> r y z = T)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    44
  by (induct ys) simp+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    45
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    46
theorem list_all_scoc [simp]: "list_all P (xs @ [x]) = (P x \<and> list_all P xs)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    47
  by (induct xs, simp+, rules)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    48
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    49
theorem list_all_lemma: 
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    50
  "list_all P xs \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> list_all Q xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    51
proof -
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    52
  assume PQ: "\<And>x. P x \<Longrightarrow> Q x"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    53
  show "list_all P xs \<Longrightarrow> list_all Q xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    54
  proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    55
    case Nil
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    56
    show ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    57
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    58
    case (Cons y ys)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    59
    hence Py: "P y" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    60
    from Cons have Pys: "list_all P ys" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    61
    show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    62
      by simp (rule conjI PQ Py Cons Pys)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    63
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    64
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    65
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    66
theorem lemma1: "\<And>p. is_path r p i j k \<Longrightarrow> is_path r p (Suc i) j k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    67
  apply (unfold is_path_def)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    68
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    69
  apply (erule conjE)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    70
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    71
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    72
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    73
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    74
theorem lemma2: "\<And>p. is_path r p 0 j k \<Longrightarrow> r j k = T"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    75
  apply (unfold is_path_def)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    76
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    77
  apply (case_tac "aa")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    78
  apply simp+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    79
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    80
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    81
theorem is_path'_conc: "is_path' r j xs i \<Longrightarrow> is_path' r i ys k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    82
  is_path' r j (xs @ i # ys) k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    83
proof -
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    84
  assume pys: "is_path' r i ys k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    85
  show "\<And>j. is_path' r j xs i \<Longrightarrow> is_path' r j (xs @ i # ys) k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    86
  proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    87
    case (Nil j)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    88
    hence "r j i = T" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    89
    with pys show ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    90
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    91
    case (Cons z zs j)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    92
    hence jzr: "r j z = T" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    93
    from Cons have pzs: "is_path' r z zs i" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    94
    show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    95
      by simp (rule conjI jzr Cons pzs)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    96
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    97
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    98
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    99
theorem lemma3:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   100
  "\<And>p q. is_path r p i j i \<Longrightarrow> is_path r q i i k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   101
   is_path r (conc p q) (Suc i) j k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   102
  apply (unfold is_path_def conc_def)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   103
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   104
  apply (erule conjE)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   105
  apply (rule conjI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   106
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   107
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   108
  apply (rule conjI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   109
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   110
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   111
  apply (rule is_path'_conc)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   112
  apply assumption+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   113
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   114
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   115
theorem lemma5:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   116
  "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> ~ is_path r p i j k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   117
   (\<exists>q. is_path r q i j i) \<and> (\<exists>q'. is_path r q' i i k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   118
proof (simp cong add: conj_cong add: split_paired_all is_path_def, (erule conjE)+)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   119
  fix xs
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   120
  assume "list_all (\<lambda>x. x < Suc i) xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   121
  assume "is_path' r j xs k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   122
  assume "\<not> list_all (\<lambda>x. x < i) xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   123
  show "(\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i) \<and>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   124
    (\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   125
  proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   126
    show "\<And>j. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   127
      \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   128
    \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i" (is "PROP ?ih xs")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   129
    proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   130
      case Nil
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   131
      thus ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   132
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   133
      case (Cons a as j)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   134
      show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   135
      proof (cases "a=i")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   136
      	case True
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   137
      	show ?thesis
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   138
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   139
	  from True and Cons have "r j i = T" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   140
	  thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   141
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   142
      next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   143
      	case False
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   144
      	have "PROP ?ih as" by (rule Cons)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   145
      	then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   146
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   147
	  from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   148
	  from Cons show "is_path' r a as k" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   149
	  from Cons and False show "\<not> list_all (\<lambda>x. x < i) as"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   150
	    by (simp, arith)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   151
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   152
      	show ?thesis
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   153
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   154
	  from Cons False ys
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   155
	  show "list_all (\<lambda>x. x < i) (a # ys) \<and> is_path' r j (a # ys) i"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   156
	    by (simp, arith)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   157
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   158
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   159
    qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   160
    show "\<And>k. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   161
      \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   162
      \<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k" (is "PROP ?ih xs")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   163
    proof (induct xs rule: rev_induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   164
      case Nil
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   165
      thus ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   166
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   167
      case (snoc a as k)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   168
      show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   169
      proof (cases "a=i")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   170
      	case True
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   171
      	show ?thesis
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   172
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   173
	  from True and snoc have "r i k = T" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   174
	  thus "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   175
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   176
      next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   177
      	case False
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   178
      	have "PROP ?ih as" by (rule snoc)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   179
      	then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   180
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   181
	  from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   182
	  from snoc show "is_path' r j as a" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   183
	  from snoc and False show "\<not> list_all (\<lambda>x. x < i) as"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   184
	    by (simp, arith)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   185
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   186
      	show ?thesis
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   187
      	proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   188
	  from snoc False ys
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   189
	  show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   190
	    by (simp, arith)  
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   191
      	qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   192
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   193
    qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   194
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   195
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   196
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   197
theorem lemma5':
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   198
  "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   199
   \<not> (\<forall>q. \<not> is_path r q i j i) \<and> \<not> (\<forall>q'. \<not> is_path r q' i i k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   200
  by (rules dest: lemma5)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   201
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   202
theorem warshall: 
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   203
  "\<And>j k. \<not> (\<exists>p. is_path r p i j k) \<or> (\<exists>p. is_path r p i j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   204
proof (induct i)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   205
  case (0 j k)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   206
  show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   207
  proof (cases "r j k")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   208
    assume "r j k = T"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   209
    hence "is_path r (j, [], k) 0 j k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   210
      by (simp add: is_path_def)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   211
    hence "\<exists>p. is_path r p 0 j k" ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   212
    thus ?thesis ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   213
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   214
    assume "r j k = F"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   215
    hence "r j k ~= T" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   216
    hence "\<not> (\<exists>p. is_path r p 0 j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   217
      by (rules dest: lemma2)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   218
    thus ?thesis ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   219
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   220
next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   221
  case (Suc i j k)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   222
  thus ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   223
  proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   224
    assume h1: "\<not> (\<exists>p. is_path r p i j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   225
    from Suc show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   226
    proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   227
      assume "\<not> (\<exists>p. is_path r p i j i)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   228
      with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   229
	by (rules dest: lemma5')
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   230
      thus ?case ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   231
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   232
      assume "\<exists>p. is_path r p i j i"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   233
      then obtain p where h2: "is_path r p i j i" ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   234
      from Suc show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   235
      proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   236
	assume "\<not> (\<exists>p. is_path r p i i k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   237
	with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   238
	  by (rules dest: lemma5')
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   239
	thus ?case ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   240
      next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   241
	assume "\<exists>q. is_path r q i i k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   242
	then obtain q where "is_path r q i i k" ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   243
	with h2 have "is_path r (conc p q) (Suc i) j k" 
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   244
	  by (rule lemma3)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   245
	hence "\<exists>pq. is_path r pq (Suc i) j k" ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   246
	thus ?case ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   247
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   248
    qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   249
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   250
    assume "\<exists>p. is_path r p i j k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   251
    hence "\<exists>p. is_path r p (Suc i) j k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   252
      by (rules intro: lemma1)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   253
    thus ?case ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   254
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   255
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   256
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   257
extract warshall
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   258
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   259
text {*
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   260
  The program extracted from the above proof looks as follows
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   261
  @{thm [display] warshall_def [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   262
  The corresponding correctness theorem is
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   263
  @{thm [display] warshall_correctness [no_vars]}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   264
*}
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   265
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   266
end
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   267