| author | wenzelm | 
| Wed, 04 Oct 2017 20:16:53 +0200 | |
| changeset 66764 | 006deaf5c3dc | 
| parent 66453 | cc19f7ca2ed6 | 
| child 68249 | 949d93804740 | 
| 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  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
64267 
diff
changeset
 | 
2  | 
imports "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  | 
|
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
5  | 
(*  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
6  | 
section {* Sets *}
 | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
lemma "x \<in> {(1::nat)} ==> False"
 | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
9  | 
quickcheck[generator=predicate_compile_wo_ff, iterations=10]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
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
 | 
13  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
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
 | 
17  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
18  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
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
 | 
21  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
section {* Numerals *}
 | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
27  | 
  "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
 | 
28  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
29  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
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
 | 
32  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
36  | 
  "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
 | 
37  | 
quickcheck[generator=predicate_compile_wo_ff]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
oops  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
39  | 
*)  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
|
| 63167 | 41  | 
section \<open>Equivalences\<close>  | 
| 
39650
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
42  | 
|
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
43  | 
inductive is_ten :: "nat => bool"  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
44  | 
where  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
45  | 
"is_ten 10"  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
46  | 
|
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
47  | 
inductive is_eleven :: "nat => bool"  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
48  | 
where  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
49  | 
"is_eleven 11"  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
50  | 
|
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
51  | 
lemma  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
52  | 
"is_ten x = is_eleven x"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
53  | 
quickcheck[tester = smart_exhaustive, iterations = 1, size = 1, expect = counterexample]  | 
| 
39650
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
54  | 
oops  | 
| 
 
2a35950ec495
handling equivalences smarter in the predicate compiler
 
bulwahn 
parents: 
39191 
diff
changeset
 | 
