src/HOL/Proofs/Extraction/Warshall.thy
author wenzelm
Mon, 13 Apr 2020 22:08:14 +0200
changeset 71751 abf3e80bd815
parent 67320 6afba546f0e5
child 71989 bad75618fb82
permissions -rw-r--r--
tuned NEWS;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39157
b98909faaea8 more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
wenzelm
parents: 37599
diff changeset
     1
(*  Title:      HOL/Proofs/Extraction/Warshall.thy
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     2
    Author:     Stefan Berghofer, TU Muenchen
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     3
*)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     4
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 60595
diff changeset
     5
section \<open>Warshall's algorithm\<close>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
     6
16761
99549528ce76 proof needed updating because of arith
nipkow
parents: 16417
diff changeset
     7
theory Warshall
67320
6afba546f0e5 updated dependencies + compile
blanchet
parents: 66453
diff changeset
     8
imports "HOL-Library.Realizers"
16761
99549528ce76 proof needed updating because of arith
nipkow
parents: 16417
diff changeset
     9
begin
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    10
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 60595
diff changeset
    11
text \<open>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    12
  Derivation of Warshall's algorithm using program extraction,
58622
aa99568f56de more antiquotations;
wenzelm
parents: 58372
diff changeset
    13
  based on Berger, Schwichtenberg and Seisenberger @{cite "Berger-JAR-2001"}.
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 60595
diff changeset
    14
\<close>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    15
58310
91ea607a34d8 updated news
blanchet
parents: 58279
diff changeset
    16
datatype b = T | F
58272
61d94335ef6c ported HOL-Proofs-Extraction to new datatypes
blanchet
parents: 51272
diff changeset
    17
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    18
primrec is_path' :: "('a \<Rightarrow> 'a \<Rightarrow> b) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 'a \<Rightarrow> bool"
25976
11c6811f232c modernized primrec;
wenzelm
parents: 23373
diff changeset
    19
where
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    20
  "is_path' r x [] z \<longleftrightarrow> r x z = T"
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    21
| "is_path' r x (y # ys) z \<longleftrightarrow> r x y = T \<and> is_path' r y ys z"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    22
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    23
definition is_path :: "(nat \<Rightarrow> nat \<Rightarrow> b) \<Rightarrow> (nat * nat list * nat) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    24
  where "is_path r p i j k \<longleftrightarrow>
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    25
    fst p = j \<and> snd (snd p) = k \<and>
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    26
    list_all (\<lambda>x. x < i) (fst (snd p)) \<and>
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    27
    is_path' r (fst p) (fst (snd p)) (snd (snd p))"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    28
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    29
definition conc :: "'a \<times> 'a list \<times> 'a \<Rightarrow> 'a \<times> 'a list \<times> 'a \<Rightarrow> 'a \<times> 'a list * 'a"
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    30
  where "conc p q = (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))"
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    31
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    32
theorem is_path'_snoc [simp]: "\<And>x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \<and> r y z = T)"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    33
  by (induct ys) simp+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    34
37599
b8e3400dab19 tuned syntax
haftmann
parents: 32960
diff changeset
    35
theorem list_all_scoc [simp]: "list_all P (xs @ [x]) \<longleftrightarrow> P x \<and> list_all P xs"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    36
  by (induct xs) (simp+, iprover)
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    37
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    38
theorem list_all_lemma: "list_all P xs \<Longrightarrow> (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> list_all Q xs"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    39
proof -
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    40
  assume PQ: "\<And>x. P x \<Longrightarrow> Q x"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    41
  show "list_all P xs \<Longrightarrow> list_all Q xs"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    42
  proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    43
    case Nil
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    44
    show ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    45
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    46
    case (Cons y ys)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    47
    then have Py: "P y" by simp
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    48
    from Cons have Pys: "list_all P ys" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    49
    show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    50
      by simp (rule conjI PQ Py Cons Pys)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    51
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    52
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    53
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    54
theorem lemma1: "\<And>p. is_path r p i j k \<Longrightarrow> is_path r p (Suc i) j k"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    55
  unfolding is_path_def
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    56
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    57
  apply (erule conjE)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    58
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    59
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    60
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    61
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    62
theorem lemma2: "\<And>p. is_path r p 0 j k \<Longrightarrow> r j k = T"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    63
  unfolding is_path_def
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    64
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    65
  apply (case_tac "aa")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    66
  apply simp+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    67
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    68
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    69
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
    70
  is_path' r j (xs @ i # ys) k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    71
