| author | blanchet | 
| Thu, 15 Dec 2016 15:05:35 +0100 | |
| changeset 64561 | a7664ca9ffc5 | 
| parent 63167 | 0909deb8059b | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
1  | 
theory Specialisation_Examples  | 
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40054 
diff
changeset
 | 
2  | 
imports Main "~~/src/HOL/Library/Predicate_Compile_Alternative_Defs"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
3  | 
begin  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
4  | 
|
| 
45125
 
c15b0faeb70a
increasing values_timeout to avoid SML_makeall failures on our current tests
 
bulwahn 
parents: 
42463 
diff
changeset
 | 
5  | 
declare [[values_timeout = 960.0]]  | 
| 
42142
 
d24a93025feb
raised various timeouts to accommodate sluggish SML/NJ
 
krauss 
parents: 
41413 
diff
changeset
 | 
6  | 
|
| 63167 | 7  | 
section \<open>Specialisation Examples\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
8  | 
|
| 
39919
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
9  | 
primrec nth_el'  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
10  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
11  | 
"nth_el' [] i = None"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
12  | 
| "nth_el' (x # xs) i = (case i of 0 => Some x | Suc j => nth_el' xs j)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
13  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
14  | 
definition  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
15  | 
"greater_than_index xs = (\<forall>i x. nth_el' xs i = Some x --> x > i)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
16  | 
|
| 
36257
 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 
bulwahn 
parents: 
36054 
diff
changeset
 | 
