diff -r 997aa3a3e4bb -r c9f428269b38 src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy --- a/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Tue Feb 23 10:02:14 2010 +0100 +++ b/src/HOL/ex/Predicate_Compile_Quickcheck_ex.thy Tue Feb 23 13:36:15 2010 +0100 @@ -1,45 +1,39 @@ theory Predicate_Compile_Quickcheck_ex imports Predicate_Compile_Quickcheck - Predicate_Compile_Alternative_Defs begin -ML {* Predicate_Compile_Alternative_Defs.get *} - section {* Sets *} -(* + lemma "x \ {(1::nat)} ==> False" -quickcheck[generator=predicate_compile, iterations=10] +quickcheck[generator=predicate_compile_wo_ff, iterations=10] oops -*) -(* TODO: some error with doubled negation *) -(* + lemma "x \ {Suc 0, Suc (Suc 0)} ==> x \ Suc 0" -quickcheck[generator=predicate_compile] +quickcheck[generator=predicate_compile_wo_ff] oops -*) -(* + lemma "x \ {Suc 0, Suc (Suc 0)} ==> x = Suc 0" -quickcheck[generator=predicate_compile] +quickcheck[generator=predicate_compile_wo_ff] oops -*) + lemma "x \ {Suc 0, Suc (Suc 0)} ==> x <= Suc 0" -(*quickcheck[generator=predicate_compile]*) +quickcheck[generator=predicate_compile_wo_ff] oops section {* Numerals *} -(* + lemma "x \ {1, 2, (3::nat)} ==> x = 1 \ x = 2" -quickcheck[generator=predicate_compile] +quickcheck[generator=predicate_compile_wo_ff] oops -*) + lemma "x \ {1, 2, (3::nat)} ==> x < 3" -(*quickcheck[generator=predicate_compile]*) +quickcheck[generator=predicate_compile_wo_ff] oops lemma "x \ {1, 2} \ {3, 4} ==> x = (1::nat) \ x = (2::nat)" -(*quickcheck[generator=predicate_compile]*) +quickcheck[generator=predicate_compile_wo_ff] oops section {* Context Free Grammar *} @@ -53,33 +47,15 @@ | "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" -(* -code_pred [random_dseq inductify] "S\<^isub>1p" . -*) -(*thm B\<^isub>1p.random_dseq_equation*) -(* -values [random_dseq 2, 2, 4] 10 "{x. S\<^isub>1p x}" -values [random_dseq 1, 1, 5] 20 "{x. S\<^isub>1p x}" -ML {* set ML_Context.trace *} -*) -ML {* set Toplevel.debug *} -(* -quickcheck[generator = predicate_compile, size = 10, iterations = 1] -oops -*) -ML {* Spec_Rules.get *} -ML {* Item_Net.retrieve *} -local_setup {* Local_Theory.checkpoint *} -ML {* Predicate_Compile_Data.get_specification @{theory} @{term "append"} *} lemma - "w \ S\<^isub>1p \ w = []" -quickcheck[generator = predicate_compile, iterations=1] + "w \ S\<^isub>1 \ w = []" +quickcheck[generator = predicate_compile_ff_nofs, iterations=1] oops theorem S\<^isub>1_sound: -"w \ S\<^isub>1p \ length [x \ w. x = a] = length [x \ w. x = b]" -quickcheck[generator=predicate_compile, size=15] +"w \ S\<^isub>1 \ length [x \ w. x = a] = length [x \ w. x = b]" +quickcheck[generator=predicate_compile_ff_nofs, size=15] oops @@ -90,7 +66,7 @@ | "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" - +(* code_pred [random_dseq inductify] S\<^isub>2 . thm S\<^isub>2.random_dseq_equation thm A\<^isub>2.random_dseq_equation @@ -118,10 +94,10 @@ "w \ S\<^isub>2 ==> length [x \ w. x = a] <= Suc (Suc 0)" quickcheck[generator=predicate_compile, size = 10, iterations = 1] oops - +*) theorem S\<^isub>2_sound: "w \ S\<^isub>2 \ length [x \ w. x = a] = length [x \ w. x = b]" -quickcheck[generator=predicate_compile, size=15, iterations=1] +quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10] oops inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where @@ -141,17 +117,17 @@ lemma S\<^isub>3_sound: "w \ S\<^isub>3 \ length [x \ w. x = a] = length [x \ w. x = b]" -quickcheck[generator=predicate_compile, size=10, iterations=10] +quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10] oops lemma "\ (length w > 2) \ \ (length [x \ w. x = a] = length [x \ w. x = b])" -quickcheck[size=10, generator = predicate_compile] +quickcheck[size=10, generator = predicate_compile_ff_fs] oops theorem S\<^isub>3_complete: "length [x \ w. x = a] = length [x \ w. b = x] \ w \ S\<^isub>3" (*quickcheck[generator=SML]*) -quickcheck[generator=predicate_compile, size=10, iterations=100] +quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100] oops @@ -166,20 +142,23 @@ theorem S\<^isub>4_sound: "w \ S\<^isub>4 \ length [x \ w. x = a] = length [x \ w. x = b]" -quickcheck[generator = predicate_compile, size=5, iterations=1] +quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1] oops theorem S\<^isub>4_complete: "length [x \ w. x = a] = length [x \ w. x = b] \ w \ S\<^isub>4" -quickcheck[generator = predicate_compile, size=5, iterations=1] +quickcheck[generator = predicate_compile_ff_nofs, size=5, iterations=1] oops -hide const b +hide const a b subsection {* Lexicographic order *} +(* TODO *) +(* lemma "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r" - +oops +*) subsection {* IMP *} types @@ -208,7 +187,7 @@ lemma "exec c s s' ==> exec (Seq c c) s s'" -quickcheck[generator = predicate_compile, size=3, iterations=1] +(*quickcheck[generator = predicate_compile_wo_ff, size=2, iterations=10]*) oops subsection {* Lambda *} @@ -263,28 +242,9 @@ lemma "\ \ t : U \ t \\<^sub>\ t' \ \ \ t' : U" -quickcheck[generator = predicate_compile, size = 7, iterations = 10] +quickcheck[generator = predicate_compile_ff_fs, size = 7, iterations = 10] oops -(* -code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing . -thm typing.equation - -code_pred (modes: i => i => bool, i => o => bool as reduce') beta . -thm beta.equation - -values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \\<^sub>\ x}" - -definition "reduce t = Predicate.the (reduce' t)" - -value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))" - -code_pred [random] typing . -code_pred [random_dseq] typing . - -(*values [random] 1 "{(\, t, T). \ \ t : T}" -*)*) - subsection {* JAD *} definition matrix :: "('a :: semiring_0) list list \ nat \ nat \ bool" where @@ -300,9 +260,17 @@ lemma [code_pred_intro]: "matrix [] 0 m" "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m" -sorry +proof - + show "matrix [] 0 m" unfolding matrix_def by auto +next + show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m" + unfolding matrix_def by auto +qed -code_pred [random_dseq inductify] matrix sorry +code_pred [random_dseq inductify] matrix + apply (cases x) + unfolding matrix_def apply fastsimp + apply fastsimp done values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}" @@ -344,10 +312,10 @@ definition "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)" - +(* definition "transpose M = [map (\ xs. xs ! i) (takeWhile (\ xs. i < length xs) M). i \ [0 ..< length (M ! 0)]]" - +*) definition "inflate upds = foldr (\ (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)" @@ -356,15 +324,14 @@ definition "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\ (i, x). v ! i * x)))" -ML {* ML_Context.trace := false *} lemma "matrix (M::int list list) rs cs \ False" -quickcheck[generator = predicate_compile, size = 6] +quickcheck[generator = predicate_compile_ff_nofs, size = 6] oops lemma "\ matrix M rs cs ; length v = cs \ \ jad_mv v (jad M) = mv M v" -(*quickcheck[generator = predicate_compile]*) +quickcheck[generator = predicate_compile_wo_ff] oops end \ No newline at end of file