55  | 
|
| 63167 | 56  | 
section \<open>Context Free Grammar\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
57  | 
|
| 58310 | 58  | 
datatype alphabet = a | b  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
59  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
60  | 
inductive_set S\<^sub>1 and A\<^sub>1 and B\<^sub>1 where  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
61  | 
"[] \<in> S\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
62  | 
| "w \<in> A\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
63  | 
| "w \<in> B\<^sub>1 \<Longrightarrow> a # w \<in> S\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
64  | 
| "w \<in> S\<^sub>1 \<Longrightarrow> a # w \<in> A\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
65  | 
| "w \<in> S\<^sub>1 \<Longrightarrow> b # w \<in> S\<^sub>1"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
66  | 
| "\<lbrakk>v \<in> B\<^sub>1; v \<in> B\<^sub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>1"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
67  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
68  | 
lemma  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
69  | 
"S\<^sub>1p w \<Longrightarrow> w = []"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
70  | 
quickcheck[tester = smart_exhaustive, iterations=1]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
71  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
72  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
73  | 
theorem S\<^sub>1_sound:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
74  | 
"S\<^sub>1p w \<Longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
75  | 
quickcheck[tester=smart_exhaustive, size=15]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
76  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
77  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
79  | 
inductive_set S\<^sub>2 and A\<^sub>2 and B\<^sub>2 where  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
80  | 
"[] \<in> S\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
81  | 
| "w \<in> A\<^sub>2 \<Longrightarrow> b # w \<in> S\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
82  | 
| "w \<in> B\<^sub>2 \<Longrightarrow> a # w \<in> S\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
83  | 
| "w \<in> S\<^sub>2 \<Longrightarrow> a # w \<in> A\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
84  | 
| "w \<in> S\<^sub>2 \<Longrightarrow> b # w \<in> B\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
85  | 
| "\<lbrakk>v \<in> B\<^sub>2; v \<in> B\<^sub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>2"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
(*  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
87  | 
code_pred [random_dseq inductify] S\<^sub>2 .  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
88  | 
thm S\<^sub>2.random_dseq_equation  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
89  | 
thm A\<^sub>2.random_dseq_equation  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
90  | 
thm B\<^sub>2.random_dseq_equation  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
91  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
92  | 
values [random_dseq 1, 2, 8] 10 "{x. S\<^sub>2 x}"
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
94  | 
lemma "w \<in> S\<^sub>2 ==> w \<noteq> [] ==> w \<noteq> [b, a] ==> w \<in> {}"
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
quickcheck[generator=predicate_compile, size=8]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
lemma "[x <- w. x = a] = []"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
quickcheck[generator=predicate_compile]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
100  | 
oops  | 
| 
 
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  | 
declare list.size(3,4)[code_pred_def]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
(*  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
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
 | 
106  | 
quickcheck[generator=predicate_compile]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
107  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
108  | 
*)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
110  | 
lemma  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
111  | 
"w \<in> S\<^sub>2 ==> length [x \<leftarrow> w. x = a] <= Suc (Suc 0)"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
112  | 
quickcheck[generator=predicate_compile, size = 10, iterations = 1]  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
*)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
115  | 
theorem S\<^sub>2_sound:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
116  | 
"S\<^sub>2p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
117  | 
quickcheck[tester=smart_exhaustive, size=5, iterations=10]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
119  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
120  | 
inductive_set S\<^sub>3 and A\<^sub>3 and B\<^sub>3 where  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
121  | 
"[] \<in> S\<^sub>3"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
122  | 
| "w \<in> A\<^sub>3 \<Longrightarrow> b # w \<in> S\<^sub>3"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
123  | 
| "w \<in> B\<^sub>3 \<Longrightarrow> a # w \<in> S\<^sub>3"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
124  | 
| "w \<in> S\<^sub>3 \<Longrightarrow> a # w \<in> A\<^sub>3"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
125  | 
| "w \<in> S\<^sub>3 \<Longrightarrow> b # w \<in> B\<^sub>3"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
126  | 
| "\<lbrakk>v \<in> B\<^sub>3; w \<in> B\<^sub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>3"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
127  | 
|
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
128  | 
code_pred [inductify, skip_proof] S\<^sub>3p .  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
129  | 
thm S\<^sub>3p.equation  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
130  | 
(*  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
131  | 
values 10 "{x. S\<^sub>3 x}"
 | 
| 
35955
 
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  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
134  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
135  | 
lemma S\<^sub>3_sound:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
136  | 
"S\<^sub>3p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
137  | 
quickcheck[tester=smart_exhaustive, size=10, iterations=10]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
139  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
lemma "\<not> (length w > 2) \<or> \<not> (length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b])"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
141  | 
quickcheck[size=10, tester = smart_exhaustive]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
142  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
143  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
144  | 
theorem S\<^sub>3_complete:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
145  | 
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. b = x] \<longrightarrow> S\<^sub>3p w"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
146  | 
(*quickcheck[generator=SML]*)  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
147  | 
quickcheck[tester=smart_exhaustive, size=10, iterations=100]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
148  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
149  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
150  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
151  | 
inductive_set S\<^sub>4 and A\<^sub>4 and B\<^sub>4 where  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
152  | 
"[] \<in> S\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
153  | 
| "w \<in> A\<^sub>4 \<Longrightarrow> b # w \<in> S\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
154  | 
| "w \<in> B\<^sub>4 \<Longrightarrow> a # w \<in> S\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
155  | 
| "w \<in> S\<^sub>4 \<Longrightarrow> a # w \<in> A\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
156  | 
| "\<lbrakk>v \<in> A\<^sub>4; w \<in> A\<^sub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
157  | 
| "w \<in> S\<^sub>4 \<Longrightarrow> b # w \<in> B\<^sub>4"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
158  | 
| "\<lbrakk>v \<in> B\<^sub>4; w \<in> B\<^sub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^sub>4"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
159  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
160  | 
theorem S\<^sub>4_sound:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
161  | 
"S\<^sub>4p w \<longrightarrow> length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b]"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
162  | 
quickcheck[tester = smart_exhaustive, size=5, iterations=1]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
163  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
164  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
44890 
diff
changeset
 | 