17  | 
code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index .  | 
| 63167 | 18  | 
ML_val \<open>Core_Data.intros_of @{context} @{const_name specialised_nth_el'P}\<close>
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
19  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
20  | 
thm greater_than_index.equation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
22  | 
values [expected "{()}"] "{x. greater_than_index [1,2,4,6]}"
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
23  | 
values [expected "{}"] "{x. greater_than_index [0,2,3,2]}"
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
24  | 
|
| 63167 | 25  | 
subsection \<open>Common subterms\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
26  | 
|
| 63167 | 27  | 
text \<open>If a predicate is called with common subterms as arguments,  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
28  | 
this predicate should be specialised.  | 
| 63167 | 29  | 
\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
31  | 
definition max_nat :: "nat => nat => nat"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
32  | 
where "max_nat a b = (if a <= b then b else a)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
34  | 
lemma [code_pred_inline]:  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
35  | 
"max = max_nat"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39198 
diff
changeset
 | 
36  | 
by (simp add: fun_eq_iff max_def max_nat_def)  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
38  | 
definition  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
39  | 
"max_of_my_Suc x = max x (Suc x)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
40  | 
|
| 63167 | 41  | 
text \<open>In this example, max is specialised, hence the mode o => i => bool is possible\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
42  | 
|
| 
36257
 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 
bulwahn 
parents: 
36054 
diff
changeset
 | 
43  | 
code_pred (modes: o => i => bool) [inductify, specialise, skip_proof] max_of_my_Suc .  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
44  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
45  | 
thm max_of_my_SucP.equation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
46  | 
|
| 63167 | 47  | 
ML_val \<open>Core_Data.intros_of @{context} @{const_name specialised_max_natP}\<close>
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
49  | 
values "{x. max_of_my_SucP x 6}"
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
50  | 
|
| 63167 | 51  | 
subsection \<open>Sorts\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
52  | 
|
| 
39919
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
53  | 
declare sorted.Nil [code_pred_intro]  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
54  | 
sorted_single [code_pred_intro]  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
55  | 
sorted_many [code_pred_intro]  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
56  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
57  | 
code_pred sorted  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
58  | 
proof -  | 
| 
39919
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
59  | 
assume "sorted xa"  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
60  | 
assume 1: "xa = [] \<Longrightarrow> thesis"  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
61  | 
assume 2: "\<And>x. xa = [x] \<Longrightarrow> thesis"  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
62  | 
assume 3: "\<And>x y zs. xa = x # y # zs \<Longrightarrow> x \<le> y \<Longrightarrow> sorted (y # zs) \<Longrightarrow> thesis"  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
63  | 
show thesis proof (cases xa)  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
64  | 
case Nil with 1 show ?thesis .  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
65  | 
next  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
66  | 
case (Cons x xs) show ?thesis proof (cases xs)  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
67  | 
case Nil with Cons 2 show ?thesis by simp  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
68  | 
next  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
69  | 
case (Cons y zs)  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
70  | 
show ?thesis  | 
| 
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
71  | 
proof (rule 3)  | 
| 63167 | 72  | 
from Cons \<open>xa = x # xs\<close> show "xa = x # y # zs" by simp  | 
73  | 
with \<open>sorted xa\<close> show "x \<le> y" and "sorted (y # zs)" by simp_all  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53072 
diff
changeset
 | 
74  | 
qed  | 
| 
39919
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
75  | 
qed  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
76  | 
qed  | 
| 
 
9f6503aaa77d
adjusted to inductive characterization of sorted
 
haftmann 
parents: 
39302 
diff
changeset
 | 
77  | 
qed  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
78  | 
thm sorted.equation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
79  | 
|
| 63167 | 80  | 
section \<open>Specialisation in POPLmark theory\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
81  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
82  | 
notation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
83  | 
  Some ("\<lfloor>_\<rfloor>")
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
84  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
85  | 
notation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
86  | 
  None ("\<bottom>")
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
87  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
88  | 
notation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
89  | 
  length ("\<parallel>_\<parallel>")
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
91  | 
notation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
92  | 
  Cons ("_ \<Colon>/ _" [66, 65] 65)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
94  | 
primrec  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
95  | 
  nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
96  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
97  | 
"[]\<langle>i\<rangle> = \<bottom>"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
98  | 
| "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> \<lfloor>x\<rfloor> | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
99  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
100  | 
primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" ("_\<langle>_\<rangle>\<^sub>?" [90, 0] 91)
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
101  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
102  | 
"[]\<langle>a\<rangle>\<^sub>? = \<bottom>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
103  | 
| "(x # xs)\<langle>a\<rangle>\<^sub>? = (if fst x = a then \<lfloor>snd x\<rfloor> else xs\<langle>a\<rangle>\<^sub>?)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
105  | 
primrec unique :: "('a \<times> 'b) list \<Rightarrow> bool"
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
106  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
107  | 
"unique [] = True"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
108  | 
| "unique (x # xs) = (xs\<langle>fst x\<rangle>\<^sub>? = \<bottom> \<and> unique xs)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
109  | 
|
| 58310 | 110  | 
datatype type =  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
111  | 
TVar nat  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
112  | 
| Top  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
113  | 
| Fun type type (infixr "\<rightarrow>" 200)  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
114  | 
  | TyAll type type  ("(3\<forall><:_./ _)" [0, 10] 10)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
115  | 
|
| 58310 | 116  | 
datatype binding = VarB type | TVarB type  | 
| 42463 | 117  | 
type_synonym env = "binding list"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
119  | 
primrec is_TVarB :: "binding \<Rightarrow> bool"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
120  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
121  | 
"is_TVarB (VarB T) = False"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
122  | 
| "is_TVarB (TVarB T) = True"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
123  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
124  | 
primrec type_ofB :: "binding \<Rightarrow> type"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
125  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
126  | 
"type_ofB (VarB T) = T"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
127  | 
| "type_ofB (TVarB T) = T"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
128  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
129  | 
primrec mapB :: "(type \<Rightarrow> type) \<Rightarrow> binding \<Rightarrow> binding"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
130  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
131  | 
"mapB f (VarB T) = VarB (f T)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
132  | 
| "mapB f (TVarB T) = TVarB (f T)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
133  | 
|
| 58310 | 134  | 
datatype trm =  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
135  | 
Var nat  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
136  | 
  | Abs type trm   ("(3\<lambda>:_./ _)" [0, 10] 10)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
137  | 
  | TAbs type trm  ("(3\<lambda><:_./ _)" [0, 10] 10)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
138  | 
| App trm trm (infixl "\<bullet>" 200)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
139  | 
| TApp trm type (infixl "\<bullet>\<^sub>\<tau>" 200)  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
140  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
141  | 
primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<up>\<^sub>\<tau>")
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
142  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
143  | 
"\<up>\<^sub>\<tau> n k (TVar i) = (if i < k then TVar i else TVar (i + n))"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
144  | 
| "\<up>\<^sub>\<tau> n k Top = Top"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
145  | 
| "\<up>\<^sub>\<tau> n k (T \<rightarrow> U) = \<up>\<^sub>\<tau> n k T \<rightarrow> \<up>\<^sub>\<tau> n k U"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
146  | 
| "\<up>\<^sub>\<tau> n k (\<forall><:T. U) = (\<forall><:\<up>\<^sub>\<tau> n k T. \<up>\<^sub>\<tau> n (k + 1) U)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
147  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
148  | 
primrec lift :: "nat \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("\<up>")
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
149  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
150  | 
"\<up> n k (Var i) = (if i < k then Var i else Var (i + n))"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
151  | 
| "\<up> n k (\<lambda>:T. t) = (\<lambda>:\<up>\<^sub>\<tau> n k T. \<up> n (k + 1) t)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
152  | 
| "\<up> n k (\<lambda><:T. t) = (\<lambda><:\<up>\<^sub>\<tau> n k T. \<up> n (k + 1) t)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
153  | 
| "\<up> n k (s \<bullet> t) = \<up> n k s \<bullet> \<up> n k t"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
154  | 
| "\<up> n k (t \<bullet>\<^sub>\<tau> T) = \<up> n k t \<bullet>\<^sub>\<tau> \<up>\<^sub>\<tau> n k T"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
155  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
156  | 
primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type"  ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>\<tau>" [300, 0, 0] 300)
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
157  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
158  | 
"(TVar i)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
159  | 
(if k < i then TVar (i - 1) else if i = k then \<up>\<^sub>\<tau> k 0 S else TVar i)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
160  | 
| "Top[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = Top"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
161  | 
| "(T \<rightarrow> U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> \<rightarrow> U[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
162  | 
| "(\<forall><:T. U)[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau> = (\<forall><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. U[k+1 \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
163  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
164  | 
primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type"  ("\<down>\<^sub>\<tau>")
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
165  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
166  | 
"\<down>\<^sub>\<tau> 0 k T = T"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
167  | 
| "\<down>\<^sub>\<tau> (Suc n) k T = \<down>\<^sub>\<tau> n k (T[k \<mapsto>\<^sub>\<tau> Top]\<^sub>\<tau>)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
168  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
169  | 
primrec subst :: "trm \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm"  ("_[_ \<mapsto> _]" [300, 0, 0] 300)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
170  | 
where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
171  | 
"(Var i)[k \<mapsto> s] = (if k < i then Var (i - 1) else if i = k then \<up> k 0 s else Var i)"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
172  | 
| "(t \<bullet> u)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet> u[k \<mapsto> s]"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
173  | 
| "(t \<bullet>\<^sub>\<tau> T)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet>\<^sub>\<tau> \<down>\<^sub>\<tau> 1 k T"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
174  | 
| "(\<lambda>:T. t)[k \<mapsto> s] = (\<lambda>:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
175  | 
| "(\<lambda><:T. t)[k \<mapsto> s] = (\<lambda><:\<down>\<^sub>\<tau> 1 k T. t[k+1 \<mapsto> s])"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
176  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
177  | 
primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm"    ("_[_ \<mapsto>\<^sub>\<tau> _]" [300, 0, 0] 300)
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
178  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
179  | 
"(Var i)[k \<mapsto>\<^sub>\<tau> S] = (if k < i then Var (i - 1) else Var i)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
180  | 
| "(t \<bullet> u)[k \<mapsto>\<^sub>\<tau> S] = t[k \<mapsto>\<^sub>\<tau> S] \<bullet> u[k \<mapsto>\<^sub>\<tau> S]"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
181  | 
| "(t \<bullet>\<^sub>\<tau> T)[k \<mapsto>\<^sub>\<tau> S] = t[k \<mapsto>\<^sub>\<tau> S] \<bullet>\<^sub>\<tau> T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
182  | 
| "(\<lambda>:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda>:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
183  | 
| "(\<lambda><:T. t)[k \<mapsto>\<^sub>\<tau> S] = (\<lambda><:T[k \<mapsto>\<^sub>\<tau> S]\<^sub>\<tau>. t[k+1 \<mapsto>\<^sub>\<tau> S])"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
184  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
185  | 
primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<up>\<^sub>e")
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
186  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
187  | 
"\<up>\<^sub>e n k [] = []"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
188  | 
| "\<up>\<^sub>e n k (B \<Colon> \<Gamma>) = mapB (\<up>\<^sub>\<tau> n (k + \<parallel>\<Gamma>\<parallel>)) B \<Colon> \<up>\<^sub>e n k \<Gamma>"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
189  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
190  | 
primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env"  ("_[_ \<mapsto>\<^sub>\<tau> _]\<^sub>e" [300, 0, 0] 300)
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
191  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
192  | 
"[][k \<mapsto>\<^sub>\<tau> T]\<^sub>e = []"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
193  | 
| "(B \<Colon> \<Gamma>)[k \<mapsto>\<^sub>\<tau> T]\<^sub>e = mapB (\<lambda>U. U[k + \<parallel>\<Gamma>\<parallel> \<mapsto>\<^sub>\<tau> T]\<^sub>\<tau>) B \<Colon> \<Gamma>[k \<mapsto>\<^sub>\<tau> T]\<^sub>e"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
194  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
195  | 
primrec decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env"  ("\<down>\<^sub>e")
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
196  | 
where  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
197  | 
"\<down>\<^sub>e 0 k \<Gamma> = \<Gamma>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
198  | 
| "\<down>\<^sub>e (Suc n) k \<Gamma> = \<down>\<^sub>e n k (\<Gamma>[k \<mapsto>\<^sub>\<tau> Top]\<^sub>e)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
199  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
200  | 
inductive  | 
| 53072 | 201  | 
  well_formed :: "env \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile>\<^sub>w\<^sub>f _" [50, 50] 50)
 | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
