# HG changeset patch # User wenzelm # Date 1201301433 -3600 # Node ID 11c6811f232c3a8c5779c207e96121a333ab05ed # Parent bcb1e9b7644b05040f4edbd728e982e343e9cc61 modernized primrec; diff -r bcb1e9b7644b -r 11c6811f232c src/HOL/Extraction/Euclid.thy --- a/src/HOL/Extraction/Euclid.thy Fri Jan 25 23:05:27 2008 +0100 +++ b/src/HOL/Extraction/Euclid.thy Fri Jan 25 23:50:33 2008 +0100 @@ -128,10 +128,10 @@ lemma dvd_prod [iff]: "n dvd prod (n # ns)" by simp -consts fact :: "nat \ nat" ("(_!)" [1000] 999) -primrec - "0! = 1" - "(Suc n)! = n! * Suc n" +primrec fact :: "nat \ nat" ("(_!)" [1000] 999) +where + "0! = 1" + | "(Suc n)! = n! * Suc n" lemma fact_greater_0 [iff]: "0 < n!" by (induct n) simp_all diff -r bcb1e9b7644b -r 11c6811f232c src/HOL/Extraction/Higman.thy --- a/src/HOL/Extraction/Higman.thy Fri Jan 25 23:05:27 2008 +0100 +++ b/src/HOL/Extraction/Higman.thy Fri Jan 25 23:50:33 2008 +0100 @@ -234,12 +234,11 @@ qed qed -consts +primrec is_prefix :: "'a list \ (nat \ 'a) \ bool" - -primrec - "is_prefix [] f = True" - "is_prefix (x # xs) f = (x = f (length xs) \ is_prefix xs f)" +where + "is_prefix [] f = True" + | "is_prefix (x # xs) f = (x = f (length xs) \ is_prefix xs f)" theorem L_idx: assumes L: "L w ws" diff -r bcb1e9b7644b -r 11c6811f232c src/HOL/Extraction/Warshall.thy --- a/src/HOL/Extraction/Warshall.thy Fri Jan 25 23:05:27 2008 +0100 +++ b/src/HOL/Extraction/Warshall.thy Fri Jan 25 23:50:33 2008 +0100 @@ -16,22 +16,24 @@ datatype b = T | F -consts +primrec is_path' :: "('a \ 'a \ b) \ 'a \ 'a list \ 'a \ bool" +where + "is_path' r x [] z = (r x z = T)" + | "is_path' r x (y # ys) z = (r x y = T \ is_path' r y ys z)" -primrec - "is_path' r x [] z = (r x z = T)" - "is_path' r x (y # ys) z = (r x y = T \ is_path' r y ys z)" - -constdefs +definition is_path :: "(nat \ nat \ b) \ (nat * nat list * nat) \ nat \ nat \ nat \ bool" - "is_path r p i j k == fst p = j \ snd (snd p) = k \ +where + "is_path r p i j k \ fst p = j \ snd (snd p) = k \ list_all (\x. x < i) (fst (snd p)) \ is_path' r (fst p) (fst (snd p)) (snd (snd p))" +definition conc :: "('a * 'a list * 'a) \ ('a * 'a list * 'a) \ ('a * 'a list * 'a)" - "conc p q == (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))" +where + "conc p q = (fst p, fst (snd p) @ fst q # fst (snd q), snd (snd q))" theorem is_path'_snoc [simp]: "\x. is_path' r x (ys @ [y]) z = (is_path' r x ys y \ r y z = T)"