165  | 
theorem S\<^sub>4_complete:  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
166  | 
"length [x \<leftarrow> w. x = a] = length [x \<leftarrow> w. x = b] \<longrightarrow> S\<^sub>4p w"  | 
| 
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
167  | 
quickcheck[tester = smart_exhaustive, size=5, iterations=1]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
168  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
169  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
36040 
diff
changeset
 | 
170  | 
hide_const a b  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
171  | 
|
| 63167 | 172  | 
subsection \<open>Lexicographic order\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
173  | 
(* TODO *)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
174  | 
(*  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
175  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
176  | 
"(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
 | 
177  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
178  | 
*)  | 
| 63167 | 179  | 
subsection \<open>IMP\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
180  | 
|
| 42463 | 181  | 
type_synonym var = nat  | 
182  | 
type_synonym state = "int list"  | 
|
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
183  | 
|
| 58310 | 184  | 
datatype com =  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
185  | 
Skip |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
186  | 
Ass var "int" |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
187  | 
Seq com com |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
188  | 
IF "state list" com com |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
189  | 
While "state list" com  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
190  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
191  | 
inductive exec :: "com => state => state => bool" where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
192  | 
"exec Skip s s" |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
193  | 
"exec (Ass x e) s (s[x := e])" |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
194  | 
"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
 | 
195  | 
"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
 | 
196  | 
"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
 | 
197  | 
"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
 | 
198  | 
"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
 | 
199  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
200  | 
code_pred [random_dseq] exec .  | 
| 
 
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  | 
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
 | 
203  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
204  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
205  | 
"exec c s s' ==> exec (Seq c c) s s'"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
206  | 
quickcheck[tester = smart_exhaustive, size=2, iterations=20, expect = counterexample]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
207  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
208  | 
|
| 63167 | 209  | 
subsection \<open>Lambda\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
210  | 
|
| 58310 | 211  | 
datatype type =  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
212  | 
Atom nat  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
213  | 
| Fun type type (infixr "\<Rightarrow>" 200)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
214  | 
|
| 58310 | 215  | 
datatype dB =  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
216  | 
Var nat  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
217  | 
| App dB dB (infixl "\<degree>" 200)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
218  | 
| Abs type dB  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
219  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
220  | 
primrec  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
221  | 
  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
 | 
222  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
223  | 
"[]\<langle>i\<rangle> = None"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
224  | 
| "(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
 | 
225  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
226  | 
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
 | 
227  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
228  | 
"nth_el' (x # xs) 0 x"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
229  | 
| "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
 | 
230  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
231  | 
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
 | 
232  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
233  | 
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
 | 
234  | 
| 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
 | 
235  | 
| 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
 | 
236  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
237  | 
primrec  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
238  | 
lift :: "[dB, nat] => dB"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
239  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
240  | 
"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
 | 
241  | 
| "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
 | 
242  | 
| "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
 | 
243  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
244  | 
primrec  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
245  | 
  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
 | 