202  | 
where  | 
| 53072 | 203  | 
wf_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB T\<rfloor> \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TVar i"  | 
204  | 
| wf_Top: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f Top"  | 
|
205  | 
| wf_arrow: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f T \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f U \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f T \<rightarrow> U"  | 
|
206  | 
| wf_all: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f T \<Longrightarrow> TVarB T \<Colon> \<Gamma> \<turnstile>\<^sub>w\<^sub>f U \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f (\<forall><:T. U)"  | 
|
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
207  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
208  | 
inductive  | 
| 53072 | 209  | 
  well_formedE :: "env \<Rightarrow> bool"  ("_ \<turnstile>\<^sub>w\<^sub>f" [50] 50)
 | 
210  | 
  and well_formedB :: "env \<Rightarrow> binding \<Rightarrow> bool"  ("_ \<turnstile>\<^sub>w\<^sub>f\<^sub>B _" [50, 50] 50)
 | 
|
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
211  | 
where  | 
| 53072 | 212  | 
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f\<^sub>B B \<equiv> \<Gamma> \<turnstile>\<^sub>w\<^sub>f type_ofB B"  | 
213  | 
| wf_Nil: "[] \<turnstile>\<^sub>w\<^sub>f"  | 
|
214  | 
| wf_Cons: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f\<^sub>B B \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> B \<Colon> \<Gamma> \<turnstile>\<^sub>w\<^sub>f"  | 
|
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
216  | 
inductive_cases well_formed_cases:  | 
| 53072 | 217  | 
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f TVar i"  | 
218  | 
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f Top"  | 
|
219  | 
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f T \<rightarrow> U"  | 
|
220  | 
"\<Gamma> \<turnstile>\<^sub>w\<^sub>f (\<forall><:T. U)"  | 
|
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
221  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
222  | 
inductive_cases well_formedE_cases:  | 
| 53072 | 223  | 
"B \<Colon> \<Gamma> \<turnstile>\<^sub>w\<^sub>f"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
224  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
225  | 
inductive  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
226  | 
  subtyping :: "env \<Rightarrow> type \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ <: _" [50, 50, 50] 50)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