proof -
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    72
  assume pys: "is_path' r i ys k"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    73
  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
    74
  proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    75
    case (Nil j)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    76
    then have "r j i = T" by simp
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    77
    with pys show ?case by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    78
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    79
    case (Cons z zs j)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    80
    then have jzr: "r j z = T" by simp
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    81
    from Cons have pzs: "is_path' r z zs i" by simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    82
    show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    83
      by simp (rule conjI jzr Cons pzs)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    84
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    85
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    86
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    87
theorem lemma3:
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    88
  "\<And>p q. is_path r p i j i \<Longrightarrow> is_path r q i i k \<Longrightarrow>
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
    89
    is_path r (conc p q) (Suc i) j k"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    90
  apply (unfold is_path_def conc_def)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    91
  apply (simp cong add: conj_cong add: split_paired_all)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    92
  apply (erule conjE)+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    93
  apply (rule conjI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    94
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    95
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    96
  apply (rule conjI)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    97
  apply (erule list_all_lemma)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    98
  apply simp
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
    99
  apply (rule is_path'_conc)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   100
  apply assumption+
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   101
  done
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   102
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   103
theorem lemma5:
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   104
  "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow>
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   105
    (\<exists>q. is_path r q i j i) \<and> (\<exists>q'. is_path r q' i i k)"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   106
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
   107
  fix xs
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 17604
diff changeset
   108
  assume asms:
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 17604
diff changeset
   109
    "list_all (\<lambda>x. x < Suc i) xs"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 17604
diff changeset
   110
    "is_path' r j xs k"
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 17604
diff changeset
   111
    "\<not> list_all (\<lambda>x. x < i) xs"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   112
  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
   113
    (\<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
   114
  proof
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 58889
diff changeset
   115
    have "\<And>j. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   116
      \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   117
    \<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
   118
    proof (induct xs)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   119
      case Nil
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   120
      then show ?case by simp
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   121
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   122
      case (Cons a as j)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   123
      show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   124
      proof (cases "a=i")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   125
        case True
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   126
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   127
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   128
          from True and Cons have "r j i = T" by simp
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   129
          then show "list_all (\<lambda>x. x < i) [] \<and> is_path' r j [] i" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   130
        qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   131
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   132
        case False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   133
        have "PROP ?ih as" by (rule Cons)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   134
        then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r a ys i"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   135
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   136
          from Cons show "list_all (\<lambda>x. x < Suc i) as" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   137
          from Cons show "is_path' r a as k" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   138
          from Cons and False show "\<not> list_all (\<lambda>x. x < i) as" by (simp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   139
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   140
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   141
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   142
          from Cons False ys
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   143
          show "list_all (\<lambda>x. x<i) (a#ys) \<and> is_path' r j (a#ys) i" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   144
        qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   145
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   146
    qed
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 58889
diff changeset
   147
    from this asms show "\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r j ys i" .
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 58889
diff changeset
   148
    have "\<And>k. list_all (\<lambda>x. x < Suc i) xs \<Longrightarrow> is_path' r j xs k \<Longrightarrow>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   149
      \<not> list_all (\<lambda>x. x < i) xs \<Longrightarrow>
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   150
      \<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
   151
    proof (induct xs rule: rev_induct)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   152
      case Nil
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   153
      then show ?case by simp
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   154
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   155
      case (snoc a as k)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   156
      show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   157
      proof (cases "a=i")
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   158
        case True
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   159
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   160
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   161
          from True and snoc have "r i k = T" by simp
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   162
          then show "list_all (\<lambda>x. x < i) [] \<and> is_path' r i [] k" by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   163
        qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   164
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   165
        case False
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   166
        have "PROP ?ih as" by (rule snoc)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   167
        then obtain ys where ys: "list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys a"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   168
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   169
          from snoc show "list_all (\<lambda>x. x < Suc i) as" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   170
          from snoc show "is_path' r j as a" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   171
          from snoc and False show "\<not> list_all (\<lambda>x. x < i) as" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   172
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   173
        show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   174
        proof
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   175
          from snoc False ys
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   176
          show "list_all (\<lambda>x. x < i) (ys @ [a]) \<and> is_path' r i (ys @ [a]) k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   177
            by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   178
        qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   179
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   180
    qed
60595
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 58889
diff changeset
   181
    from this asms show "\<exists>ys. list_all (\<lambda>x. x < i) ys \<and> is_path' r i ys k" .
804dfdc82835 premises in 'show' are treated like 'assume';
wenzelm
parents: 58889
diff changeset
   182
  qed
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   183
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   184
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   185
theorem lemma5':
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   186
  "\<And>p. is_path r p (Suc i) j k \<Longrightarrow> \<not> is_path r p i j k \<Longrightarrow>
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   187
    \<not> (\<forall>q. \<not> is_path r q i j i) \<and> \<not> (\<forall>q'. \<not> is_path r q' i i k)"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 16761
diff changeset
   188
  by (iprover dest: lemma5)
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   189
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   190
theorem warshall: "\<And>j k. \<not> (\<exists>p. is_path r p i j k) \<or> (\<exists>p. is_path r p i j k)"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   191
proof (induct i)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   192
  case (0 j k)
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   193
  show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   194
  proof (cases "r j k")
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   195
    assume "r j k = T"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   196
    then have "is_path r (j, [], k) 0 j k"
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   197
      by (simp add: is_path_def)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   198
    then have "\<exists>p. is_path r p 0 j k" ..
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   199
    then show ?thesis ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   200
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   201
    assume "r j k = F"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   202
    then have "r j k \<noteq> T" by simp
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   203
    then have "\<not> (\<exists>p. is_path r p 0 j k)"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 16761
diff changeset
   204
      by (iprover dest: lemma2)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   205
    then show ?thesis ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   206
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   207
next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   208
  case (Suc i j k)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   209
  then show ?case
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   210
  proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   211
    assume h1: "\<not> (\<exists>p. is_path r p i j k)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   212
    from Suc show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   213
    proof
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   214
      assume "\<not> (\<exists>p. is_path r p i j i)"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   215
      with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   216
        by (iprover dest: lemma5')
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   217
      then show ?case ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   218
    next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   219
      assume "\<exists>p. is_path r p i j i"
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   220
      then obtain p where h2: "is_path r p i j i" ..
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   221
      from Suc show ?case
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   222
      proof
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   223
        assume "\<not> (\<exists>p. is_path r p i i k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   224
        with h1 have "\<not> (\<exists>p. is_path r p (Suc i) j k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   225
          by (iprover dest: lemma5')
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   226
        then show ?case ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   227
      next
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   228
        assume "\<exists>q. is_path r q i i k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   229
        then obtain q where "is_path r q i i k" ..
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   230
        with h2 have "is_path r (conc p q) (Suc i) j k" 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 29823
diff changeset
   231
          by (rule lemma3)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   232
        then have "\<exists>pq. is_path r pq (Suc i) j k" ..
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   233
        then show ?case ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   234
      qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   235
    qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   236
  next
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   237
    assume "\<exists>p. is_path r p i j k"
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   238
    then have "\<exists>p. is_path r p (Suc i) j k"
17604
5f30179fbf44 rules -> iprover
nipkow
parents: 16761
diff changeset
   239
      by (iprover intro: lemma1)
63361
d10eab0672f9 misc tuning and modernization;
wenzelm
parents: 61986
diff changeset
   240
    then show ?case ..
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   241
  qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   242
qed
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   243
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   244
extract warshall
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   245
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 60595
diff changeset
   246
text \<open>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   247
  The program extracted from the above proof looks as follows
13671
eec2582923f6 Eta contraction is now switched off when printing extracted program.
berghofe
parents: 13471
diff changeset
   248
  @{thm [display, eta_contract=false] warshall_def [no_vars]}
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   249
  The corresponding correctness theorem is
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   250
  @{thm [display] warshall_correctness [no_vars]}
61986
2461779da2b8 isabelle update_cartouches -c -t;
wenzelm
parents: 60595
diff changeset
   251
\<close>
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   252
51272
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 39157
diff changeset
   253
ML_val "@{code warshall}"
27982
2aaa4a5569a6 default replaces arbitrary
haftmann
parents: 25976
diff changeset
   254
13405
d20a4e67afc8 Examples for program extraction in HOL.
berghofe
parents:
diff changeset
   255
end