diff -r bb18eed53ed6 -r a1119cf551e8 src/HOL/Nitpick_Examples/Manual_Nits.thy --- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Aug 13 14:20:22 2013 +0200 +++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Aug 13 16:25:47 2013 +0200 @@ -250,25 +250,25 @@ "loose (Lam t) k = loose t (Suc k)" | "loose (App t u) k = (loose t k \ loose u k)" -primrec subst\<^isub>1 where -"subst\<^isub>1 \ (Var j) = \ j" | -"subst\<^isub>1 \ (Lam t) = - Lam (subst\<^isub>1 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 1) t)" | -"subst\<^isub>1 \ (App t u) = App (subst\<^isub>1 \ t) (subst\<^isub>1 \ u)" +primrec subst\<^sub>1 where +"subst\<^sub>1 \ (Var j) = \ j" | +"subst\<^sub>1 \ (Lam t) = + Lam (subst\<^sub>1 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 1) t)" | +"subst\<^sub>1 \ (App t u) = App (subst\<^sub>1 \ t) (subst\<^sub>1 \ u)" -lemma "\ loose t 0 \ subst\<^isub>1 \ t = t" +lemma "\ loose t 0 \ subst\<^sub>1 \ t = t" nitpick [verbose, expect = genuine] -nitpick [eval = "subst\<^isub>1 \ t", expect = genuine] +nitpick [eval = "subst\<^sub>1 \ t", expect = genuine] (* nitpick [dont_box, expect = unknown] *) oops -primrec subst\<^isub>2 where -"subst\<^isub>2 \ (Var j) = \ j" | -"subst\<^isub>2 \ (Lam t) = - Lam (subst\<^isub>2 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 0) t)" | -"subst\<^isub>2 \ (App t u) = App (subst\<^isub>2 \ t) (subst\<^isub>2 \ u)" +primrec subst\<^sub>2 where +"subst\<^sub>2 \ (Var j) = \ j" | +"subst\<^sub>2 \ (Lam t) = + Lam (subst\<^sub>2 (\n. case n of 0 \ Var 0 | Suc m \ lift (\ m) 0) t)" | +"subst\<^sub>2 \ (App t u) = App (subst\<^sub>2 \ t) (subst\<^sub>2 \ u)" -lemma "\ loose t 0 \ subst\<^isub>2 \ t = t" +lemma "\ loose t 0 \ subst\<^sub>2 \ t = t" nitpick [card = 1\5, expect = none] sorry @@ -354,73 +354,73 @@ datatype alphabet = a | b -inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where - "[] \ S\<^isub>1" -| "w \ A\<^isub>1 \ b # w \ S\<^isub>1" -| "w \ B\<^isub>1 \ a # w \ S\<^isub>1" -| "w \ S\<^isub>1 \ a # w \ A\<^isub>1" -| "w \ S\<^isub>1 \ b # w \ S\<^isub>1" -| "\v \ B\<^isub>1; v \ B\<^isub>1\ \ a # v @ w \ B\<^isub>1" +inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where + "[] \ S\<^sub>1" +| "w \ A\<^sub>1 \ b # w \ S\<^sub>1" +| "w \ B\<^sub>1 \ a # w \ S\<^sub>1" +| "w \ S\<^sub>1 \ a # w \ A\<^sub>1" +| "w \ S\<^sub>1 \ b # w \ S\<^sub>1" +| "\v \ B\<^sub>1; v \ B\<^sub>1\ \ a # v @ w \ B\<^sub>1" -theorem S\<^isub>1_sound: -"w \ S\<^isub>1 \ length [x \ w. x = a] = length [x \ w. x = b]" +theorem S\<^sub>1_sound: +"w \ S\<^sub>1 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [expect = genuine] oops -inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where - "[] \ S\<^isub>2" -| "w \ A\<^isub>2 \ b # w \ S\<^isub>2" -| "w \ B\<^isub>2 \ a # w \ S\<^isub>2" -| "w \ S\<^isub>2 \ a # w \ A\<^isub>2" -| "w \ S\<^isub>2 \ b # w \ B\<^isub>2" -| "\v \ B\<^isub>2; v \ B\<^isub>2\ \ a # v @ w \ B\<^isub>2" +inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where + "[] \ S\<^sub>2" +| "w \ A\<^sub>2 \ b # w \ S\<^sub>2" +| "w \ B\<^sub>2 \ a # w \ S\<^sub>2" +| "w \ S\<^sub>2 \ a # w \ A\<^sub>2" +| "w \ S\<^sub>2 \ b # w \ B\<^sub>2" +| "\v \ B\<^sub>2; v \ B\<^sub>2\ \ a # v @ w \ B\<^sub>2" -theorem S\<^isub>2_sound: -"w \ S\<^isub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" +theorem S\<^sub>2_sound: +"w \ S\<^sub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [expect = genuine] oops -inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where - "[] \ S\<^isub>3" -| "w \ A\<^isub>3 \ b # w \ S\<^isub>3" -| "w \ B\<^isub>3 \ a # w \ S\<^isub>3" -| "w \ S\<^isub>3 \ a # w \ A\<^isub>3" -| "w \ S\<^isub>3 \ b # w \ B\<^isub>3" -| "\v \ B\<^isub>3; w \ B\<^isub>3\ \ a # v @ w \ B\<^isub>3" +inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where + "[] \ S\<^sub>3" +| "w \ A\<^sub>3 \ b # w \ S\<^sub>3" +| "w \ B\<^sub>3 \ a # w \ S\<^sub>3" +| "w \ S\<^sub>3 \ a # w \ A\<^sub>3" +| "w \ S\<^sub>3 \ b # w \ B\<^sub>3" +| "\v \ B\<^sub>3; w \ B\<^sub>3\ \ a # v @ w \ B\<^sub>3" -theorem S\<^isub>3_sound: -"w \ S\<^isub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" +theorem S\<^sub>3_sound: +"w \ S\<^sub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [card = 1\5, expect = none] sorry -theorem S\<^isub>3_complete: -"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>3" +theorem S\<^sub>3_complete: +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^sub>3" nitpick [expect = genuine] oops -inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where - "[] \ S\<^isub>4" -| "w \ A\<^isub>4 \ b # w \ S\<^isub>4" -| "w \ B\<^isub>4 \ a # w \ S\<^isub>4" -| "w \ S\<^isub>4 \ a # w \ A\<^isub>4" -| "\v \ A\<^isub>4; w \ A\<^isub>4\ \ b # v @ w \ A\<^isub>4" -| "w \ S\<^isub>4 \ b # w \ B\<^isub>4" -| "\v \ B\<^isub>4; w \ B\<^isub>4\ \ a # v @ w \ B\<^isub>4" +inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where + "[] \ S\<^sub>4" +| "w \ A\<^sub>4 \ b # w \ S\<^sub>4" +| "w \ B\<^sub>4 \ a # w \ S\<^sub>4" +| "w \ S\<^sub>4 \ a # w \ A\<^sub>4" +| "\v \ A\<^sub>4; w \ A\<^sub>4\ \ b # v @ w \ A\<^sub>4" +| "w \ S\<^sub>4 \ b # w \ B\<^sub>4" +| "\v \ B\<^sub>4; w \ B\<^sub>4\ \ a # v @ w \ B\<^sub>4" -theorem S\<^isub>4_sound: -"w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +theorem S\<^sub>4_sound: +"w \ S\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" nitpick [card = 1\5, expect = none] sorry -theorem S\<^isub>4_complete: -"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>4" +theorem S\<^sub>4_complete: +"length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^sub>4" nitpick [card = 1\5, expect = none] sorry -theorem S\<^isub>4_A\<^isub>4_B\<^isub>4_sound_and_complete: -"w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" -"w \ A\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b] + 1" -"w \ B\<^isub>4 \ length [x \ w. x = b] = length [x \ w. x = a] + 1" +theorem S\<^sub>4_A\<^sub>4_B\<^sub>4_sound_and_complete: +"w \ S\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" +"w \ A\<^sub>4 \ length [x \ w. x = a] = length [x \ w. x = b] + 1" +"w \ B\<^sub>4 \ length [x \ w. x = b] = length [x \ w. x = a] + 1" nitpick [card = 1\5, expect = none] sorry @@ -442,11 +442,11 @@ primrec left where "left \ = \" | -"left (N _ _ t\<^isub>1 _) = t\<^isub>1" +"left (N _ _ t\<^sub>1 _) = t\<^sub>1" primrec right where "right \ = \" | -"right (N _ _ _ t\<^isub>2) = t\<^isub>2" +"right (N _ _ _ t\<^sub>2) = t\<^sub>2" fun wf where "wf \ = True" | @@ -484,31 +484,31 @@ nitpick [card = 1\5, expect = none] sorry -primrec insort\<^isub>1 where -"insort\<^isub>1 \ x = N x 1 \ \" | -"insort\<^isub>1 (N y k t u) x = - (* (split \ skew) *) (N y k (if x < y then insort\<^isub>1 t x else t) - (if x > y then insort\<^isub>1 u x else u))" +primrec insort\<^sub>1 where +"insort\<^sub>1 \ x = N x 1 \ \" | +"insort\<^sub>1 (N y k t u) x = + (* (split \ skew) *) (N y k (if x < y then insort\<^sub>1 t x else t) + (if x > y then insort\<^sub>1 u x else u))" -theorem wf_insort\<^isub>1: "wf t \ wf (insort\<^isub>1 t x)" +theorem wf_insort\<^sub>1: "wf t \ wf (insort\<^sub>1 t x)" nitpick [expect = genuine] oops -theorem wf_insort\<^isub>1_nat: "wf t \ wf (insort\<^isub>1 t (x\nat))" -nitpick [eval = "insort\<^isub>1 t x", expect = genuine] +theorem wf_insort\<^sub>1_nat: "wf t \ wf (insort\<^sub>1 t (x\nat))" +nitpick [eval = "insort\<^sub>1 t x", expect = genuine] oops -primrec insort\<^isub>2 where -"insort\<^isub>2 \ x = N x 1 \ \" | -"insort\<^isub>2 (N y k t u) x = - (split \ skew) (N y k (if x < y then insort\<^isub>2 t x else t) - (if x > y then insort\<^isub>2 u x else u))" +primrec insort\<^sub>2 where +"insort\<^sub>2 \ x = N x 1 \ \" | +"insort\<^sub>2 (N y k t u) x = + (split \ skew) (N y k (if x < y then insort\<^sub>2 t x else t) + (if x > y then insort\<^sub>2 u x else u))" -theorem wf_insort\<^isub>2: "wf t \ wf (insort\<^isub>2 t x)" +theorem wf_insort\<^sub>2: "wf t \ wf (insort\<^sub>2 t x)" nitpick [card = 1\5, expect = none] sorry -theorem dataset_insort\<^isub>2: "dataset (insort\<^isub>2 t x) = {x} \ dataset t" +theorem dataset_insort\<^sub>2: "dataset (insort\<^sub>2 t x) = {x} \ dataset t" nitpick [card = 1\5, expect = none] sorry