227  | 
where  | 
| 53072 | 228  | 
SA_Top: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f S \<Longrightarrow> \<Gamma> \<turnstile> S <: Top"  | 
229  | 
| SA_refl_TVar: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma> \<turnstile>\<^sub>w\<^sub>f TVar i \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: TVar i"  | 
|
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
230  | 
| SA_trans_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB U\<rfloor> \<Longrightarrow>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
231  | 
\<Gamma> \<turnstile> \<up>\<^sub>\<tau> (Suc i) 0 U <: T \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: T"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
232  | 
| SA_arrow: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<Longrightarrow> \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> S\<^sub>1 \<rightarrow> S\<^sub>2 <: T\<^sub>1 \<rightarrow> T\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
233  | 
| SA_all: "\<Gamma> \<turnstile> T\<^sub>1 <: S\<^sub>1 \<Longrightarrow> TVarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> S\<^sub>2 <: T\<^sub>2 \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
234  | 
\<Gamma> \<turnstile> (\<forall><:S\<^sub>1. S\<^sub>2) <: (\<forall><:T\<^sub>1. T\<^sub>2)"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
236  | 
inductive  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
237  | 
  typing :: "env \<Rightarrow> trm \<Rightarrow> type \<Rightarrow> bool"    ("_ \<turnstile> _ : _" [50, 50, 50] 50)
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
238  | 
where  | 
| 53072 | 239  | 
T_Var: "\<Gamma> \<turnstile>\<^sub>w\<^sub>f \<Longrightarrow> \<Gamma>\<langle>i\<rangle> = \<lfloor>VarB U\<rfloor> \<Longrightarrow> T = \<up>\<^sub>\<tau> (Suc i) 0 U \<Longrightarrow> \<Gamma> \<turnstile> Var i : T"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
240  | 
| T_Abs: "VarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>:T\<^sub>1. t\<^sub>2) : T\<^sub>1 \<rightarrow> \<down>\<^sub>\<tau> 1 0 T\<^sub>2"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
241  | 
| T_App: "\<Gamma> \<turnstile> t\<^sub>1 : T\<^sub>11 \<rightarrow> T\<^sub>12 \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>11 \<Longrightarrow> \<Gamma> \<turnstile> t\<^sub>1 \<bullet> t\<^sub>2 : T\<^sub>12"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
242  | 
| T_TAbs: "TVarB T\<^sub>1 \<Colon> \<Gamma> \<turnstile> t\<^sub>2 : T\<^sub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda><:T\<^sub>1. t\<^sub>2) : (\<forall><:T\<^sub>1. T\<^sub>2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
243  | 
| T_TApp: "\<Gamma> \<turnstile> t\<^sub>1 : (\<forall><:T\<^sub>11. T\<^sub>12) \<Longrightarrow> \<Gamma> \<turnstile> T\<^sub>2 <: T\<^sub>11 \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
51272 
diff
changeset
 | 