246  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
247  | 
subst_Var: "(Var i)[s/k] =  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
248  | 
(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
 | 
249  | 
| 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
 | 
250  | 
| 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
 | 
251  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
252  | 
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
 | 
253  | 
where  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
254  | 
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
 | 
255  | 
| 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
 | 
256  | 
| 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
 | 
257  | 
| 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
 | 
258  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
259  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
260  | 
"\<Gamma> \<turnstile> t : U \<Longrightarrow> t \<rightarrow>\<^sub>\<beta> t' \<Longrightarrow> \<Gamma> \<turnstile> t' : U"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
261  | 
quickcheck[tester = smart_exhaustive, size = 7, iterations = 10]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
262  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
263  | 
|
| 63167 | 264  | 
subsection \<open>JAD\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
266  | 
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
 | 
267  | 
"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
 | 
268  | 
(*  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
269  | 
code_pred [random_dseq inductify] matrix .  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
270  | 
thm matrix.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  | 
thm matrix_aux.random_dseq_equation  | 
| 
 
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  | 
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
 | 
275  | 
*)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
276  | 
lemma [code_pred_intro]:  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
277  | 
"matrix [] 0 m"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
278  | 
"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
 | 
279  | 
proof -  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
280  | 
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
 | 
281  | 
next  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
282  | 
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
 | 
283  | 
unfolding matrix_def by auto  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
284  | 
qed  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
285  | 
|
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
286  | 
code_pred [random_dseq] matrix  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
287  | 
apply (cases x)  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
288  | 
unfolding matrix_def apply fastforce  | 
| 
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
42463 
diff
changeset
 | 
289  | 
apply fastforce done  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
290  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
291  | 
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
 | 
292  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
293  | 
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
 | 
294  | 
|
| 61142 | 295  | 
definition mv :: "('a :: semiring_0) list list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
296  | 
where [simp]: "mv M v = map (scalar_product v) M"  | 
| 63167 | 297  | 
text \<open>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
298  | 
  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
 | 
299  | 
"matrix M m n \<and> length v = n"} must hold.  | 
| 63167 | 300  | 
\<close>  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
301  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
302  | 
subsection "Compressed matrix"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
303  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
304  | 
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
 | 
305  | 
(*  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
306  | 
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
 | 
307  | 
by (auto simp: sparsify_def set_zip)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
308  | 
|
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63167 
diff
changeset
 | 
309  | 
lemma sum_list_sparsify[simp]:  | 
| 61142 | 310  | 
  fixes v :: "('a :: semiring_0) list"
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
311  | 
assumes "length w = length v"  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
312  | 
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
 | 
313  | 
(is "(\<Sum>x\<leftarrow>_. ?f x) = _")  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
314  | 
unfolding sparsify_def scalar_product_def  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63167 
diff
changeset
 | 
315  | 
using assms sum_list_map_filter[where f="?f" and P="\<lambda> i. snd i \<noteq> (0::'a)"]  | 
| 64267 | 316  | 
by (simp add: sum_list_sum)  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
317  | 
*)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
318  | 
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
 | 
319  | 
|
| 61142 | 320  | 
primrec insert :: "('a \<Rightarrow> 'b :: linorder) => 'a \<Rightarrow> 'a list => 'a list" where
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
321  | 
"insert f x [] = [x]" |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
322  | 
"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
 | 
323  | 
|
| 61142 | 324  | 
primrec sort :: "('a \<Rightarrow> 'b :: linorder) \<Rightarrow> 'a list => 'a list" where
 | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
325  | 
"sort f [] = []" |  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
326  | 
"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
 | 
327  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
328  | 
definition  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
329  | 
"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
 | 
330  | 
(*  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
331  | 
definition  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
332  | 
"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
 | 
333  | 
*)  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
334  | 
definition  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
335  | 
"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
 | 
336  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
337  | 
definition  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
338  | 
"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
 | 
339  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
340  | 
definition  | 
| 
63882
 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 
nipkow 
parents: 
63167 
diff
changeset
 | 
341  | 
"jad_mv v = inflate o case_prod zip o apsnd (map sum_list o transpose o map (map (\<lambda> (i, x). v ! i * x)))"  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
342  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
343  | 
lemma "matrix (M::int list list) rs cs \<Longrightarrow> False"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
344  | 
quickcheck[tester = smart_exhaustive, size = 6]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
345  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
346  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
347  | 
lemma  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
348  | 
"\<lbrakk> matrix M rs cs ; length v = cs \<rbrakk> \<Longrightarrow> jad_mv v (jad M) = mv M v"  | 
| 
61140
 
78ece168f5b5
reactivate examples with predicate compiler and quickcheck
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
349  | 
quickcheck[tester = smart_exhaustive]  | 
| 
35955
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
350  | 
oops  | 
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
351  | 
|
| 
 
e657fb805c68
added predicate compiler quickcheck examples to new session Predicate_Compile_Examples
 
bulwahn 
parents:  
diff
changeset
 | 
352  | 
end  |