| author | bulwahn | 
| Mon, 10 Oct 2011 16:47:45 +0200 | |
| changeset 45125 | c15b0faeb70a | 
| parent 44890 | 22f665a2e91c | 
| child 53015 | a1119cf551e8 | 
| permissions | -rw-r--r-- | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 1 | theory Predicate_Compile_Quickcheck_Examples | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40924diff
changeset | 2 | imports "~~/src/HOL/Library/Predicate_Compile_Quickcheck" | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 3 | begin | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 4 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 5 | section {* Sets *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 6 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 7 | lemma "x \<in> {(1::nat)} ==> False"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 8 | quickcheck[generator=predicate_compile_wo_ff, iterations=10] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 9 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 10 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 11 | lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x \<noteq> Suc 0"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 12 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 13 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 14 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 15 | lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x = Suc 0"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 16 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 17 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 18 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 19 | lemma "x \<in> {Suc 0, Suc (Suc 0)} ==> x <= Suc 0"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 20 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 21 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 22 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 23 | section {* Numerals *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 24 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 25 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 26 |   "x \<in> {1, 2, (3::nat)} ==> x = 1 \<or> x = 2"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 27 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 28 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 29 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 30 | lemma "x \<in> {1, 2, (3::nat)} ==> x < 3"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 31 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 32 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 33 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 34 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 35 |   "x \<in> {1, 2} \<union> {3, 4} ==> x = (1::nat) \<or> x = (2::nat)"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 36 | quickcheck[generator=predicate_compile_wo_ff] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 37 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 38 | |
| 39650 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 39 | section {* Equivalences *}
 | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 40 | |
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 41 | inductive is_ten :: "nat => bool" | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 42 | where | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 43 | "is_ten 10" | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 44 | |
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 45 | inductive is_eleven :: "nat => bool" | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 46 | where | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 47 | "is_eleven 11" | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 48 | |
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 49 | lemma | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 50 | "is_ten x = is_eleven x" | 
| 40924 | 51 | quickcheck[tester = predicate_compile_wo_ff, iterations = 1, size = 1, expect = counterexample] | 
| 39650 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 52 | oops | 
| 
2a35950ec495
handling equivalences smarter in the predicate compiler
 bulwahn parents: 
39191diff
changeset | 53 | |
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 54 | section {* Context Free Grammar *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 55 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 56 | datatype alphabet = a | b | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 57 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 58 | inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 59 | "[] \<in> S\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 60 | | "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 61 | | "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 62 | | "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 63 | | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 64 | | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 65 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 66 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 67 | "w \<in> S\<^isub>1 \<Longrightarrow> w = []" | 
| 40924 | 68 | quickcheck[tester = predicate_compile_ff_nofs, iterations=1] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 69 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 70 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 71 | theorem S\<^isub>1_sound: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 72 | "w \<in> S\<^isub>1 \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 73 | quickcheck[generator=predicate_compile_ff_nofs, size=15] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 74 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 75 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 76 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 77 | inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 78 | "[] \<in> S\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 79 | | "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 80 | | "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 81 | | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 82 | | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 83 | | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 84 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 85 | code_pred [random_dseq inductify] S\<^isub>2 . | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 86 | thm S\<^isub>2.random_dseq_equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 87 | thm A\<^isub>2.random_dseq_equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 88 | thm B\<^isub>2.random_dseq_equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 89 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 90 | values [random_dseq 1, 2, 8] 10 "{x. S\<^isub>2 x}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 91 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 92 | lemma "w \<in> S\<^isub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 93 | quickcheck[generator=predicate_compile, size=8] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 94 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 95 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 96 | lemma "[x <- w. x = a] = []" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 97 | quickcheck[generator=predicate_compile] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 98 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 99 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 100 | declare list.size(3,4)[code_pred_def] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 101 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 102 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 103 | lemma "length ([x \<leftarrow> w. x = a]) = (0::nat)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 104 | quickcheck[generator=predicate_compile] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 105 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 106 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 107 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 108 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 109 | "w \<in> S\<^isub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 110 | quickcheck[generator=predicate_compile, size = 10, iterations = 1] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 111 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 112 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 113 | theorem S\<^isub>2_sound: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 114 | "w \<in> S\<^isub>2 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 115 | quickcheck[generator=predicate_compile_ff_nofs, size=5, iterations=10] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 116 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 117 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 118 | inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 119 | "[] \<in> S\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 120 | | "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 121 | | "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 122 | | "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 123 | | "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 124 | | "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 125 | |
| 36040 
fcd7bea01a93
adding skip_proof in the examples because proof procedure cannot handle alternative compilations yet
 bulwahn parents: 
35955diff
changeset | 126 | code_pred [inductify, skip_proof] S\<^isub>3 . | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 127 | thm S\<^isub>3.equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 128 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 129 | values 10 "{x. S\<^isub>3 x}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 130 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 131 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 132 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 133 | lemma S\<^isub>3_sound: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 134 | "w \<in> S\<^isub>3 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 135 | quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=10] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 136 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 137 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 138 | lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])" | 
| 40924 | 139 | quickcheck[size=10, tester = predicate_compile_ff_fs] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 140 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 141 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 142 | theorem S\<^isub>3_complete: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 143 | "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> w \<in> S\<^isub>3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 144 | (*quickcheck[generator=SML]*) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 145 | quickcheck[generator=predicate_compile_ff_fs, size=10, iterations=100] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 146 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 147 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 148 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 149 | inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 150 | "[] \<in> S\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 151 | | "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 152 | | "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 153 | | "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 154 | | "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 155 | | "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 156 | | "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 157 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 158 | theorem S\<^isub>4_sound: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 159 | "w \<in> S\<^isub>4 \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]" | 
| 40924 | 160 | quickcheck[tester = predicate_compile_ff_nofs, size=5, iterations=1] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 161 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 162 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 163 | theorem S\<^isub>4_complete: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 164 | "length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> w \<in> S\<^isub>4" | 
| 40924 | 165 | quickcheck[tester = predicate_compile_ff_nofs, size=5, iterations=1] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 166 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 167 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
36040diff
changeset | 168 | hide_const a b | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 169 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 170 | subsection {* Lexicographic order *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 171 | (* TODO *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 172 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 173 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 174 | "(u, v) : lexord r ==> (x @ u, y @ v) : lexord r" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 175 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 176 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 177 | subsection {* IMP *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 178 | |
| 42463 | 179 | type_synonym var = nat | 
| 180 | type_synonym state = "int list" | |
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 181 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 182 | datatype com = | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 183 | Skip | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 184 | Ass var "int" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 185 | Seq com com | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 186 | IF "state list" com com | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 187 | While "state list" com | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 188 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 189 | inductive exec :: "com => state => state => bool" where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 190 | "exec Skip s s" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 191 | "exec (Ass x e) s (s[x := e])" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 192 | "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 193 | "s \<in> set b ==> exec c1 s t ==> exec (IF b c1 c2) s t" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 194 | "s \<notin> set b ==> exec c2 s t ==> exec (IF b c1 c2) s t" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 195 | "s \<notin> set b ==> exec (While b c) s s" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 196 | "s1 \<in> set b ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 197 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 198 | code_pred [random_dseq] exec . | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 199 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 200 | values [random_dseq 1, 2, 3] 10 "{(c, s, s'). exec c s s'}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 201 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 202 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 203 | "exec c s s' ==> exec (Seq c c) s s'" | 
| 40924 | 204 | quickcheck[tester = predicate_compile_wo_ff, size=2, iterations=20, expect = counterexample] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 205 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 206 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 207 | subsection {* Lambda *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 208 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 209 | datatype type = | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 210 | Atom nat | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 211 | | Fun type type (infixr "\<Rightarrow>" 200) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 212 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 213 | datatype dB = | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 214 | Var nat | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 215 | | App dB dB (infixl "\<degree>" 200) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 216 | | Abs type dB | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 217 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 218 | primrec | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 219 |   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 220 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 221 | "[]\<langle>i\<rangle> = None" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 222 | | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 223 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 224 | inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 225 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 226 | "nth_el' (x # xs) 0 x" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 227 | | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 228 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 229 | inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 230 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 231 | Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 232 | | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 233 | | App [intro!]: "env \<turnstile> s : U \<Rightarrow> T \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 234 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 235 | primrec | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 236 | lift :: "[dB, nat] => dB" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 237 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 238 | "lift (Var i) k = (if i < k then Var i else Var (i + 1))" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 239 | | "lift (s \<degree> t) k = lift s k \<degree> lift t k" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 240 | | "lift (Abs T s) k = Abs T (lift s (k + 1))" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 241 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 242 | primrec | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 243 |   subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 244 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 245 | subst_Var: "(Var i)[s/k] = | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 246 | (if k < i then Var (i - 1) else if i = k then s else Var i)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 247 | | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 248 | | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 249 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 250 | inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 251 | where | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 252 | beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 253 | | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 254 | | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 255 | | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 256 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 257 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 258 | "\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U" | 
| 40924 | 259 | quickcheck[tester = predicate_compile_ff_fs, size = 7, iterations = 10] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 260 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 261 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 262 | subsection {* JAD *}
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 263 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 264 | definition matrix :: "('a :: semiring_0) list list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 265 | "matrix M rs cs \<longleftrightarrow> (\<forall> row \<in> set M. length row = cs) \<and> length M = rs" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 266 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 267 | code_pred [random_dseq inductify] matrix . | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 268 | thm matrix.random_dseq_equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 269 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 270 | thm matrix_aux.random_dseq_equation | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 271 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 272 | values [random_dseq 3, 2] 10 "{(M, rs, cs). matrix (M:: int list list) rs cs}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 273 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 274 | lemma [code_pred_intro]: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 275 | "matrix [] 0 m" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 276 | "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 277 | proof - | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 278 | show "matrix [] 0 m" unfolding matrix_def by auto | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 279 | next | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 280 | show "matrix xss n m ==> length xs = m ==> matrix (xs # xss) (Suc n) m" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 281 | unfolding matrix_def by auto | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 282 | qed | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 283 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 284 | code_pred [random_dseq inductify] matrix | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 285 | apply (cases x) | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42463diff
changeset | 286 | unfolding matrix_def apply fastforce | 
| 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
42463diff
changeset | 287 | apply fastforce done | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 288 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 289 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 290 | values [random_dseq 2, 2, 15] 6 "{(M::int list list, n, m). matrix M n m}"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 291 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 292 | definition "scalar_product v w = (\<Sum> (x, y)\<leftarrow>zip v w. x * y)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 293 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 294 | definition mv :: "('a \<Colon> semiring_0) list list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 295 | where [simp]: "mv M v = map (scalar_product v) M" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 296 | text {*
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 297 |   This defines the matrix vector multiplication. To work properly @{term
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 298 | "matrix M m n \<and> length v = n"} must hold. | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 299 | *} | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 300 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 301 | subsection "Compressed matrix" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 302 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 303 | definition "sparsify xs = [i \<leftarrow> zip [0..<length xs] xs. snd i \<noteq> 0]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 304 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 305 | lemma sparsify_length: "(i, x) \<in> set (sparsify xs) \<Longrightarrow> i < length xs" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 306 | by (auto simp: sparsify_def set_zip) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 307 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 308 | lemma listsum_sparsify[simp]: | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 309 |   fixes v :: "('a \<Colon> semiring_0) list"
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 310 | assumes "length w = length v" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 311 | shows "(\<Sum>x\<leftarrow>sparsify w. (\<lambda>(i, x). v ! i) x * snd x) = scalar_product v w" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 312 | (is "(\<Sum>x\<leftarrow>_. ?f x) = _") | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 313 | unfolding sparsify_def scalar_product_def | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 314 | using assms listsum_map_filter[where f="?f" and P="\<lambda> i. snd i \<noteq> (0::'a)"] | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 315 | by (simp add: listsum_setsum) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 316 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 317 | definition [simp]: "unzip w = (map fst w, map snd w)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 318 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 319 | primrec insert :: "('a \<Rightarrow> 'b \<Colon> linorder) => 'a \<Rightarrow> 'a list => 'a list" where
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 320 | "insert f x [] = [x]" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 321 | "insert f x (y # ys) = (if f y < f x then y # insert f x ys else x # y # ys)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 322 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 323 | primrec sort :: "('a \<Rightarrow> 'b \<Colon> linorder) \<Rightarrow> 'a list => 'a list" where
 | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 324 | "sort f [] = []" | | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 325 | "sort f (x # xs) = insert f x (sort f xs)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 326 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 327 | definition | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 328 | "length_permutate M = (unzip o sort (length o snd)) (zip [0 ..< length M] M)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 329 | (* | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 330 | definition | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 331 | "transpose M = [map (\<lambda> xs. xs ! i) (takeWhile (\<lambda> xs. i < length xs) M). i \<leftarrow> [0 ..< length (M ! 0)]]" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 332 | *) | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 333 | definition | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 334 | "inflate upds = foldr (\<lambda> (i, x) upds. upds[i := x]) upds (replicate (length upds) 0)" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 335 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 336 | definition | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 337 | "jad = apsnd transpose o length_permutate o map sparsify" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 338 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 339 | definition | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 340 | "jad_mv v = inflate o split zip o apsnd (map listsum o transpose o map (map (\<lambda> (i, x). v ! i * x)))" | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 341 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 342 | lemma "matrix (M::int list list) rs cs \<Longrightarrow> False" | 
| 40924 | 343 | quickcheck[tester = predicate_compile_ff_nofs, size = 6] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 344 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 345 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 346 | lemma | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 347 | "\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v" | 
| 40924 | 348 | quickcheck[tester = predicate_compile_wo_ff] | 
| 35955 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 349 | oops | 
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 350 | |
| 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 bulwahn parents: diff
changeset | 351 | end |