244  | 
\<Gamma> \<turnstile> t\<^sub>1 \<bullet>\<^sub>\<tau> T\<^sub>2 : T\<^sub>12[0 \<mapsto>\<^sub>\<tau> T\<^sub>2]\<^sub>\<tau>"  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
245  | 
| T_Sub: "\<Gamma> \<turnstile> t : S \<Longrightarrow> \<Gamma> \<turnstile> S <: T \<Longrightarrow> \<Gamma> \<turnstile> t : T"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
246  | 
|
| 
36257
 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 
bulwahn 
parents: 
36054 
diff
changeset
 | 
247  | 
code_pred [inductify, skip_proof, specialise] typing .  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
248  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
249  | 
thm typing.equation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
250  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
251  | 
values 6 "{(E, t, T). typing E t T}"
 | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
252  | 
|
| 63167 | 253  | 
subsection \<open>Higher-order predicate\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
254  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
255  | 
code_pred [inductify] mapB .  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
256  | 
|
| 63167 | 257  | 
subsection \<open>Multiple instances\<close>  | 
| 
36033
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
258  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
259  | 
inductive subtype_refl' where  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
260  | 
"\<Gamma> \<turnstile> t : T ==> \<not> (\<Gamma> \<turnstile> T <: T) ==> subtype_refl' t T"  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
261  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
262  | 
code_pred (modes: i => i => bool, o => i => bool, i => o => bool, o => o => bool) [inductify] subtype_refl' .  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
263  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
264  | 
thm subtype_refl'.equation  | 
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
7106f079bd05
adding specialisation examples of the predicate compiler
 
bulwahn 
parents:  
diff
changeset
 | 
266  | 
|
| 
45125
 
c15b0faeb70a
increasing values_timeout to avoid SML_makeall failures on our current tests
 
bulwahn 
parents: 
42463 
diff
changeset
 | 
267  | 
end  |