| author | haftmann | 
| Mon, 28 Jun 2010 15:32:26 +0200 | |
| changeset 37606 | b47dd044a1f1 | 
| parent 37008 | 8da3b51726ac | 
| child 39198 | f967a16dfcdd | 
| permissions | -rw-r--r-- | 
| 36033 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 1 | theory Specialisation_Examples | 
| 36054 
93d62439506c
adopting specialisation examples to moving the alternative defs in the library
 bulwahn parents: 
36040diff
changeset | 2 | imports Main 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 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 5 | section {* Specialisation Examples *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 6 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 7 | fun nth_el' | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 8 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 9 | "nth_el' [] i = None" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 10 | | "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 | 11 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 12 | definition | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 13 | "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 | 14 | |
| 36257 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 bulwahn parents: 
36054diff
changeset | 15 | code_pred (expected_modes: i => bool) [inductify, skip_proof, specialise] greater_than_index . | 
| 37008 | 16 | ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_nth_el'P} *}
 | 
| 36033 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 17 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 18 | thm greater_than_index.equation | 
| 
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 | values [expected "{()}"] "{x. greater_than_index [1,2,4,6]}"
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 21 | values [expected "{}"] "{x. greater_than_index [0,2,3,2]}"
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 22 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 23 | subsection {* Common subterms *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 24 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 25 | text {* If a predicate is called with common subterms as arguments,
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 26 | this predicate should be specialised. | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 27 | *} | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 28 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 29 | definition max_nat :: "nat => nat => nat" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 30 | where "max_nat a b = (if a <= b then b else a)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 31 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 32 | lemma [code_pred_inline]: | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 33 | "max = max_nat" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 34 | by (simp add: expand_fun_eq max_def max_nat_def) | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 35 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 36 | definition | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 37 | "max_of_my_Suc x = max x (Suc x)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 38 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 39 | text {* In this example, max is specialised, hence the mode o => i => bool is possible *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 40 | |
| 36257 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 bulwahn parents: 
36054diff
changeset | 41 | 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 | 42 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 43 | thm max_of_my_SucP.equation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 44 | |
| 37008 | 45 | ML {* Predicate_Compile_Core.intros_of @{context} @{const_name specialised_max_natP} *}
 | 
| 36033 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 46 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 47 | values "{x. max_of_my_SucP x 6}"
 | 
| 
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 | subsection {* Sorts *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 50 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 51 | code_pred [inductify] sorted . | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 52 | thm sorted.equation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 53 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 54 | section {* Specialisation in POPLmark theory *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 55 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 56 | notation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 57 |   Some ("\<lfloor>_\<rfloor>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 58 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 59 | notation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 60 |   None ("\<bottom>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 61 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 62 | notation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 63 |   length ("\<parallel>_\<parallel>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 64 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 65 | notation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 66 |   Cons ("_ \<Colon>/ _" [66, 65] 65)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 67 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 68 | primrec | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 69 |   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 | 70 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 71 | "[]\<langle>i\<rangle> = \<bottom>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 72 | | "(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 | 73 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 74 | primrec assoc :: "('a \<times> 'b) list \<Rightarrow> 'a \<Rightarrow> 'b option" ("_\<langle>_\<rangle>\<^isub>?" [90, 0] 91)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 75 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 76 | "[]\<langle>a\<rangle>\<^isub>? = \<bottom>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 77 | | "(x # xs)\<langle>a\<rangle>\<^isub>? = (if fst x = a then \<lfloor>snd x\<rfloor> else xs\<langle>a\<rangle>\<^isub>?)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 78 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 79 | primrec unique :: "('a \<times> 'b) list \<Rightarrow> bool"
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 80 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 81 | "unique [] = True" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 82 | | "unique (x # xs) = (xs\<langle>fst x\<rangle>\<^isub>? = \<bottom> \<and> unique xs)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 83 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 84 | datatype type = | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 85 | TVar nat | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 86 | | Top | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 87 | | Fun type type (infixr "\<rightarrow>" 200) | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 88 |   | TyAll type type  ("(3\<forall><:_./ _)" [0, 10] 10)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 89 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 90 | datatype binding = VarB type | TVarB type | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 91 | types env = "binding list" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 92 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 93 | primrec is_TVarB :: "binding \<Rightarrow> bool" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 94 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 95 | "is_TVarB (VarB T) = False" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 96 | | "is_TVarB (TVarB T) = True" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 97 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 98 | primrec type_ofB :: "binding \<Rightarrow> type" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 99 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 100 | "type_ofB (VarB T) = T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 101 | | "type_ofB (TVarB T) = T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 102 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 103 | primrec mapB :: "(type \<Rightarrow> type) \<Rightarrow> binding \<Rightarrow> binding" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 104 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 105 | "mapB f (VarB T) = VarB (f T)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 106 | | "mapB f (TVarB T) = TVarB (f T)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 107 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 108 | datatype trm = | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 109 | Var nat | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 110 |   | Abs type trm   ("(3\<lambda>:_./ _)" [0, 10] 10)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 111 |   | TAbs type trm  ("(3\<lambda><:_./ _)" [0, 10] 10)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 112 | | App trm trm (infixl "\<bullet>" 200) | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 113 | | TApp trm type (infixl "\<bullet>\<^isub>\<tau>" 200) | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 114 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 115 | primrec liftT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type" ("\<up>\<^isub>\<tau>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 116 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 117 | "\<up>\<^isub>\<tau> n k (TVar i) = (if i < k then TVar i else TVar (i + n))" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 118 | | "\<up>\<^isub>\<tau> n k Top = Top" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 119 | | "\<up>\<^isub>\<tau> n k (T \<rightarrow> U) = \<up>\<^isub>\<tau> n k T \<rightarrow> \<up>\<^isub>\<tau> n k U" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 120 | | "\<up>\<^isub>\<tau> n k (\<forall><:T. U) = (\<forall><:\<up>\<^isub>\<tau> n k T. \<up>\<^isub>\<tau> n (k + 1) U)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 121 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 122 | primrec lift :: "nat \<Rightarrow> nat \<Rightarrow> trm \<Rightarrow> trm" ("\<up>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 123 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 124 | "\<up> n k (Var i) = (if i < k then Var i else Var (i + n))" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 125 | | "\<up> n k (\<lambda>:T. t) = (\<lambda>:\<up>\<^isub>\<tau> n k T. \<up> n (k + 1) t)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 126 | | "\<up> n k (\<lambda><:T. t) = (\<lambda><:\<up>\<^isub>\<tau> n k T. \<up> n (k + 1) t)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 127 | | "\<up> n k (s \<bullet> t) = \<up> n k s \<bullet> \<up> n k t" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 128 | | "\<up> n k (t \<bullet>\<^isub>\<tau> T) = \<up> n k t \<bullet>\<^isub>\<tau> \<up>\<^isub>\<tau> n k T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 129 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 130 | primrec substTT :: "type \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type"  ("_[_ \<mapsto>\<^isub>\<tau> _]\<^isub>\<tau>" [300, 0, 0] 300)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 131 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 132 | "(TVar i)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 133 | (if k < i then TVar (i - 1) else if i = k then \<up>\<^isub>\<tau> k 0 S else TVar i)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 134 | | "Top[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = Top" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 135 | | "(T \<rightarrow> U)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> \<rightarrow> U[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 136 | | "(\<forall><:T. U)[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau> = (\<forall><:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. U[k+1 \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 137 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 138 | primrec decT :: "nat \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> type"  ("\<down>\<^isub>\<tau>")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 139 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 140 | "\<down>\<^isub>\<tau> 0 k T = T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 141 | | "\<down>\<^isub>\<tau> (Suc n) k T = \<down>\<^isub>\<tau> n k (T[k \<mapsto>\<^isub>\<tau> Top]\<^isub>\<tau>)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 142 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 143 | 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 | 144 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 145 | "(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 | 146 | | "(t \<bullet> u)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet> u[k \<mapsto> s]" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 147 | | "(t \<bullet>\<^isub>\<tau> T)[k \<mapsto> s] = t[k \<mapsto> s] \<bullet>\<^isub>\<tau> \<down>\<^isub>\<tau> 1 k T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 148 | | "(\<lambda>:T. t)[k \<mapsto> s] = (\<lambda>:\<down>\<^isub>\<tau> 1 k T. t[k+1 \<mapsto> s])" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 149 | | "(\<lambda><:T. t)[k \<mapsto> s] = (\<lambda><:\<down>\<^isub>\<tau> 1 k T. t[k+1 \<mapsto> s])" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 150 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 151 | primrec substT :: "trm \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> trm"    ("_[_ \<mapsto>\<^isub>\<tau> _]" [300, 0, 0] 300)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 152 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 153 | "(Var i)[k \<mapsto>\<^isub>\<tau> S] = (if k < i then Var (i - 1) else Var i)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 154 | | "(t \<bullet> u)[k \<mapsto>\<^isub>\<tau> S] = t[k \<mapsto>\<^isub>\<tau> S] \<bullet> u[k \<mapsto>\<^isub>\<tau> S]" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 155 | | "(t \<bullet>\<^isub>\<tau> T)[k \<mapsto>\<^isub>\<tau> S] = t[k \<mapsto>\<^isub>\<tau> S] \<bullet>\<^isub>\<tau> T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 156 | | "(\<lambda>:T. t)[k \<mapsto>\<^isub>\<tau> S] = (\<lambda>:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. t[k+1 \<mapsto>\<^isub>\<tau> S])" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 157 | | "(\<lambda><:T. t)[k \<mapsto>\<^isub>\<tau> S] = (\<lambda><:T[k \<mapsto>\<^isub>\<tau> S]\<^isub>\<tau>. t[k+1 \<mapsto>\<^isub>\<tau> S])" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 158 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 159 | primrec liftE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env" ("\<up>\<^isub>e")
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 160 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 161 | "\<up>\<^isub>e n k [] = []" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 162 | | "\<up>\<^isub>e n k (B \<Colon> \<Gamma>) = mapB (\<up>\<^isub>\<tau> n (k + \<parallel>\<Gamma>\<parallel>)) B \<Colon> \<up>\<^isub>e n k \<Gamma>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 163 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 164 | primrec substE :: "env \<Rightarrow> nat \<Rightarrow> type \<Rightarrow> env"  ("_[_ \<mapsto>\<^isub>\<tau> _]\<^isub>e" [300, 0, 0] 300)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 165 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 166 | "[][k \<mapsto>\<^isub>\<tau> T]\<^isub>e = []" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 167 | | "(B \<Colon> \<Gamma>)[k \<mapsto>\<^isub>\<tau> T]\<^isub>e = mapB (\<lambda>U. U[k + \<parallel>\<Gamma>\<parallel> \<mapsto>\<^isub>\<tau> T]\<^isub>\<tau>) B \<Colon> \<Gamma>[k \<mapsto>\<^isub>\<tau> T]\<^isub>e" | 
| 
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 decE :: "nat \<Rightarrow> nat \<Rightarrow> env \<Rightarrow> env"  ("\<down>\<^isub>e")
 | 
| 
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 | "\<down>\<^isub>e 0 k \<Gamma> = \<Gamma>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 172 | | "\<down>\<^isub>e (Suc n) k \<Gamma> = \<down>\<^isub>e n k (\<Gamma>[k \<mapsto>\<^isub>\<tau> Top]\<^isub>e)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 173 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 174 | inductive | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 175 |   well_formed :: "env \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile>\<^bsub>wf\<^esub> _" [50, 50] 50)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 176 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 177 | wf_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB T\<rfloor> \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> TVar i" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 178 | | wf_Top: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> Top" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 179 | | wf_arrow: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> T \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> U \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> T \<rightarrow> U" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 180 | | wf_all: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> T \<Longrightarrow> TVarB T \<Colon> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> U \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> (\<forall><:T. U)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 181 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 182 | inductive | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 183 |   well_formedE :: "env \<Rightarrow> bool"  ("_ \<turnstile>\<^bsub>wf\<^esub>" [50] 50)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 184 |   and well_formedB :: "env \<Rightarrow> binding \<Rightarrow> bool"  ("_ \<turnstile>\<^bsub>wfB\<^esub> _" [50, 50] 50)
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 185 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 186 | "\<Gamma> \<turnstile>\<^bsub>wfB\<^esub> B \<equiv> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> type_ofB B" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 187 | | wf_Nil: "[] \<turnstile>\<^bsub>wf\<^esub>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 188 | | wf_Cons: "\<Gamma> \<turnstile>\<^bsub>wfB\<^esub> B \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> B \<Colon> \<Gamma> \<turnstile>\<^bsub>wf\<^esub>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 189 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 190 | inductive_cases well_formed_cases: | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 191 | "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> TVar i" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 192 | "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> Top" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 193 | "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> T \<rightarrow> U" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 194 | "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> (\<forall><:T. U)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 195 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 196 | inductive_cases well_formedE_cases: | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 197 | "B \<Colon> \<Gamma> \<turnstile>\<^bsub>wf\<^esub>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 198 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 199 | inductive | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 200 |   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 | 201 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 202 | SA_Top: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> S \<Longrightarrow> \<Gamma> \<turnstile> S <: Top" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 203 | | SA_refl_TVar: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma> \<turnstile>\<^bsub>wf\<^esub> TVar i \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: TVar i" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 204 | | SA_trans_TVar: "\<Gamma>\<langle>i\<rangle> = \<lfloor>TVarB U\<rfloor> \<Longrightarrow> | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 205 | \<Gamma> \<turnstile> \<up>\<^isub>\<tau> (Suc i) 0 U <: T \<Longrightarrow> \<Gamma> \<turnstile> TVar i <: T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 206 | | SA_arrow: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> S\<^isub>1 \<rightarrow> S\<^isub>2 <: T\<^isub>1 \<rightarrow> T\<^isub>2" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 207 | | SA_all: "\<Gamma> \<turnstile> T\<^isub>1 <: S\<^isub>1 \<Longrightarrow> TVarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> S\<^isub>2 <: T\<^isub>2 \<Longrightarrow> | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 208 | \<Gamma> \<turnstile> (\<forall><:S\<^isub>1. S\<^isub>2) <: (\<forall><:T\<^isub>1. T\<^isub>2)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 209 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 210 | inductive | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 211 |   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 | 212 | where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 213 | T_Var: "\<Gamma> \<turnstile>\<^bsub>wf\<^esub> \<Longrightarrow> \<Gamma>\<langle>i\<rangle> = \<lfloor>VarB U\<rfloor> \<Longrightarrow> T = \<up>\<^isub>\<tau> (Suc i) 0 U \<Longrightarrow> \<Gamma> \<turnstile> Var i : T" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 214 | | T_Abs: "VarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda>:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \<rightarrow> \<down>\<^isub>\<tau> 1 0 T\<^isub>2" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 215 | | T_App: "\<Gamma> \<turnstile> t\<^isub>1 : T\<^isub>1\<^isub>1 \<rightarrow> T\<^isub>1\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>1\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> t\<^isub>1 \<bullet> t\<^isub>2 : T\<^isub>1\<^isub>2" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 216 | | T_TAbs: "TVarB T\<^isub>1 \<Colon> \<Gamma> \<turnstile> t\<^isub>2 : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> (\<lambda><:T\<^isub>1. t\<^isub>2) : (\<forall><:T\<^isub>1. T\<^isub>2)" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 217 | | T_TApp: "\<Gamma> \<turnstile> t\<^isub>1 : (\<forall><:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2) \<Longrightarrow> \<Gamma> \<turnstile> T\<^isub>2 <: T\<^isub>1\<^isub>1 \<Longrightarrow> | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 218 | \<Gamma> \<turnstile> t\<^isub>1 \<bullet>\<^isub>\<tau> T\<^isub>2 : T\<^isub>1\<^isub>2[0 \<mapsto>\<^isub>\<tau> T\<^isub>2]\<^isub>\<tau>" | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 219 | | 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 | 220 | |
| 36257 
3f3e9f18f302
adopting examples to changes in the predicate compiler
 bulwahn parents: 
36054diff
changeset | 221 | code_pred [inductify, skip_proof, specialise] typing . | 
| 36033 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 222 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 223 | thm typing.equation | 
| 
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 | values 6 "{(E, t, T). typing E t T}"
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 226 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 227 | subsection {* Higher-order predicate *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 228 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 229 | code_pred [inductify] mapB . | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 230 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 231 | subsection {* Multiple instances *}
 | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 232 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 233 | inductive subtype_refl' where | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 234 | "\<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 | 235 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 236 | 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 | 237 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 238 | thm subtype_refl'.equation | 
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 239 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 240 | |
| 
7106f079bd05
adding specialisation examples of the predicate compiler
 bulwahn parents: diff
changeset | 241 | end |